Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *   INRIA, CNRS and contributors - Copyright 1999-2018       *)
(* <O___,, *       (see CREDITS file for the list of authors)           *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)

(* Finite map library.  *)

FMapAVL

This module implements maps using AVL trees. It follows the implementation from Ocaml's standard library.
See the comments at the beginning of FSetAVL for more details.
Require Import FunInd FMapInterface FMapList ZArith Int.

Set Implicit Arguments.
Unset Strict Implicit.
Notations and helper lemma about pairs
Declare Scope pair_scope.
Notation "s #1" := (fst s) (at level 9, format "s '#1'") : pair_scope.
Notation "s #2" := (snd s) (at level 9, format "s '#2'") : pair_scope.

The Raw functor

Functor of pure functions + separate proofs of invariant preservation
Module Raw (Import I:Int)(X: OrderedType).
Local Open Scope pair_scope.
Local Open Scope lazy_bool_scope.
Local Open Scope Int_scope.
Notation int := I.t.

Definition key := X.t.
Hint Transparent key : core.

Trees

Trees

The fifth field of Node is the height of the tree
#[universes(template)]
Inductive tree {elt : Type} :=
  | Leaf : tree
  | Node : tree -> key -> elt -> tree -> int -> tree.
Arguments tree : clear implicits.

Section Elt.

Variable elt : Type.

Notation t := (tree elt).

Implicit Types m : t.

Basic functions on trees: height and cardinal

Definition height (m : t) : int :=
  match m with
  | Leaf => 0
  | Node _ _ _ _ h => h
  end.

Fixpoint cardinal (m : t) : nat :=
  match m with
   | Leaf => 0%nat
   | Node l _ _ r _ => S (cardinal l + cardinal r)
  end.

Empty Map

Definition empty : t := Leaf.

Emptyness test

Definition is_empty m := match m with Leaf => true | _ => false end.

Membership

The mem function is deciding membership. It exploits the bst property to achieve logarithmic complexity.
Fixpoint mem x m : bool :=
   match m with
     |  Leaf => false
     |  Node l y _ r _ => match X.compare x y with
             | LT _ => mem x l
             | EQ _ => true
             | GT _ => mem x r
         end
   end.

Fixpoint find x m : option elt :=
   match m with
     |  Leaf => None
     |  Node l y d r _ => match X.compare x y with
             | LT _ => find x l
             | EQ _ => Some d
             | GT _ => find x r
         end
   end.

Helper functions

create l x r creates a node, assuming l and r to be balanced and |height l - height r| 2.
Definition create l x e r :=
   Node l x e r (max (height l) (height r) + 1).
bal l x e r acts as create, but performs one step of rebalancing if necessary, i.e. assumes |height l - height r| 3.
Definition assert_false := create.

Fixpoint bal l x d r :=
  let hl := height l in
  let hr := height r in
  if gt_le_dec hl (hr+2) then
    match l with
     | Leaf => assert_false l x d r
     | Node ll lx ld lr _ =>
       if ge_lt_dec (height ll) (height lr) then
         create ll lx ld (create lr x d r)
       else
         match lr with
          | Leaf => assert_false l x d r
          | Node lrl lrx lrd lrr _ =>
              create (create ll lx ld lrl) lrx lrd (create lrr x d r)
         end
    end
  else
    if gt_le_dec hr (hl+2) then
      match r with
       | Leaf => assert_false l x d r
       | Node rl rx rd rr _ =>
         if ge_lt_dec (height rr) (height rl) then
            create (create l x d rl) rx rd rr
         else
           match rl with
            | Leaf => assert_false l x d r
            | Node rll rlx rld rlr _ =>
                create (create l x d rll) rlx rld (create rlr rx rd rr)
           end
      end
    else
      create l x d r.

Insertion

Fixpoint add x d m :=
  match m with
   | Leaf => Node Leaf x d Leaf 1
   | Node l y d' r h =>
      match X.compare x y with
         | LT _ => bal (add x d l) y d' r
         | EQ _ => Node l y d r h
         | GT _ => bal l y d' (add x d r)
      end
  end.

Extraction of minimum binding

Morally, remove_min is to be applied to a non-empty tree t = Node l x e r h. Since we can't deal here with assert false for t=Leaf, we pre-unpack t (and forget about h).
Fixpoint remove_min l x d r : t*(key*elt) :=
  match l with
    | Leaf => (r,(x,d))
    | Node ll lx ld lr lh =>
       let (l',m) := remove_min ll lx ld lr in
       (bal l' x d r, m)
  end.

Merging two trees

merge t1 t2 builds the union of t1 and t2 assuming all elements of t1 to be smaller than all elements of t2, and |height t1 - height t2| 2.
Fixpoint merge s1 s2 :=  match s1,s2 with
  | Leaf, _ => s2
  | _, Leaf => s1
  | _, Node l2 x2 d2 r2 h2 =>
    match remove_min l2 x2 d2 r2 with
      (s2',(x,d)) => bal s1 x d s2'
    end
end.

Deletion

Fixpoint remove x m := match m with
  | Leaf => Leaf
  | Node l y d r h =>
      match X.compare x y with
         | LT _ => bal (remove x l) y d r
         | EQ _ => merge l r
         | GT _ => bal l y d (remove x r)
      end
   end.

join

Same as bal but does not assume anything regarding heights of l and r.
Fixpoint join l : key -> elt -> t -> t :=
  match l with
    | Leaf => add
    | Node ll lx ld lr lh => fun x d =>
       fix join_aux (r:t) : t := match r with
          | Leaf =>  add x d l
          | Node rl rx rd rr rh =>
               if gt_le_dec lh (rh+2) then bal ll lx ld (join lr x d r)
               else if gt_le_dec rh (lh+2) then bal (join_aux rl) rx rd rr
               else create l x d r
          end
  end.

Splitting

split x m returns a triple (l, o, r) where
Record triple := mktriple { t_left:t; t_opt:option elt; t_right:t }.
Notation "<< l , b , r >>" := (mktriple l b r) (at level 9).

Fixpoint split x m : triple := match m with
  | Leaf => << Leaf, None, Leaf >>
  | Node l y d r h =>
     match X.compare x y with
      | LT _ => let (ll,o,rl) := split x l in << ll, o, join rl y d r >>
      | EQ _ => << l, Some d, r >>
      | GT _ => let (rl,o,rr) := split x r in << join l y d rl, o, rr >>
     end
 end.

Concatenation

Same as merge but does not assume anything about heights.
Definition concat m1 m2 :=
   match m1, m2 with
      | Leaf, _ => m2
      | _ , Leaf => m1
      | _, Node l2 x2 d2 r2 _ =>
            let (m2',xd) := remove_min l2 x2 d2 r2 in
            join m1 xd#1 xd#2 m2'
   end.

Elements

elements_tree_aux acc t catenates the elements of t in infix order to the list acc
Fixpoint elements_aux (acc : list (key*elt)) m : list (key*elt) :=
  match m with
   | Leaf => acc
   | Node l x d r _ => elements_aux ((x,d) :: elements_aux acc r) l
  end.
then elements is an instantiation with an empty acc
Definition elements := elements_aux nil.

Fold

Fixpoint fold (A : Type) (f : key -> elt -> A -> A) (m : t) : A -> A :=
 fun a => match m with
  | Leaf => a
  | Node l x d r _ => fold f r (f x d (fold f l a))
 end.

Comparison

Variable cmp : elt->elt->bool.

Enumeration of the elements of a tree

Inductive enumeration :=
 | End : enumeration
 | More : key -> elt -> t -> enumeration -> enumeration.
cons m e adds the elements of tree m on the head of enumeration e.
Fixpoint cons m e : enumeration :=
 match m with
  | Leaf => e
  | Node l x d r h => cons l (More x d r e)
 end.
One step of comparison of elements
Definition equal_more x1 d1 (cont:enumeration->bool) e2 :=
 match e2 with
 | End => false
 | More x2 d2 r2 e2 =>
     match X.compare x1 x2 with
      | EQ _ => cmp d1 d2 &&& cont (cons r2 e2)
      | _ => false
     end
 end.
Comparison of left tree, middle element, then right tree
Fixpoint equal_cont m1 (cont:enumeration->bool) e2 :=
 match m1 with
  | Leaf => cont e2
  | Node l1 x1 d1 r1 _ =>
     equal_cont l1 (equal_more x1 d1 (equal_cont r1 cont)) e2
  end.
Initial continuation
Definition equal_end e2 := match e2 with End => true | _ => false end.
The complete comparison
Definition equal m1 m2 := equal_cont m1 equal_end (cons m2 End).

End Elt.
Notation t := tree.
Arguments Leaf : clear implicits.
Arguments Node [elt].

Notation "<< l , b , r >>" := (mktriple l b r) (at level 9).
Notation "t #l" := (t_left t) (at level 9, format "t '#l'").
Notation "t #o" := (t_opt t) (at level 9, format "t '#o'").
Notation "t #r" := (t_right t) (at level 9, format "t '#r'").

Map

Fixpoint map (elt elt' : Type)(f : elt -> elt')(m : t elt) : t elt' :=
  match m with
   | Leaf _   => Leaf _
   | Node l x d r h => Node (map f l) x (f d) (map f r) h
  end.

(* * Mapi *)

Fixpoint mapi (elt elt' : Type)(f : key -> elt -> elt')(m : t elt) : t elt' :=
  match m with
   | Leaf _ => Leaf _
   | Node l x d r h => Node (mapi f l) x (f x d) (mapi f r) h
  end.

Map with removal

Fixpoint map_option (elt elt' : Type)(f : key -> elt -> option elt')(m : t elt)
  : t elt' :=
  match m with
   | Leaf _ => Leaf _
   | Node l x d r h =>
      match f x d with
       | Some d' => join (map_option f l) x d' (map_option f r)
       | None => concat (map_option f l) (map_option f r)
      end
  end.

Optimized map2

Suggestion by B. Gregoire: a map2 function with specialized arguments that allows bypassing some tree traversal. Instead of one f0 of type key option elt option elt' option elt'', we ask here for:
The idea is that mapl and mapr can be instantaneous (e.g. the identity or some constant function).
Section Map2_opt.
Variable elt elt' elt'' : Type.
Variable f : key -> elt -> option elt' -> option elt''.
Variable mapl : t elt -> t elt''.
Variable mapr : t elt' -> t elt''.

Fixpoint map2_opt m1 m2 :=
 match m1, m2 with
  | Leaf _, _ => mapr m2
  | _, Leaf _ => mapl m1
  | Node l1 x1 d1 r1 h1, _ =>
     let (l2',o2,r2') := split x1 m2 in
     match f x1 d1 o2 with
      | Some e => join (map2_opt l1 l2') x1 e (map2_opt r1 r2')
      | None => concat (map2_opt l1 l2') (map2_opt r1 r2')
     end
 end.

End Map2_opt.

Map2

The map2 function of the Map interface can be implemented via map2_opt and map_option.
Section Map2.
Variable elt elt' elt'' : Type.
Variable f : option elt -> option elt' -> option elt''.

Definition map2 : t elt -> t elt' -> t elt'' :=
 map2_opt
   (fun _ d o => f (Some d) o)
   (map_option (fun _ d => f (Some d) None))
   (map_option (fun _ d' => f None (Some d'))).

End Map2.

Invariants

Section Invariants.
Variable elt : Type.

Occurrence in a tree

Inductive MapsTo (x : key)(e : elt) : t elt -> Prop :=
  | MapsRoot : forall l r h y,
      X.eq x y -> MapsTo x e (Node l y e r h)
  | MapsLeft : forall l r h y e',
      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) : t elt -> Prop :=
  | InRoot : forall l r h y e,
      X.eq x y -> In x (Node l y e r h)
  | InLeft : forall l r h y e',
      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 m := 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 m := forall y, In y m -> X.lt y x.
Definition gt_tree x m := forall y, In y m -> X.lt x y.
bst t : t is a binary search tree
Inductive bst : t elt -> Prop :=
  | BSLeaf : bst (Leaf _)
  | BSNode : forall x e l r h, bst l -> bst r ->
     lt_tree x l -> gt_tree x r -> bst (Node l x e r h).

End Invariants.

Correctness proofs, isolated in a sub-module

Module Proofs.
 Module MX := OrderedTypeFacts X.
 Module PX := KeyOrderedType X.
 Module L := FMapList.Raw X.

Functional Scheme mem_ind := Induction for mem Sort Prop.
Functional Scheme find_ind := Induction for find Sort Prop.
Functional Scheme bal_ind := Induction for bal Sort Prop.
Functional Scheme add_ind := Induction for add Sort Prop.
Functional Scheme remove_min_ind := Induction for remove_min Sort Prop.
Functional Scheme merge_ind := Induction for merge Sort Prop.
Functional Scheme remove_ind := Induction for remove Sort Prop.
Functional Scheme concat_ind := Induction for concat Sort Prop.
Functional Scheme split_ind := Induction for split Sort Prop.
Functional Scheme map_option_ind := Induction for map_option Sort Prop.
Functional Scheme map2_opt_ind := Induction for map2_opt Sort Prop.

Automation and dedicated tactics.

Hint Constructors tree MapsTo In bst : core.
Hint Unfold lt_tree gt_tree : core.

Tactic Notation "factornode" ident(l) ident(x) ident(d) ident(r) ident(h)
 "as" ident(s) :=
 set (s:=Node l x d r h) in *; clearbody s; clear l x d r h.
A tactic for cleaning hypothesis after use of functional induction.
Ltac clearf :=
 match goal with
  | H : (@Logic.eq (Compare _ _ _ _) _ _) |- _ => clear H; clearf
  | H : (@Logic.eq (sumbool _ _) _ _) |- _ => clear H; clearf
  | _ => idtac
 end.
A tactic to repeat inversion_clear on all hyps of the form (f (Node ...))
Ltac inv f :=
  match goal with
     | H:f (Leaf _) |- _ => inversion_clear H; inv f
     | 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 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.
Helper tactic concerning order of elements.
Ltac order := match goal with
 | U: lt_tree _ ?s, V: In _ ?s |- _ => generalize (U _ V); clear U; order
 | U: gt_tree _ ?s, V: In _ ?s |- _ => generalize (U _ V); clear U; order
 | _ => MX.order
end.

Ltac intuition_in := repeat (intuition; inv In; inv MapsTo).

(* Function/Functional Scheme can't deal with internal fix.
   Let's do its job by hand: *)

Ltac join_tac :=
 intros l; induction l as [| ll _ lx ld lr Hlr lh];
   [ | intros x d r; induction r as [| rl Hrl rx rd rr _ rh]; unfold join;
     [ | destruct (gt_le_dec lh (rh+2)) as [GT|LE];
       [ match goal with |- context [ bal ?u ?v ?w ?z ] =>
           replace (bal u v w z)
           with (bal ll lx ld (join lr x d (Node rl rx rd rr rh))); [ | auto]
         end
       | destruct (gt_le_dec rh (lh+2)) as [GT'|LE'];
         [ match goal with |- context [ bal ?u ?v ?w ?z ] =>
             replace (bal u v w z)
             with (bal (join (Node ll lx ld lr lh) x d rl) rx rd rr); [ | auto]
           end
         | ] ] ] ]; intros.

Section Elt.
Variable elt:Type.
Implicit Types m r : t elt.

Basic results about MapsTo, In, lt_tree, gt_tree, height

Facts about MapsTo and In.
elt:Type

forall (k : key) (e : elt) (m : t elt), MapsTo k e m -> In k m
elt:Type

forall (k : key) (e : elt) (m : t elt), MapsTo k e m -> In k m
induction 1; auto. Qed. Hint Resolve MapsTo_In : core.
elt:Type

forall (k : key) (m : t elt), In k m -> exists e : elt, MapsTo k e m
elt:Type

forall (k : key) (m : t elt), In k m -> exists e : elt, MapsTo k e m
induction 1; try destruct IHIn as (e,He); exists e; auto. Qed.
elt:Type

forall (k : key) (m : t elt), In0 k m <-> In k m
elt:Type

forall (k : key) (m : t elt), In0 k m <-> In k m
elt:Type
k:key
m:t elt

In0 k m -> In k m
elt:Type
k:key
m:t elt
In k m -> In0 k m
elt:Type
k:key
m:t elt

In k m -> In0 k m
unfold In0; apply In_MapsTo; auto. Qed.
elt:Type

forall (m : t elt) (x y : X.t) (e : elt), X.eq x y -> MapsTo x e m -> MapsTo y e m
elt:Type

forall (m : t elt) (x y : X.t) (e : elt), X.eq x y -> MapsTo x e m -> MapsTo y e m
induction m; simpl; intuition_in; eauto. Qed. Hint Immediate MapsTo_1 : core.
elt:Type

forall (m : t elt) (x y : X.t), X.eq x y -> In x m -> In y m
elt:Type

forall (m : t elt) (x y : X.t), X.eq x y -> In x m -> In y m
intros m x y; induction m; simpl; intuition_in; eauto. Qed.
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int) (y : key), In y (Node l x e r h) <-> In y l \/ X.eq y x \/ In y r
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int) (y : key), In y (Node l x e r h) <-> In y l \/ X.eq y x \/ In y r
intuition_in. Qed.
Results about lt_tree and gt_tree
elt:Type

forall x : X.t, lt_tree x (Leaf elt)
elt:Type

forall x : X.t, lt_tree x (Leaf elt)
unfold lt_tree; intros; intuition_in. Qed.
elt:Type

forall x : X.t, gt_tree x (Leaf elt)
elt:Type

forall x : X.t, gt_tree x (Leaf elt)
unfold gt_tree; intros; intuition_in. Qed.
elt:Type

forall (x y : X.t) (l r : t elt) (e : elt) (h : int), lt_tree x l -> lt_tree x r -> X.lt y x -> lt_tree x (Node l y e r h)
elt:Type

forall (x y : X.t) (l r : t elt) (e : elt) (h : int), lt_tree x l -> lt_tree x r -> X.lt y x -> lt_tree x (Node l y e r h)
unfold lt_tree in *; intuition_in; order. Qed.
elt:Type

forall (x y : X.t) (l r : t elt) (e : elt) (h : int), gt_tree x l -> gt_tree x r -> X.lt x y -> gt_tree x (Node l y e r h)
elt:Type

forall (x y : X.t) (l r : t elt) (e : elt) (h : int), gt_tree x l -> gt_tree x r -> X.lt x y -> gt_tree x (Node l y e r h)
unfold gt_tree in *; intuition_in; order. Qed. Hint Resolve lt_leaf gt_leaf lt_tree_node gt_tree_node : core.
elt:Type

forall (x : X.t) (y : key) (l r : t elt) (e : elt) (h : int), lt_tree x (Node l y e r h) -> lt_tree x l
elt:Type

forall (x : X.t) (y : key) (l r : t elt) (e : elt) (h : int), lt_tree x (Node l y e r h) -> lt_tree x l
intuition_in. Qed.
elt:Type

forall (x : X.t) (y : key) (l r : t elt) (e : elt) (h : int), lt_tree x (Node l y e r h) -> lt_tree x r
elt:Type

forall (x : X.t) (y : key) (l r : t elt) (e : elt) (h : int), lt_tree x (Node l y e r h) -> lt_tree x r
intuition_in. Qed.
elt:Type

forall (x : X.t) (y : key) (l r : t elt) (e : elt) (h : int), gt_tree x (Node l y e r h) -> gt_tree x l
elt:Type

forall (x : X.t) (y : key) (l r : t elt) (e : elt) (h : int), gt_tree x (Node l y e r h) -> gt_tree x l
intuition_in. Qed.
elt:Type

forall (x : X.t) (y : key) (l r : t elt) (e : elt) (h : int), gt_tree x (Node l y e r h) -> gt_tree x r
elt:Type

forall (x : X.t) (y : key) (l r : t elt) (e : elt) (h : int), gt_tree x (Node l y e r h) -> gt_tree x r
intuition_in. Qed. Hint Resolve lt_left lt_right gt_left gt_right : core.
elt:Type

forall (x : X.t) (m : t elt), lt_tree x m -> ~ In x m
elt:Type

forall (x : X.t) (m : t elt), lt_tree x m -> ~ In x m
intros; intro; generalize (H _ H0); order. Qed.
elt:Type

forall x y : X.t, X.lt x y -> forall m : t elt, lt_tree x m -> lt_tree y m
elt:Type

forall x y : X.t, X.lt x y -> forall m : t elt, lt_tree x m -> lt_tree y m
eauto. Qed.
elt:Type

forall (x : X.t) (m : t elt), gt_tree x m -> ~ In x m
elt:Type

forall (x : X.t) (m : t elt), gt_tree x m -> ~ In x m
intros; intro; generalize (H _ H0); order. Qed.
elt:Type

forall x y : X.t, X.lt y x -> forall m : t elt, gt_tree x m -> gt_tree y m
elt:Type

forall x y : X.t, X.lt y x -> forall m : t elt, gt_tree x m -> gt_tree y m
eauto. Qed. Hint Resolve lt_tree_not_in lt_tree_trans gt_tree_not_in gt_tree_trans : core.

Empty map

Definition Empty m := forall (a:key)(e:elt) , ~ MapsTo a e m.

elt:Type

bst (empty elt)
elt:Type

bst (empty elt)
unfold empty; auto. Qed.
elt:Type

Empty (empty elt)
elt:Type

Empty (empty elt)
unfold empty, Empty; intuition_in. Qed.

Emptyness test

elt:Type

forall m : t elt, Empty m -> is_empty m = true
elt:Type

forall m : t elt, Empty m -> is_empty m = true
elt:Type
r:t elt
x:key
e:elt
l:t elt
h:int

Empty (Node r x e l h) -> false = true
intro H; elim (H x e); auto. Qed.
elt:Type

forall m : t elt, is_empty m = true -> Empty m
elt:Type

forall m : t elt, is_empty m = true -> Empty m
destruct m; simpl; intros; try discriminate; red; intuition_in. Qed.

Membership

elt:Type

forall (m : t elt) (x : key), bst m -> In x m -> mem x m = true
elt:Type

forall (m : t elt) (x : key), bst m -> In x m -> mem x m = true
intros m x; functional induction (mem x m); auto; intros; clearf; inv bst; intuition_in; order. Qed.
elt:Type

forall (m : t elt) (x : X.t), mem x m = true -> In x m
elt:Type

forall (m : t elt) (x : X.t), mem x m = true -> In x m
intros m x; functional induction (mem x m); auto; intros; discriminate. Qed.
elt:Type

forall (m : t elt) (x : key) (e : elt), bst m -> MapsTo x e m -> find x m = Some e
elt:Type

forall (m : t elt) (x : key) (e : elt), bst m -> MapsTo x e m -> find x m = Some e
intros m x; functional induction (find x m); auto; intros; clearf; inv bst; intuition_in; simpl; auto; try solve [order | absurd (X.lt x y); eauto | absurd (X.lt y x); eauto]. Qed.
elt:Type

forall (m : t elt) (x : X.t) (e : elt), find x m = Some e -> MapsTo x e m
elt:Type

forall (m : t elt) (x : X.t) (e : elt), find x m = Some e -> MapsTo x e m
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
IHo:forall e0 : elt, find x l = Some e0 -> MapsTo x e0 l
e:elt
H:find x l = Some e

MapsTo x e (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e:elt
H:Some d = Some e
MapsTo x e (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
IHo:forall e0 : elt, find x r = Some e0 -> MapsTo x e0 r
e:elt
H:find x r = Some e
MapsTo x e (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e:elt
H:Some d = Some e

MapsTo x e (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
IHo:forall e0 : elt, find x r = Some e0 -> MapsTo x e0 r
e:elt
H:find x r = Some e
MapsTo x e (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
IHo:forall e0 : elt, find x r = Some e0 -> MapsTo x e0 r
e:elt
H:find x r = Some e

MapsTo x e (Node l y d r _x)
constructor 3; auto. Qed.
elt:Type

forall (m : t elt) (x : X.t) (e : elt), bst m -> find x m = Some e <-> MapsTo x e m
elt:Type

forall (m : t elt) (x : X.t) (e : elt), bst m -> find x m = Some e <-> MapsTo x e m
split; auto using find_1, find_2. Qed.
elt:Type

forall (m : t elt) (x : X.t), find x m <> None -> In x m
elt:Type

forall (m : t elt) (x : X.t), find x m <> None -> In x m
elt:Type
m:t elt
x:X.t
H:find x m <> None

In x m
elt:Type
m:t elt
x:X.t
H:find x m <> None
e:elt
H0:find x m = Some e

In x m
apply MapsTo_In with e; apply find_2; auto. Qed.
elt:Type

forall (m : t elt) (x : key), bst m -> In x m -> find x m <> None
elt:Type

forall (m : t elt) (x : key), bst m -> In x m -> find x m <> None
elt:Type
m:t elt
x:key
H:bst m
H0:In x m

find x m <> None
elt:Type
m:t elt
x:key
H:bst m
H0:In x m
d:elt
Hd:MapsTo x d m

find x m <> None
rewrite (find_1 H Hd); discriminate. Qed.
elt:Type

forall (m : t elt) (x : X.t), bst m -> find x m <> None <-> In x m
elt:Type

forall (m : t elt) (x : X.t), bst m -> find x m <> None <-> In x m
split; auto using find_in, in_find. Qed.
elt:Type

forall (m : t elt) (x : X.t), bst m -> find x m = None <-> ~ In x m
elt:Type

forall (m : t elt) (x : X.t), bst m -> find x m = None <-> ~ In x m
elt:Type
m:t elt
x:X.t
H:bst m
H0:find x m = None

~ In x m
elt:Type
m:t elt
x:X.t
H:bst m
H0:~ In x m
find x m = None
elt:Type
m:t elt
x:X.t
H:bst m
H0:find x m = None
H1:In x m

False
elt:Type
m:t elt
x:X.t
H:bst m
H0:~ In x m
find x m = None
elt:Type
m:t elt
x:X.t
H:bst m
H0:~ In x m

find x m = None
elt:Type
m:t elt
x:X.t
H:bst m
H0:~ In x m
e:elt
H1:find x m = Some e

Some e = None
elim H0; apply find_in; congruence. Qed.
elt:Type

forall (m m' : t elt) (x : X.t), find x m = find x m' <-> (forall d : elt, find x m = Some d <-> find x m' = Some d)
elt:Type

forall (m m' : t elt) (x : X.t), find x m = find x m' <-> (forall d : elt, find x m = Some d <-> find x m' = Some d)
elt:Type
m, m':t elt
x:X.t
e, e0:elt
H:forall d : elt, Some e = Some d <-> Some e0 = Some d

Some e = Some e0
elt:Type
m, m':t elt
x:X.t
e:elt
H:forall d : elt, Some e = Some d <-> None = Some d
Some e = None
elt:Type
m, m':t elt
x:X.t
e:elt
H:forall d : elt, None = Some d <-> Some e = Some d
None = Some e
elt:Type
m, m':t elt
x:X.t
e:elt
H:forall d : elt, Some e = Some d <-> None = Some d

Some e = None
elt:Type
m, m':t elt
x:X.t
e:elt
H:forall d : elt, None = Some d <-> Some e = Some d
None = Some e
elt:Type
m, m':t elt
x:X.t
e:elt
H:forall d : elt, None = Some d <-> Some e = Some d

None = Some e
rewrite H; auto. Qed.
elt:Type

forall (m m' : t elt) (x : X.t), bst m -> bst m' -> find x m = find x m' <-> (forall d : elt, MapsTo x d m <-> MapsTo x d m')
elt:Type

forall (m m' : t elt) (x : X.t), bst m -> bst m' -> find x m = find x m' <-> (forall d : elt, MapsTo x d m <-> MapsTo x d m')
elt:Type
m, m':t elt
x:X.t
Hm:bst m
Hm':bst m'

find x m = find x m' <-> (forall d : elt, MapsTo x d m <-> MapsTo x d m')
elt:Type
m, m':t elt
x:X.t
Hm:bst m
Hm':bst m'

(forall d : elt, find x m = Some d <-> find x m' = Some d) <-> (forall d : elt, MapsTo x d m <-> MapsTo x d m')
elt:Type
m, m':t elt
x:X.t
Hm:bst m
Hm':bst m'
d:elt
H:find x m = Some d <-> find x m' = Some d

MapsTo x d m <-> MapsTo x d m'
elt:Type
m, m':t elt
x:X.t
Hm:bst m
Hm':bst m'
d:elt
H:MapsTo x d m <-> MapsTo x d m'
find x m = Some d <-> find x m' = Some d
elt:Type
m, m':t elt
x:X.t
Hm:bst m
Hm':bst m'
d:elt
H:MapsTo x d m <-> MapsTo x d m'

find x m = Some d <-> find x m' = Some d
rewrite 2 find_iff; auto. Qed.
elt:Type

forall (m m' : t elt) (x : X.t), bst m -> bst m' -> find x m = find x m' -> In x m <-> In x m'
elt:Type

forall (m m' : t elt) (x : X.t), bst m -> bst m' -> find x m = find x m' -> In x m <-> In x m'
split; intros; apply find_in; [ rewrite <- H1 | rewrite H1 ]; apply in_find; auto. Qed.

Helper functions

elt:Type

forall (l : t elt) (x : X.t) (e : elt) (r : t elt), bst l -> bst r -> lt_tree x l -> gt_tree x r -> bst (create l x e r)
elt:Type

forall (l : t elt) (x : X.t) (e : elt) (r : t elt), bst l -> bst r -> lt_tree x l -> gt_tree x r -> bst (create l x e r)
unfold create; auto. Qed. Hint Resolve create_bst : core.
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (y : key), In y (create l x e r) <-> X.eq y x \/ In y l \/ In y r
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (y : key), In y (create l x e r) <-> X.eq y x \/ In y l \/ In y r
unfold create; split; [ inversion_clear 1 | ]; intuition. Qed.
elt:Type

forall (l : t elt) (x : X.t) (e : elt) (r : t elt), bst l -> bst r -> lt_tree x l -> gt_tree x r -> bst (bal l x e r)
elt:Type

forall (l : t elt) (x : X.t) (e : elt) (r : t elt), bst l -> bst r -> lt_tree x l -> gt_tree x r -> bst (bal l x e r)
intros l x e r; functional induction (bal l x e r); intros; clearf; inv bst; repeat apply create_bst; auto; unfold create; try constructor; (apply lt_tree_node || apply gt_tree_node); auto; (eapply lt_tree_trans || eapply gt_tree_trans); eauto. Qed. Hint Resolve bal_bst : core.
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (y : key), In y (bal l x e r) <-> X.eq y x \/ In y l \/ In y r
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (y : key), In y (bal l x e r) <-> X.eq y x \/ In y l \/ In y r
intros l x e r; functional induction (bal l x e r); intros; clearf; rewrite !create_in; intuition_in. Qed.
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (y : key) (e' : elt), MapsTo y e' (bal l x e r) <-> MapsTo y e' (create l x e r)
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (y : key) (e' : elt), MapsTo y e' (bal l x e r) <-> MapsTo y e' (create l x e r)
intros l x e r; functional induction (bal l x e r); intros; clearf; unfold assert_false, create; intuition_in. Qed.
elt:Type

forall (l : t elt) (x : X.t) (e : elt) (r : t elt) (y : X.t), bst l -> bst r -> lt_tree x l -> gt_tree x r -> find y (bal l x e r) = find y (create l x e r)
elt:Type

forall (l : t elt) (x : X.t) (e : elt) (r : t elt) (y : X.t), bst l -> bst r -> lt_tree x l -> gt_tree x r -> find y (bal l x e r) = find y (create l x e r)
intros; rewrite find_mapsto_equiv; auto; intros; apply bal_mapsto. Qed.

Insertion

elt:Type

forall (m : t elt) (x y : key) (e : elt), In y (add x e m) <-> X.eq y x \/ In y m
elt:Type

forall (m : t elt) (x y : key) (e : elt), In y (add x e m) <-> X.eq y x \/ In y m
elt:Type
x, y:key
e:elt
l:t elt
y0:key
d':elt
r:t elt
h:int
_x:X.eq x y0
e1:X.compare x y0 = EQ _x
H0:X.eq y x

In y (Node l y0 e r h)
apply In_1 with x; auto. Qed.
elt:Type

forall (m : t elt) (x : key) (e : elt), bst m -> bst (add x e m)
elt:Type

forall (m : t elt) (x : key) (e : elt), bst m -> bst (add x e m)
elt:Type
x:key
e:elt
l:t elt
y:key
d':elt
r:t elt
h:int
_x:X.lt x y
e1:X.compare x y = LT _x
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
z:key
H4:X.eq z x
H:bst (add x e l)

X.lt z y
elt:Type
x:key
e:elt
l:t elt
y:key
d':elt
r:t elt
h:int
_x:X.lt y x
e1:X.compare x y = GT _x
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
z:key
H4:X.eq z x
H:bst (add x e r)
X.lt y z
elt:Type
x:key
e:elt
l:t elt
y:key
d':elt
r:t elt
h:int
_x:X.lt y x
e1:X.compare x y = GT _x
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
z:key
H4:X.eq z x
H:bst (add x e r)

X.lt y z
apply MX.lt_eq with x; auto. Qed. Hint Resolve add_bst : core.
elt:Type

forall (m : t elt) (x y : X.t) (e : elt), X.eq x y -> MapsTo y e (add x e m)
elt:Type

forall (m : t elt) (x y : X.t) (e : elt), X.eq x y -> MapsTo y e (add x e m)
intros m x y e; functional induction (add x e m); intros; inv bst; try rewrite bal_mapsto; unfold create; eauto. Qed.
elt:Type

forall (m : t elt) (x y : X.t) (e e' : elt), ~ X.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m)
elt:Type

forall (m : t elt) (x y : X.t) (e e' : elt), ~ X.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m)
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x, y:X.t
e, e':elt
IHm1:~ X.eq x y -> MapsTo y e m1 -> MapsTo y e (add x e' m1)
IHm2:~ X.eq x y -> MapsTo y e m2 -> MapsTo y e (add x e' m2)

~ X.eq x y -> MapsTo y e (Node m1 k e0 m2 t) -> MapsTo y e match X.compare x k with | LT _ => bal (add x e' m1) k e0 m2 | EQ _ => Node m1 k e' m2 t | GT _ => bal m1 k e0 (add x e' m2) end
destruct (X.compare x k); intros; inv bst; try rewrite bal_mapsto; unfold create; auto; inv MapsTo; auto; order. Qed.
elt:Type

forall (m : t elt) (x y : X.t) (e e' : elt), ~ X.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e m
elt:Type

forall (m : t elt) (x y : X.t) (e e' : elt), ~ X.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e m
elt:Type
x, y:X.t
e, e':elt

~ X.eq x y -> MapsTo y e (Node (Leaf elt) x e' (Leaf elt) 1) -> MapsTo y e (Leaf elt)
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x, y:X.t
e, e':elt
IHm1:~ X.eq x y -> MapsTo y e (add x e' m1) -> MapsTo y e m1
IHm2:~ X.eq x y -> MapsTo y e (add x e' m2) -> MapsTo y e m2
~ X.eq x y -> MapsTo y e match X.compare x k with | LT _ => bal (add x e' m1) k e0 m2 | EQ _ => Node m1 k e' m2 t | GT _ => bal m1 k e0 (add x e' m2) end -> MapsTo y e (Node m1 k e0 m2 t)
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x, y:X.t
e, e':elt
IHm1:~ X.eq x y -> MapsTo y e (add x e' m1) -> MapsTo y e m1
IHm2:~ X.eq x y -> MapsTo y e (add x e' m2) -> MapsTo y e m2

~ X.eq x y -> MapsTo y e match X.compare x k with | LT _ => bal (add x e' m1) k e0 m2 | EQ _ => Node m1 k e' m2 t | GT _ => bal m1 k e0 (add x e' m2) end -> MapsTo y e (Node m1 k e0 m2 t)
destruct (X.compare x k); intro; try rewrite bal_mapsto; auto; unfold create; intros; inv MapsTo; auto; order. Qed.
elt:Type

forall (m : t elt) (x : key) (y : X.t) (e : elt), bst m -> find y (add x e m) = match X.compare y x with | EQ _ => Some e | _ => find y m end
elt:Type

forall (m : t elt) (x : key) (y : X.t) (e : elt), bst m -> find y (add x e m) = match X.compare y x with | EQ _ => Some e | _ => find y m end
elt:Type
m:t elt
x:key
y:X.t
e:elt
H:bst m

find y (add x e m) = match X.compare y x with | EQ _ => Some e | _ => find y m end
elt:Type
m:t elt
x:key
y:X.t
e:elt
H:bst m

~ X.eq x y -> find y (add x e m) = find y m
elt:Type
m:t elt
x:key
y:X.t
e:elt
H:bst m
H0:~ X.eq x y -> find y (add x e m) = find y m
find y (add x e m) = match X.compare y x with | EQ _ => Some e | _ => find y m end
elt:Type
m:t elt
x:key
y:X.t
e:elt
H:bst m
H0:~ X.eq x y

forall d : elt, MapsTo y d (add x e m) <-> MapsTo y d m
elt:Type
m:t elt
x:key
y:X.t
e:elt
H:bst m
H0:~ X.eq x y -> find y (add x e m) = find y m
find y (add x e m) = match X.compare y x with | EQ _ => Some e | _ => find y m end
elt:Type
m:t elt
x:key
y:X.t
e:elt
H:bst m
H0:~ X.eq x y -> find y (add x e m) = find y m

find y (add x e m) = match X.compare y x with | EQ _ => Some e | _ => find y m end
elt:Type
m:t elt
x:key
y:X.t
e:elt
H:bst m
H0:~ X.eq x y -> find y (add x e m) = find y m
e0:X.eq y x

find y (add x e m) = Some e
auto using find_1, add_1. Qed.

Extraction of minimum binding

elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int) (y : key), In y (Node l x e r h) <-> X.eq y ((remove_min l x e r)#2)#1 \/ In y (remove_min l x e r)#1
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int) (y : key), In y (Node l x e r h) <-> X.eq y ((remove_min l x e r)#2)#1 \/ In y (remove_min l x e r)#1
elt:Type
x:key
d:elt
r:t elt
h:int
y:key

In y (Node (Leaf elt) x d r h) <-> X.eq y x \/ In y r
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall (h0 : int) (y0 : key), In y0 (Node ll lx ld lr h0) <-> X.eq y0 ((remove_min ll lx ld lr)#2)#1 \/ In y0 (remove_min ll lx ld lr)#1
h:int
y:key
In y (Node (Node ll lx ld lr _x) x d r h) <-> X.eq y m#1 \/ In y (bal l' x d r)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall (h0 : int) (y0 : key), In y0 (Node ll lx ld lr h0) <-> X.eq y0 ((remove_min ll lx ld lr)#2)#1 \/ In y0 (remove_min ll lx ld lr)#1
h:int
y:key

In y (Node (Node ll lx ld lr _x) x d r h) <-> X.eq y m#1 \/ In y (bal l' x d r)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall (h0 : int) (y0 : key), In y0 (Node ll lx ld lr h0) <-> X.eq y0 ((l', m)#2)#1 \/ In y0 (l', m)#1
h:int
y:key

In y (Node (Node ll lx ld lr _x) x d r h) <-> X.eq y m#1 \/ In y (bal l' x d r)
rewrite bal_in, In_node_iff, IHp; intuition. Qed.
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int) (y : key) (e' : elt), MapsTo y e' (Node l x e r h) <-> X.eq y ((remove_min l x e r)#2)#1 /\ e' = ((remove_min l x e r)#2)#2 \/ MapsTo y e' (remove_min l x e r)#1
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int) (y : key) (e' : elt), MapsTo y e' (Node l x e r h) <-> X.eq y ((remove_min l x e r)#2)#1 /\ e' = ((remove_min l x e r)#2)#2 \/ MapsTo y e' (remove_min l x e r)#1
elt:Type
x:key
d:elt
r:t elt
h:int
y:key
e':elt

MapsTo y e' (Node (Leaf elt) x d r h) <-> X.eq y x /\ e' = d \/ MapsTo y e' r
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall (h0 : int) (y0 : key) (e'0 : elt), MapsTo y0 e'0 (Node ll lx ld lr h0) <-> X.eq y0 ((remove_min ll lx ld lr)#2)#1 /\ e'0 = ((remove_min ll lx ld lr)#2)#2 \/ MapsTo y0 e'0 (remove_min ll lx ld lr)#1
h:int
y:key
e':elt
MapsTo y e' (Node (Node ll lx ld lr _x) x d r h) <-> X.eq y m#1 /\ e' = m#2 \/ MapsTo y e' (bal l' x d r)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall (h0 : int) (y0 : key) (e'0 : elt), MapsTo y0 e'0 (Node ll lx ld lr h0) <-> X.eq y0 ((remove_min ll lx ld lr)#2)#1 /\ e'0 = ((remove_min ll lx ld lr)#2)#2 \/ MapsTo y0 e'0 (remove_min ll lx ld lr)#1
h:int
y:key
e':elt

MapsTo y e' (Node (Node ll lx ld lr _x) x d r h) <-> X.eq y m#1 /\ e' = m#2 \/ MapsTo y e' (bal l' x d r)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall (h0 : int) (y0 : key) (e'0 : elt), MapsTo y0 e'0 (Node ll lx ld lr h0) <-> X.eq y0 ((l', m)#2)#1 /\ e'0 = ((l', m)#2)#2 \/ MapsTo y0 e'0 (l', m)#1
h:int
y:key
e':elt

MapsTo y e' (Node (Node ll lx ld lr _x) x d r h) <-> X.eq y m#1 /\ e' = m#2 \/ MapsTo y e' (bal l' x d r)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall (h0 : int) (y0 : key) (e'0 : elt), MapsTo y0 e'0 (Node ll lx ld lr h0) <-> X.eq y0 ((l', m)#2)#1 /\ e'0 = ((l', m)#2)#2 \/ MapsTo y0 e'0 (l', m)#1
h:int
y:key
e':elt

MapsTo y e' (Node (Node ll lx ld lr _x) x d r h) <-> X.eq y m#1 /\ e' = m#2 \/ MapsTo y e' (Node l' x d r (max (height l') (height r) + 1))
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall (h0 : int) (y0 : key) (e'0 : elt), MapsTo y0 e'0 (Node ll lx ld lr h0) <-> X.eq y0 m#1 /\ e'0 = m#2 \/ MapsTo y0 e'0 l'
h:int
y:key
e':elt
H:MapsTo y e' (Node ll lx ld lr _x) -> X.eq y m#1 /\ e' = m#2 \/ MapsTo y e' l'
H0:X.eq y m#1 /\ e' = m#2 \/ MapsTo y e' l' -> MapsTo y e' (Node ll lx ld lr _x)

MapsTo y e' (Node (Node ll lx ld lr _x) x d r h) <-> X.eq y m#1 /\ e' = m#2 \/ MapsTo y e' (Node l' x d r (max (height l') (height r) + 1))
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall (h0 : int) (y0 : key) (e'0 : elt), MapsTo y0 e'0 (Node ll lx ld lr h0) <-> X.eq y0 m#1 /\ e'0 = m#2 \/ MapsTo y0 e'0 l'
h:int
y:key
e':elt
H:MapsTo y e' (Node ll lx ld lr _x) -> X.eq y m#1 /\ e' = m#2 \/ MapsTo y e' l'
H2:MapsTo y e' l' -> MapsTo y e' (Node ll lx ld lr _x)
H0:X.eq y m#1 -> e' = m#2 -> MapsTo y e' (Node ll lx ld lr _x)
H1:MapsTo y e' (Node (Node ll lx ld lr _x) x d r h)

X.eq y m#1 /\ e' = m#2 \/ MapsTo y e' (Node l' x d r (max (height l') (height r) + 1))
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall (h0 : int) (y0 : key) (e'0 : elt), MapsTo y0 e'0 (Node ll lx ld lr h0) <-> X.eq y0 m#1 /\ e'0 = m#2 \/ MapsTo y0 e'0 l'
h:int
y:key
e':elt
H:MapsTo y e' (Node ll lx ld lr _x) -> X.eq y m#1 /\ e' = m#2 \/ MapsTo y e' l'
H2:MapsTo y e' l' -> MapsTo y e' (Node ll lx ld lr _x)
H0:X.eq y m#1 -> e' = m#2 -> MapsTo y e' (Node ll lx ld lr _x)
H3:MapsTo y e' (Node l' x d r (max (height l') (height r) + 1))
MapsTo y e' (Node (Node ll lx ld lr _x) x d r h)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall (h0 : int) (y0 : key) (e'0 : elt), MapsTo y0 e'0 (Node ll lx ld lr h0) <-> X.eq y0 m#1 /\ e'0 = m#2 \/ MapsTo y0 e'0 l'
h:int
y:key
e':elt
H:MapsTo y e' (Node ll lx ld lr _x) -> X.eq y m#1 /\ e' = m#2 \/ MapsTo y e' l'
H2:MapsTo y e' l' -> MapsTo y e' (Node ll lx ld lr _x)
H0:X.eq y m#1 -> e' = m#2 -> MapsTo y e' (Node ll lx ld lr _x)
H3:MapsTo y e' (Node l' x d r (max (height l') (height r) + 1))

MapsTo y e' (Node (Node ll lx ld lr _x) x d r h)
inversion_clear H3; intuition. Qed.
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int), bst (Node l x e r h) -> bst (remove_min l x e r)#1
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int), bst (Node l x e r h) -> bst (remove_min l x e r)#1
elt:Type
x:key
d:elt
r:t elt
h:int
H:bst (Node (Leaf elt) x d r h)

bst r
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> bst (remove_min ll lx ld lr)#1
h:int
H:bst (Node (Node ll lx ld lr _x) x d r h)
bst (bal l' x d r)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> bst (remove_min ll lx ld lr)#1
h:int
H:bst (Node (Node ll lx ld lr _x) x d r h)

bst (bal l' x d r)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> bst (remove_min ll lx ld lr)#1
h:int
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
H:bst ll
H4:bst lr
H5:lt_tree lx ll
H6:gt_tree lx lr

bst (bal l' x d r)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> bst (remove_min ll lx ld lr)#1
h:int
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
H:bst ll
H4:bst lr
H5:lt_tree lx ll
H6:gt_tree lx lr

bst l'
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> bst (remove_min ll lx ld lr)#1
h:int
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
H:bst ll
H4:bst lr
H5:lt_tree lx ll
H6:gt_tree lx lr
lt_tree x l'
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> bst (remove_min ll lx ld lr)#1
h:int
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
H:bst ll
H4:bst lr
H5:lt_tree lx ll
H6:gt_tree lx lr

lt_tree x l'
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> bst (remove_min ll lx ld lr)#1
h:int
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
H:bst ll
H4:bst lr
H5:lt_tree lx ll
H6:gt_tree lx lr
y:key
H0:In y l'

X.lt y x
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> bst (remove_min ll lx ld lr)#1
h:int
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
H:bst ll
H4:bst lr
H5:lt_tree lx ll
H6:gt_tree lx lr
y:key
H0:In y l'

In y (Node ll lx ld lr _x) <-> X.eq y ((remove_min ll lx ld lr)#2)#1 \/ In y (remove_min ll lx ld lr)#1 -> X.lt y x
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> bst (remove_min ll lx ld lr)#1
h:int
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
H:bst ll
H4:bst lr
H5:lt_tree lx ll
H6:gt_tree lx lr
y:key
H0:In y l'

In y (Node ll lx ld lr _x) <-> X.eq y m#1 \/ In y l' -> X.lt y x
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> bst (remove_min ll lx ld lr)#1
h:int
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
H:bst ll
H4:bst lr
H5:lt_tree lx ll
H6:gt_tree lx lr
y:key
H0:In y l'
H7:In y (Node ll lx ld lr _x) -> X.eq y m#1 \/ In y l'
H8:X.eq y m#1 \/ In y l' -> In y (Node ll lx ld lr _x)

X.lt y x
apply H2; intuition. Qed. Hint Resolve remove_min_bst : core.
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int), bst (Node l x e r h) -> gt_tree ((remove_min l x e r)#2)#1 (remove_min l x e r)#1
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int), bst (Node l x e r h) -> gt_tree ((remove_min l x e r)#2)#1 (remove_min l x e r)#1
elt:Type
x:key
d:elt
r:t elt
h:int
H:bst (Node (Leaf elt) x d r h)

gt_tree x r
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> gt_tree ((remove_min ll lx ld lr)#2)#1 (remove_min ll lx ld lr)#1
h:int
H:bst (Node (Node ll lx ld lr _x) x d r h)
gt_tree m#1 (bal l' x d r)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> gt_tree ((remove_min ll lx ld lr)#2)#1 (remove_min ll lx ld lr)#1
h:int
H:bst (Node (Node ll lx ld lr _x) x d r h)

gt_tree m#1 (bal l' x d r)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> gt_tree ((remove_min ll lx ld lr)#2)#1 (remove_min ll lx ld lr)#1
h:int
H0:bst (Node ll lx ld lr _x)
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r

gt_tree m#1 (bal l' x d r)
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> gt_tree ((remove_min ll lx ld lr)#2)#1 (remove_min ll lx ld lr)#1
h:int
H0:bst (Node ll lx ld lr _x)
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
y:key
H:In y (bal l' x d r)

X.lt m#1 y
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> gt_tree m#1 l'
h:int
H0:bst (Node ll lx ld lr _x)
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
y:key
H:In y (bal l' x d r)

X.lt m#1 y
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> gt_tree m#1 l'
h:int
H0:bst (Node ll lx ld lr _x)
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
y:key
H:In y (bal l' x d r)

gt_tree m#1 l' -> X.lt m#1 y
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> gt_tree m#1 l'
h:int
H0:bst (Node ll lx ld lr _x)
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
y:key
H:In y (bal l' x d r)

In m#1 (Node ll lx ld lr _x) <-> X.eq m#1 ((remove_min ll lx ld lr)#2)#1 \/ In m#1 (remove_min ll lx ld lr)#1 -> gt_tree m#1 l' -> X.lt m#1 y
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> gt_tree m#1 l'
h:int
H0:bst (Node ll lx ld lr _x)
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
y:key
H:In y (bal l' x d r)
H4:In m#1 (Node ll lx ld lr _x) <-> X.eq m#1 m#1 \/ In m#1 l'
H5:gt_tree m#1 l'

X.lt m#1 y
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> gt_tree m#1 l'
h:int
H0:bst (Node ll lx ld lr _x)
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
y:key
H:X.eq y x \/ In y l' \/ In y r
H4:In m#1 (Node ll lx ld lr _x) <-> X.eq m#1 m#1 \/ In m#1 l'
H5:gt_tree m#1 l'

X.lt m#1 y
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> gt_tree m#1 l'
h:int
H0:bst (Node ll lx ld lr _x)
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
y:key
H:X.eq y x \/ In y l' \/ In y r
H5:gt_tree m#1 l'
H6:In m#1 (Node ll lx ld lr _x)

X.lt m#1 y
elt:Type
x:key
d:elt
r, ll:t elt
lx:key
ld:elt
lr:t elt
_x:int
l':t elt
m:(key * elt)%type
e0:remove_min ll lx ld lr = (l', m)
IHp:forall h0 : int, bst (Node ll lx ld lr h0) -> gt_tree m#1 l'
h:int
H0:bst (Node ll lx ld lr _x)
H1:bst r
H2:lt_tree x (Node ll lx ld lr _x)
H3:gt_tree x r
y:key
H:X.eq y x \/ In y l' \/ In y r
H5:gt_tree m#1 l'
H6:In m#1 (Node ll lx ld lr _x)
H4:X.lt m#1 x

X.lt m#1 y
decompose [or] H; order. Qed. Hint Resolve remove_min_gt_tree : core.
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int) (y : X.t), bst (Node l x e r h) -> find y (Node l x e r h) = match X.compare y ((remove_min l x e r)#2)#1 with | LT _ => None | EQ _ => Some ((remove_min l x e r)#2)#2 | GT _ => find y (remove_min l x e r)#1 end
elt:Type

forall (l : t elt) (x : key) (e : elt) (r : t elt) (h : int) (y : X.t), bst (Node l x e r h) -> find y (Node l x e r h) = match X.compare y ((remove_min l x e r)#2)#1 with | LT _ => None | EQ _ => Some ((remove_min l x e r)#2)#2 | GT _ => find y (remove_min l x e r)#1 end
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)

find y (Node l x e r h) = match X.compare y ((remove_min l x e r)#2)#1 with | LT _ => None | EQ _ => Some ((remove_min l x e r)#2)#2 | GT _ => find y (remove_min l x e r)#1 end
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
l0:X.lt y ((remove_min l x e r)#2)#1

find y (Node l x e r h) = None
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
e0:X.eq y ((remove_min l x e r)#2)#1
find y (Node l x e r h) = Some ((remove_min l x e r)#2)#2
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
l0:X.lt ((remove_min l x e r)#2)#1 y
find y (Node l x e r h) = find y (remove_min l x e r)#1
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
l0:X.lt y ((remove_min l x e r)#2)#1

~ In y (Node l x e r h)
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
e0:X.eq y ((remove_min l x e r)#2)#1
find y (Node l x e r h) = Some ((remove_min l x e r)#2)#2
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
l0:X.lt ((remove_min l x e r)#2)#1 y
find y (Node l x e r h) = find y (remove_min l x e r)#1
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
l0:X.lt y ((remove_min l x e r)#2)#1
H':In y (remove_min l x e r)#1

False
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
e0:X.eq y ((remove_min l x e r)#2)#1
find y (Node l x e r h) = Some ((remove_min l x e r)#2)#2
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
l0:X.lt ((remove_min l x e r)#2)#1 y
find y (Node l x e r h) = find y (remove_min l x e r)#1
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
e0:X.eq y ((remove_min l x e r)#2)#1

find y (Node l x e r h) = Some ((remove_min l x e r)#2)#2
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
l0:X.lt ((remove_min l x e r)#2)#1 y
find y (Node l x e r h) = find y (remove_min l x e r)#1
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
e0:X.eq y ((remove_min l x e r)#2)#1

MapsTo y ((remove_min l x e r)#2)#2 (Node l x e r h)
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
l0:X.lt ((remove_min l x e r)#2)#1 y
find y (Node l x e r h) = find y (remove_min l x e r)#1
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
l0:X.lt ((remove_min l x e r)#2)#1 y

find y (Node l x e r h) = find y (remove_min l x e r)#1
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
y:X.t
H:bst (Node l x e r h)
l0:X.lt ((remove_min l x e r)#2)#1 y
d:elt

MapsTo y d (Node l x e r h) <-> MapsTo y d (remove_min l x e r)#1
rewrite remove_min_mapsto; intuition; order. Qed.

Merging two trees

elt:Type

forall (m1 m2 : t elt) (y : key), bst m1 -> bst m2 -> In y (merge m1 m2) <-> In y m1 \/ In y m2
elt:Type

forall (m1 m2 : t elt) (y : key), bst m1 -> bst m2 -> In y (merge m1 m2) <-> In y m1 \/ In y m2
elt:Type
s2:t elt
y:key
H:bst (Leaf elt)
H0:bst s2

In y s2 <-> In y (Leaf elt) \/ In y s2
elt:Type
y:key
m1:t elt
H:bst m1
H0:bst (Leaf elt)
In y m1 <-> In y m1 \/ In y (Leaf elt)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
y:key
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
In y (bal m1 x d s2') <-> In y m1 \/ In y (Node l2 x2 d2 r2 _x4)
elt:Type
y:key
m1:t elt
H:bst m1
H0:bst (Leaf elt)

In y m1 <-> In y m1 \/ In y (Leaf elt)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
y:key
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
In y (bal m1 x d s2') <-> In y m1 \/ In y (Node l2 x2 d2 r2 _x4)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
y:key
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)

In y (bal m1 x d s2') <-> In y m1 \/ In y (Node l2 x2 d2 r2 _x4)
rewrite bal_in, remove_min_in, e1; simpl; intuition. Qed.
elt:Type

forall (m1 m2 : t elt) (y : key) (e : elt), bst m1 -> bst m2 -> MapsTo y e (merge m1 m2) <-> MapsTo y e m1 \/ MapsTo y e m2
elt:Type

forall (m1 m2 : t elt) (y : key) (e : elt), bst m1 -> bst m2 -> MapsTo y e (merge m1 m2) <-> MapsTo y e m1 \/ MapsTo y e m2
elt:Type
s2:t elt
y:key
e:elt
H:bst (Leaf elt)
H0:bst s2

MapsTo y e s2 <-> MapsTo y e (Leaf elt) \/ MapsTo y e s2
elt:Type
y:key
e:elt
m1:t elt
H:bst m1
H0:bst (Leaf elt)
MapsTo y e m1 <-> MapsTo y e m1 \/ MapsTo y e (Leaf elt)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
y:key
e:elt
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
MapsTo y e (bal m1 x d s2') <-> MapsTo y e m1 \/ MapsTo y e (Node l2 x2 d2 r2 _x4)
elt:Type
y:key
e:elt
m1:t elt
H:bst m1
H0:bst (Leaf elt)

MapsTo y e m1 <-> MapsTo y e m1 \/ MapsTo y e (Leaf elt)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
y:key
e:elt
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
MapsTo y e (bal m1 x d s2') <-> MapsTo y e m1 \/ MapsTo y e (Node l2 x2 d2 r2 _x4)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
y:key
e:elt
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)

MapsTo y e (bal m1 x d s2') <-> MapsTo y e m1 \/ MapsTo y e (Node l2 x2 d2 r2 _x4)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
y:key
e:elt
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)

MapsTo y e (create m1 x d s2') <-> MapsTo y e m1 \/ X.eq y x /\ e = d \/ MapsTo y e s2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
y:key
e:elt
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)

MapsTo y e (Node m1 x d s2' (max (height m1) (height s2') + 1)) <-> MapsTo y e m1 \/ X.eq y x /\ e = d \/ MapsTo y e s2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
y:key
e:elt
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:MapsTo y e (Node m1 x d s2' (max (height m1) (height s2') + 1))

MapsTo y e m1 \/ X.eq y x /\ e = d \/ MapsTo y e s2'
inversion_clear H1; intuition. Qed.
elt:Type

forall m1 m2 : t elt, bst m1 -> bst m2 -> (forall y1 y2 : key, In y1 m1 -> In y2 m2 -> X.lt y1 y2) -> bst (merge m1 m2)
elt:Type

forall m1 m2 : t elt, bst m1 -> bst m2 -> (forall y1 y2 : key, In y1 m1 -> In y2 m2 -> X.lt y1 y2) -> bst (merge m1 m2)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2

bst (bal m1 x d s2')
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2

bst s2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
lt_tree x m1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
gt_tree x s2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2

lt_tree x m1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
gt_tree x s2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
y:key
H2:In y m1

X.lt y x
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
gt_tree x s2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
y:key
H2:In y m1

In x (Node l2 x2 d2 r2 _x4)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
gt_tree x s2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
s2':t elt
x:key
d:elt
e1:remove_min l2 x2 d2 r2 = (s2', (x, d))
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2

gt_tree x s2'
generalize (remove_min_gt_tree H0); rewrite e1; simpl; auto. Qed.

Deletion

elt:Type

forall (m : t elt) (x : X.t) (y : key), bst m -> In y (remove x m) <-> ~ X.eq y x /\ In y m
elt:Type

forall (m : t elt) (x : X.t) (y : key), bst m -> In y (remove x m) <-> ~ X.eq y x /\ In y m
elt:Type
x:X.t
y:key
H:bst (Leaf elt)

In y (Leaf elt) <-> ~ X.eq y x /\ In y (Leaf elt)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
e0:X.compare x y = LT _x0
IHt:forall y1 : key, bst l -> In y1 (remove x l) <-> ~ X.eq y1 x /\ In y1 l
y0:key
H:bst (Node l y d r _x)
In y0 (bal (remove x l) y d r) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
y0:key
H:bst (Node l y d r _x)
In y0 (merge l r) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:forall y1 : key, bst r -> In y1 (remove x r) <-> ~ X.eq y1 x /\ In y1 r
y0:key
H:bst (Node l y d r _x)
In y0 (bal l y d (remove x r)) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
e0:X.compare x y = LT _x0
IHt:forall y1 : key, bst l -> In y1 (remove x l) <-> ~ X.eq y1 x /\ In y1 l
y0:key
H:bst (Node l y d r _x)

In y0 (bal (remove x l) y d r) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
y0:key
H:bst (Node l y d r _x)
In y0 (merge l r) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:forall y1 : key, bst r -> In y1 (remove x r) <-> ~ X.eq y1 x /\ In y1 r
y0:key
H:bst (Node l y d r _x)
In y0 (bal l y d (remove x r)) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
(* LT *)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
IHt:forall y1 : key, bst l -> In y1 (remove x l) <-> ~ X.eq y1 x /\ In y1 l
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

In y0 (bal (remove x l) y d r) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
y0:key
H:bst (Node l y d r _x)
In y0 (merge l r) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:forall y1 : key, bst r -> In y1 (remove x r) <-> ~ X.eq y1 x /\ In y1 r
y0:key
H:bst (Node l y d r _x)
In y0 (bal l y d (remove x r)) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
IHt:forall y1 : key, bst l -> In y1 (remove x l) <-> ~ X.eq y1 x /\ In y1 l
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

X.eq y0 y \/ In y0 (remove x l) \/ In y0 r <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
y0:key
H:bst (Node l y d r _x)
In y0 (merge l r) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:forall y1 : key, bst r -> In y1 (remove x r) <-> ~ X.eq y1 x /\ In y1 r
y0:key
H:bst (Node l y d r _x)
In y0 (bal l y d (remove x r)) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
y0:key
H:bst (Node l y d r _x)

In y0 (merge l r) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:forall y1 : key, bst r -> In y1 (remove x r) <-> ~ X.eq y1 x /\ In y1 r
y0:key
H:bst (Node l y d r _x)
In y0 (bal l y d (remove x r)) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
(* EQ *)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

In y0 (merge l r) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:forall y1 : key, bst r -> In y1 (remove x r) <-> ~ X.eq y1 x /\ In y1 r
y0:key
H:bst (Node l y d r _x)
In y0 (bal l y d (remove x r)) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H4:X.eq y0 x -> False
H:X.eq y0 y

In y0 l \/ In y0 r
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:forall y1 : key, bst r -> In y1 (remove x r) <-> ~ X.eq y1 x /\ In y1 r
y0:key
H:bst (Node l y d r _x)
In y0 (bal l y d (remove x r)) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:forall y1 : key, bst r -> In y1 (remove x r) <-> ~ X.eq y1 x /\ In y1 r
y0:key
H:bst (Node l y d r _x)

In y0 (bal l y d (remove x r)) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
(* GT *)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
IHt:forall y1 : key, bst r -> In y1 (remove x r) <-> ~ X.eq y1 x /\ In y1 r
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

In y0 (bal l y d (remove x r)) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
IHt:forall y1 : key, bst r -> In y1 (remove x r) <-> ~ X.eq y1 x /\ In y1 r
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

X.eq y0 y \/ In y0 l \/ In y0 (remove x r) <-> ~ X.eq y0 x /\ In y0 (Node l y d r _x)
generalize (IHt y0 H1); intuition; [ order | order | intuition_in ]. Qed.
elt:Type

forall (m : t elt) (x : X.t), bst m -> bst (remove x m)
elt:Type

forall (m : t elt) (x : X.t), bst m -> bst (remove x m)
elt:Type
x:X.t
H:bst (Leaf elt)

bst (Leaf elt)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
e0:X.compare x y = LT _x0
IHt:bst l -> bst (remove x l)
H:bst (Node l y d r _x)
bst (bal (remove x l) y d r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
H:bst (Node l y d r _x)
bst (merge l r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H:bst (Node l y d r _x)
bst (bal l y d (remove x r))
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
e0:X.compare x y = LT _x0
IHt:bst l -> bst (remove x l)
H:bst (Node l y d r _x)

bst (bal (remove x l) y d r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
H:bst (Node l y d r _x)
bst (merge l r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H:bst (Node l y d r _x)
bst (bal l y d (remove x r))
(* LT *)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
e0:X.compare x y = LT _x0
IHt:bst l -> bst (remove x l)
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

bst (bal (remove x l) y d r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
H:bst (Node l y d r _x)
bst (merge l r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H:bst (Node l y d r _x)
bst (bal l y d (remove x r))
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
e0:X.compare x y = LT _x0
IHt:bst l -> bst (remove x l)
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

lt_tree y (remove x l)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
H:bst (Node l y d r _x)
bst (merge l r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H:bst (Node l y d r _x)
bst (bal l y d (remove x r))
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
e0:X.compare x y = LT _x0
IHt:bst l -> bst (remove x l)
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
y0:key
H:In y0 (remove x l)

X.lt y0 y
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
H:bst (Node l y d r _x)
bst (merge l r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H:bst (Node l y d r _x)
bst (bal l y d (remove x r))
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
e0:X.compare x y = LT _x0
IHt:bst l -> bst (remove x l)
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
y0:key
H:~ X.eq y0 x /\ In y0 l

X.lt y0 y
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
H:bst (Node l y d r _x)
bst (merge l r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H:bst (Node l y d r _x)
bst (bal l y d (remove x r))
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
H:bst (Node l y d r _x)

bst (merge l r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H:bst (Node l y d r _x)
bst (bal l y d (remove x r))
(* EQ *)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
e0:X.compare x y = EQ _x0
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

bst (merge l r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H:bst (Node l y d r _x)
bst (bal l y d (remove x r))
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H:bst (Node l y d r _x)

bst (bal l y d (remove x r))
(* GT *)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

bst (bal l y d (remove x r))
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

gt_tree y (remove x r)
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
y0:key
H:In y0 (remove x r)

X.lt y y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
e0:X.compare x y = GT _x0
IHt:bst r -> bst (remove x r)
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
y0:key
H:~ X.eq y0 x /\ In y0 r

X.lt y y0
destruct H; eauto. Qed.
elt:Type

forall (m : t elt) (x y : X.t), bst m -> X.eq x y -> ~ In y (remove x m)
elt:Type

forall (m : t elt) (x y : X.t), bst m -> X.eq x y -> ~ In y (remove x m)
intros; rewrite remove_in; intuition. Qed.
elt:Type

forall (m : t elt) (x y : X.t) (e : elt), bst m -> ~ X.eq x y -> MapsTo y e m -> MapsTo y e (remove x m)
elt:Type

forall (m : t elt) (x y : X.t) (e : elt), bst m -> ~ X.eq x y -> MapsTo y e m -> MapsTo y e (remove x m)
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x, y:X.t
e:elt
IHm1:bst m1 -> ~ X.eq x y -> MapsTo y e m1 -> MapsTo y e (remove x m1)
IHm2:bst m2 -> ~ X.eq x y -> MapsTo y e m2 -> MapsTo y e (remove x m2)

bst (Node m1 k e0 m2 t) -> ~ X.eq x y -> MapsTo y e (Node m1 k e0 m2 t) -> MapsTo y e match X.compare x k with | LT _ => bal (remove x m1) k e0 m2 | EQ _ => merge m1 m2 | GT _ => bal m1 k e0 (remove x m2) end
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x, y:X.t
e:elt
IHm1:bst m1 -> ~ X.eq x y -> MapsTo y e m1 -> MapsTo y e (remove x m1)
IHm2:bst m2 -> ~ X.eq x y -> MapsTo y e m2 -> MapsTo y e (remove x m2)
e1:X.eq x k
H0:~ X.eq x y
H1:MapsTo y e (Node m1 k e0 m2 t)
H2:bst m1
H3:bst m2
H4:lt_tree k m1
H5:gt_tree k m2

MapsTo y e (merge m1 m2)
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x, y:X.t
e:elt
IHm1:bst m1 -> ~ X.eq x y -> MapsTo y e m1 -> MapsTo y e (remove x m1)
IHm2:bst m2 -> ~ X.eq x y -> MapsTo y e m2 -> MapsTo y e (remove x m2)
e1:X.eq x k
H0:~ X.eq x y
H1:MapsTo y e (Node m1 k e0 m2 t)
H2:bst m1
H3:bst m2
H4:lt_tree k m1
H5:gt_tree k m2

MapsTo y e m1 \/ MapsTo y e m2
inv MapsTo; auto; order. Qed.
elt:Type

forall (m : t elt) (x : X.t) (y : key) (e : elt), bst m -> MapsTo y e (remove x m) -> MapsTo y e m
elt:Type

forall (m : t elt) (x : X.t) (y : key) (e : elt), bst m -> MapsTo y e (remove x m) -> MapsTo y e m
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x:X.t
y:key
e:elt
IHm1:bst m1 -> MapsTo y e (remove x m1) -> MapsTo y e m1
IHm2:bst m2 -> MapsTo y e (remove x m2) -> MapsTo y e m2

bst (Node m1 k e0 m2 t) -> MapsTo y e match X.compare x k with | LT _ => bal (remove x m1) k e0 m2 | EQ _ => merge m1 m2 | GT _ => bal m1 k e0 (remove x m2) end -> MapsTo y e (Node m1 k e0 m2 t)
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x:X.t
y:key
e:elt
IHm1:bst m1 -> MapsTo y e (remove x m1) -> MapsTo y e m1
IHm2:bst m2 -> MapsTo y e (remove x m2) -> MapsTo y e m2
l:X.lt x k
H:bst m1
H0:bst m2
H1:lt_tree k m1
H2:gt_tree k m2

MapsTo y e (Node (remove x m1) k e0 m2 (max (height (remove x m1)) (height m2) + 1)) -> MapsTo y e (Node m1 k e0 m2 t)
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x:X.t
y:key
e:elt
IHm1:bst m1 -> MapsTo y e (remove x m1) -> MapsTo y e m1
IHm2:bst m2 -> MapsTo y e (remove x m2) -> MapsTo y e m2
e1:X.eq x k
H:bst m1
H0:bst m2
H1:lt_tree k m1
H2:gt_tree k m2
MapsTo y e (merge m1 m2) -> MapsTo y e (Node m1 k e0 m2 t)
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x:X.t
y:key
e:elt
IHm1:bst m1 -> MapsTo y e (remove x m1) -> MapsTo y e m1
IHm2:bst m2 -> MapsTo y e (remove x m2) -> MapsTo y e m2
l:X.lt k x
H:bst m1
H0:bst m2
H1:lt_tree k m1
H2:gt_tree k m2
MapsTo y e (Node m1 k e0 (remove x m2) (max (height m1) (height (remove x m2)) + 1)) -> MapsTo y e (Node m1 k e0 m2 t)
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x:X.t
y:key
e:elt
IHm1:bst m1 -> MapsTo y e (remove x m1) -> MapsTo y e m1
IHm2:bst m2 -> MapsTo y e (remove x m2) -> MapsTo y e m2
e1:X.eq x k
H:bst m1
H0:bst m2
H1:lt_tree k m1
H2:gt_tree k m2

MapsTo y e (merge m1 m2) -> MapsTo y e (Node m1 k e0 m2 t)
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x:X.t
y:key
e:elt
IHm1:bst m1 -> MapsTo y e (remove x m1) -> MapsTo y e m1
IHm2:bst m2 -> MapsTo y e (remove x m2) -> MapsTo y e m2
l:X.lt k x
H:bst m1
H0:bst m2
H1:lt_tree k m1
H2:gt_tree k m2
MapsTo y e (Node m1 k e0 (remove x m2) (max (height m1) (height (remove x m2)) + 1)) -> MapsTo y e (Node m1 k e0 m2 t)
elt:Type
m1:Raw.t elt
k:key
e0:elt
m2:Raw.t elt
t:int
x:X.t
y:key
e:elt
IHm1:bst m1 -> MapsTo y e (remove x m1) -> MapsTo y e m1
IHm2:bst m2 -> MapsTo y e (remove x m2) -> MapsTo y e m2
l:X.lt k x
H:bst m1
H0:bst m2
H1:lt_tree k m1
H2:gt_tree k m2

MapsTo y e (Node m1 k e0 (remove x m2) (max (height m1) (height (remove x m2)) + 1)) -> MapsTo y e (Node m1 k e0 m2 t)
intros; inv MapsTo; auto. Qed.

join

elt:Type

forall (l : t elt) (x : key) (d : elt) (r : t elt) (y : key), In y (join l x d r) <-> X.eq y x \/ In y l \/ In y r
elt:Type

forall (l : t elt) (x : key) (d : elt) (r : t elt) (y : key), In y (join l x d r) <-> X.eq y x \/ In y l \/ In y r
elt:Type
x:key
d:elt
r:t elt
y:key

In y (join (Leaf elt) x d r) <-> X.eq y x \/ In y (Leaf elt) \/ In y r
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
y:key
In y (add x d (Node ll lx ld lr lh)) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Leaf elt)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
GT:lh > rh + 2
y:key
In y (bal ll lx ld (join lr x d (Node rl rx rd rr rh))) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
LE:lh <= rh + 2
GT':rh > lh + 2
y:key
In y (bal (join (Node ll lx ld lr lh) x d rl) rx rd rr) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
LE:lh <= rh + 2
LE':rh <= lh + 2
y:key
In y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh)) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
x:key
d:elt
r:t elt
y:key

In y (add x d r) <-> X.eq y x \/ In y (Leaf elt) \/ In y r
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
y:key
In y (add x d (Node ll lx ld lr lh)) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Leaf elt)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
GT:lh > rh + 2
y:key
In y (bal ll lx ld (join lr x d (Node rl rx rd rr rh))) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
LE:lh <= rh + 2
GT':rh > lh + 2
y:key
In y (bal (join (Node ll lx ld lr lh) x d rl) rx rd rr) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
LE:lh <= rh + 2
LE':rh <= lh + 2
y:key
In y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh)) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
y:key

In y (add x d (Node ll lx ld lr lh)) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Leaf elt)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
GT:lh > rh + 2
y:key
In y (bal ll lx ld (join lr x d (Node rl rx rd rr rh))) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
LE:lh <= rh + 2
GT':rh > lh + 2
y:key
In y (bal (join (Node ll lx ld lr lh) x d rl) rx rd rr) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
LE:lh <= rh + 2
LE':rh <= lh + 2
y:key
In y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh)) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
GT:lh > rh + 2
y:key

In y (bal ll lx ld (join lr x d (Node rl rx rd rr rh))) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
LE:lh <= rh + 2
GT':rh > lh + 2
y:key
In y (bal (join (Node ll lx ld lr lh) x d rl) rx rd rr) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
LE:lh <= rh + 2
LE':rh <= lh + 2
y:key
In y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh)) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
LE:lh <= rh + 2
GT':rh > lh + 2
y:key

In y (bal (join (Node ll lx ld lr lh) x d rl) rx rd rr) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
LE:lh <= rh + 2
LE':rh <= lh + 2
y:key
In y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh)) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : key) (d0 : elt) (r : t elt) (y0 : key), In y0 (join lr x0 d0 r) <-> X.eq y0 x0 \/ In y0 lr \/ In y0 r
x:key
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : key, In y0 (join (Node ll lx ld lr lh) x d rl) <-> X.eq y0 x \/ In y0 (Node ll lx ld lr lh) \/ In y0 rl
LE:lh <= rh + 2
LE':rh <= lh + 2
y:key

In y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh)) <-> X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y (Node rl rx rd rr rh)
apply create_in. Qed.
elt:Type

forall (l : t elt) (x : X.t) (d : elt) (r : t elt), bst l -> bst r -> lt_tree x l -> gt_tree x r -> bst (join l x d r)
elt:Type

forall (l : t elt) (x : X.t) (d : elt) (r : t elt), bst l -> bst r -> lt_tree x l -> gt_tree x r -> bst (join l x d r)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
GT:lh > rh + 2
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
y:key
H:X.eq y x \/ In y lr \/ In y (Node rl rx rd rr rh)

X.lt lx y
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
LE:lh <= rh + 2
GT':rh > lh + 2
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
y:key
H:X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y rl
X.lt y rx
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
LE:lh <= rh + 2
GT':rh > lh + 2
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
y:key
H:X.eq y x \/ In y (Node ll lx ld lr lh) \/ In y rl

X.lt y rx
intuition; [ apply MX.eq_lt with x | ]; eauto. Qed. Hint Resolve join_bst : core.
elt:Type

forall (l : t elt) (x : X.t) (d : elt) (r : t elt) (y : X.t), bst l -> bst r -> lt_tree x l -> gt_tree x r -> find y (join l x d r) = find y (create l x d r)
elt:Type

forall (l : t elt) (x : X.t) (d : elt) (r : t elt) (y : X.t), bst l -> bst r -> lt_tree x l -> gt_tree x r -> find y (join l x d r) = find y (create l x d r)
elt:Type
x:X.t
d:elt
r:t elt
y:X.t
H0:bst r
H1:lt_tree x (Leaf elt)
H2:gt_tree x r

match X.compare y x with | EQ _ => Some d | _ => find y r end = find y (create (Leaf elt) x d r)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Leaf elt)
H0:bst ll
H3:bst lr
H4:lt_tree lx ll
H5:gt_tree lx lr
H:X.lt lx x
match X.compare y x with | EQ _ => Some d | _ => find y (Node ll lx ld lr lh) end = find y (create (Node ll lx ld lr lh) x d (Leaf elt))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create ll lx ld (join lr x d (Node rl rx rd rr rh))) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
gt_tree lx (join lr x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create (join (Node ll lx ld lr lh) x d rl) rx rd rr) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
x:X.t
d:elt
r:t elt
y:X.t
H0:bst r
H1:lt_tree x (Leaf elt)
H2:gt_tree x r
l:X.lt y x

find y r = None
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Leaf elt)
H0:bst ll
H3:bst lr
H4:lt_tree lx ll
H5:gt_tree lx lr
H:X.lt lx x
match X.compare y x with | EQ _ => Some d | _ => find y (Node ll lx ld lr lh) end = find y (create (Node ll lx ld lr lh) x d (Leaf elt))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create ll lx ld (join lr x d (Node rl rx rd rr rh))) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
gt_tree lx (join lr x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create (join (Node ll lx ld lr lh) x d rl) rx rd rr) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Leaf elt)
H0:bst ll
H3:bst lr
H4:lt_tree lx ll
H5:gt_tree lx lr
H:X.lt lx x

match X.compare y x with | EQ _ => Some d | _ => find y (Node ll lx ld lr lh) end = find y (create (Node ll lx ld lr lh) x d (Leaf elt))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create ll lx ld (join lr x d (Node rl rx rd rr rh))) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
gt_tree lx (join lr x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create (join (Node ll lx ld lr lh) x d rl) rx rd rr) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Leaf elt)
H0:bst ll
H3:bst lr
H4:lt_tree lx ll
H5:gt_tree lx lr
H:X.lt lx x
l:X.lt x y
l0:X.lt lx y

find y lr = None
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create ll lx ld (join lr x d (Node rl rx rd rr rh))) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
gt_tree lx (join lr x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create (join (Node ll lx ld lr lh) x d rl) rx rd rr) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Leaf elt)
H0:bst ll
H3:bst lr
H4:lt_tree lx ll
H5:gt_tree lx lr
H:X.lt lx x
l:X.lt x y
l0:X.lt lx y
H6:In y lr

False
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create ll lx ld (join lr x d (Node rl rx rd rr rh))) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
gt_tree lx (join lr x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create (join (Node ll lx ld lr lh) x d rl) rx rd rr) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx

find y (create ll lx ld (join lr x d (Node rl rx rd rr rh))) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
gt_tree lx (join lr x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create (join (Node ll lx ld lr lh) x d rl) rx rd rr) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx

match X.compare y lx with | LT _ => find y ll | EQ _ => Some ld | GT _ => match X.compare y x with | LT _ => find y lr | EQ _ => Some d | GT _ => match X.compare y rx with | LT _ => find y rl | EQ _ => Some rd | GT _ => find y rr end end end = match X.compare y x with | LT _ => match X.compare y lx with | LT _ => find y ll | EQ _ => Some ld | GT _ => find y lr end | EQ _ => Some d | GT _ => match X.compare y rx with | LT _ => find y rl | EQ _ => Some rd | GT _ => find y rr end end
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
gt_tree lx (join lr x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create (join (Node ll lx ld lr lh) x d rl) rx rd rr) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx

gt_tree lx (join lr x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create (join (Node ll lx ld lr lh) x d rl) rx rd rr) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
GT:lh > rh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
u:key
Hu:X.eq u x \/ In u lr \/ In u (Node rl rx rd rr rh)

X.lt lx u
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
find y (create (join (Node ll lx ld lr lh) x d rl) rx rd rr) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx

find y (create (join (Node ll lx ld lr lh) x d rl) rx rd rr) = find y (create (Node ll lx ld lr lh) x d (Node rl rx rd rr rh))
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx

match X.compare y rx with | LT _ => match X.compare y x with | LT _ => match X.compare y lx with | LT _ => find y ll | EQ _ => Some ld | GT _ => find y lr end | EQ _ => Some d | GT _ => find y rl end | EQ _ => Some rd | GT _ => find y rr end = match X.compare y x with | LT _ => match X.compare y lx with | LT _ => find y ll | EQ _ => Some ld | GT _ => find y lr end | EQ _ => Some d | GT _ => match X.compare y rx with | LT _ => find y rl | EQ _ => Some rd | GT _ => find y rr end end
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx

lt_tree rx (join (Node ll lx ld lr lh) x d rl)
elt:Type
ll:t elt
lx:key
ld:elt
lr:t elt
lh:int
Hlr:forall (x0 : X.t) (d0 : elt) (r : t elt) (y0 : X.t), bst lr -> bst r -> lt_tree x0 lr -> gt_tree x0 r -> find y0 (join lr x0 d0 r) = find y0 (create lr x0 d0 r)
x:X.t
d:elt
rl:t elt
rx:key
rd:elt
rr:t elt
rh:int
Hrl:forall y0 : X.t, bst (Node ll lx ld lr lh) -> bst rl -> lt_tree x (Node ll lx ld lr lh) -> gt_tree x rl -> find y0 (join (Node ll lx ld lr lh) x d rl) = find y0 (create (Node ll lx ld lr lh) x d rl)
LE:lh <= rh + 2
GT':rh > lh + 2
y:X.t
H1:lt_tree x (Node ll lx ld lr lh)
H2:gt_tree x (Node rl rx rd rr rh)
H3:bst rl
H4:bst rr
H5:lt_tree rx rl
H6:gt_tree rx rr
H0:bst ll
H7:bst lr
H8:lt_tree lx ll
H9:gt_tree lx lr
H:X.lt lx x
H10:X.lt x rx
u:key
Hu:X.eq u x \/ In u (Node ll lx ld lr lh) \/ In u rl

X.lt u rx
destruct Hu as [Hu|[Hu|Hu]]; order. Qed.

split

elt:Type

forall (m : t elt) (x : X.t), bst m -> forall y : key, In y (split x m)#l <-> In y m /\ X.lt y x
elt:Type

forall (m : t elt) (x : X.t), bst m -> forall y : key, In y (split x m)#l <-> In y m /\ X.lt y x
elt:Type
x:X.t
y:key

In y (Leaf elt) <-> In y (Leaf elt) /\ X.lt y x
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
ll:t elt
o:option elt
rl:t elt
e1:split x l = << ll, o, rl >>
IHt:bst l -> forall y1 : key, In y1 (split x l)#l <-> In y1 l /\ X.lt y1 x
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 ll <-> In y0 (Node l y d r _x) /\ X.lt y0 x
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 l <-> In y0 (Node l y d r _x) /\ X.lt y0 x
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:bst r -> forall y1 : key, In y1 (split x r)#l <-> In y1 r /\ X.lt y1 x
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 (join l y d rl) <-> In y0 (Node l y d r _x) /\ X.lt y0 x
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
ll:t elt
o:option elt
rl:t elt
e1:split x l = << ll, o, rl >>
IHt:bst l -> forall y1 : key, In y1 (split x l)#l <-> In y1 l /\ X.lt y1 x
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

In y0 ll <-> In y0 (Node l y d r _x) /\ X.lt y0 x
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 l <-> In y0 (Node l y d r _x) /\ X.lt y0 x
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:bst r -> forall y1 : key, In y1 (split x r)#l <-> In y1 r /\ X.lt y1 x
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 (join l y d rl) <-> In y0 (Node l y d r _x) /\ X.lt y0 x
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

In y0 l <-> In y0 (Node l y d r _x) /\ X.lt y0 x
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:bst r -> forall y1 : key, In y1 (split x r)#l <-> In y1 r /\ X.lt y1 x
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 (join l y d rl) <-> In y0 (Node l y d r _x) /\ X.lt y0 x
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:bst r -> forall y1 : key, In y1 (split x r)#l <-> In y1 r /\ X.lt y1 x
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

In y0 (join l y d rl) <-> In y0 (Node l y d r _x) /\ X.lt y0 x
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:bst r -> forall y1 : key, In y1 (split x r)#l <-> In y1 r /\ X.lt y1 x
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

X.eq y0 y \/ In y0 l \/ In y0 rl <-> In y0 (Node l y d r _x) /\ X.lt y0 x
rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order. Qed.
elt:Type

forall (m : t elt) (x : X.t), bst m -> forall y : key, In y (split x m)#r <-> In y m /\ X.lt x y
elt:Type

forall (m : t elt) (x : X.t), bst m -> forall y : key, In y (split x m)#r <-> In y m /\ X.lt x y
elt:Type
x:X.t
y:key

In y (Leaf elt) <-> In y (Leaf elt) /\ X.lt x y
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
ll:t elt
o:option elt
rl:t elt
e1:split x l = << ll, o, rl >>
IHt:bst l -> forall y1 : key, In y1 (split x l)#r <-> In y1 l /\ X.lt x y1
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 (join rl y d r) <-> In y0 (Node l y d r _x) /\ X.lt x y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 r <-> In y0 (Node l y d r _x) /\ X.lt x y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:bst r -> forall y1 : key, In y1 (split x r)#r <-> In y1 r /\ X.lt x y1
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 rr <-> In y0 (Node l y d r _x) /\ X.lt x y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
ll:t elt
o:option elt
rl:t elt
e1:split x l = << ll, o, rl >>
IHt:bst l -> forall y1 : key, In y1 (split x l)#r <-> In y1 l /\ X.lt x y1
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

In y0 (join rl y d r) <-> In y0 (Node l y d r _x) /\ X.lt x y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 r <-> In y0 (Node l y d r _x) /\ X.lt x y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:bst r -> forall y1 : key, In y1 (split x r)#r <-> In y1 r /\ X.lt x y1
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 rr <-> In y0 (Node l y d r _x) /\ X.lt x y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
ll:t elt
o:option elt
rl:t elt
e1:split x l = << ll, o, rl >>
IHt:bst l -> forall y1 : key, In y1 (split x l)#r <-> In y1 l /\ X.lt x y1
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

X.eq y0 y \/ In y0 rl \/ In y0 r <-> In y0 (Node l y d r _x) /\ X.lt x y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 r <-> In y0 (Node l y d r _x) /\ X.lt x y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:bst r -> forall y1 : key, In y1 (split x r)#r <-> In y1 r /\ X.lt x y1
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 rr <-> In y0 (Node l y d r _x) /\ X.lt x y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

In y0 r <-> In y0 (Node l y d r _x) /\ X.lt x y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:bst r -> forall y1 : key, In y1 (split x r)#r <-> In y1 r /\ X.lt x y1
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
In y0 rr <-> In y0 (Node l y d r _x) /\ X.lt x y0
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:bst r -> forall y1 : key, In y1 (split x r)#r <-> In y1 r /\ X.lt x y1
y0:key
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

In y0 rr <-> In y0 (Node l y d r _x) /\ X.lt x y0
rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order. Qed.
elt:Type

forall (m : t elt) (x : X.t), bst m -> (split x m)#o = find x m
elt:Type

forall (m : t elt) (x : X.t), bst m -> (split x m)#o = find x m
intros m x; functional induction (split x m); subst; simpl; auto; intros; inv bst; try clear e0; destruct X.compare; try order; trivial; rewrite <- IHt, e1; auto. Qed.
elt:Type

forall (m : t elt) (x : X.t), bst m -> bst (split x m)#l /\ bst (split x m)#r
elt:Type

forall (m : t elt) (x : X.t), bst m -> bst (split x m)#l /\ bst (split x m)#r
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
ll:t elt
o:option elt
rl:t elt
e1:split x l = << ll, o, rl >>
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H4:bst ll
H5:bst rl

lt_tree y rl
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H4:bst rl
H5:bst rr
gt_tree y rl
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
ll:t elt
o:option elt
rl:t elt
e1:split x l = << ll, o, rl >>
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H4:bst ll
H5:bst rl
y0:key

In y0 rl -> X.lt y0 y
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H4:bst rl
H5:bst rr
gt_tree y rl
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H4:bst rl
H5:bst rr

gt_tree y rl
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H4:bst rl
H5:bst rr
y0:key

In y0 rl -> X.lt y y0
generalize (split_in_1 x H1 y0); rewrite e1; simpl; intuition. Qed.
elt:Type

forall (m : t elt) (x : X.t), bst m -> lt_tree x (split x m)#l
elt:Type

forall (m : t elt) (x : X.t), bst m -> lt_tree x (split x m)#l
intros m x B y Hy; rewrite split_in_1 in Hy; intuition. Qed.
elt:Type

forall (m : t elt) (x : X.t), bst m -> gt_tree x (split x m)#r
elt:Type

forall (m : t elt) (x : X.t), bst m -> gt_tree x (split x m)#r
intros m x B y Hy; rewrite split_in_2 in Hy; intuition. Qed.
elt:Type

forall (m : t elt) (x y : X.t), bst m -> find y m = match X.compare y x with | LT _ => find y (split x m)#l | EQ _ => (split x m)#o | GT _ => find y (split x m)#r end
elt:Type

forall (m : t elt) (x y : X.t), bst m -> find y m = match X.compare y x with | LT _ => find y (split x m)#l | EQ _ => (split x m)#o | GT _ => find y (split x m)#r end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
ll:t elt
o:option elt
rl:t elt
e1:split x l = << ll, o, rl >>
IHt:forall y1 : X.t, bst l -> find y1 l = match X.compare y1 x with | LT _ => find y1 ll | EQ _ => o | GT _ => find y1 rl end
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 ll <-> In y1 l /\ X.lt y1 x
H4:forall y1 : key, In y1 rl <-> In y1 l /\ X.lt x y1
H5:bst ll
H6:bst rl

match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end = match X.compare y0 x with | LT _ => find y0 ll | EQ _ => o | GT _ => find y0 (join rl y d r) end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end = match X.compare y0 x with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:forall y1 : X.t, bst r -> find y1 r = match X.compare y1 x with | LT _ => find y1 rl | EQ _ => o | GT _ => find y1 rr end
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 rl <-> In y1 r /\ X.lt y1 x
H4:forall y1 : key, In y1 rr <-> In y1 r /\ X.lt x y1
H5:bst rl
H6:bst rr
match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end = match X.compare y0 x with | LT _ => find y0 (join l y d rl) | EQ _ => o | GT _ => find y0 rr end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
ll:t elt
o:option elt
rl:t elt
e1:split x l = << ll, o, rl >>
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 ll <-> In y1 l /\ X.lt y1 x
H4:forall y1 : key, In y1 rl <-> In y1 l /\ X.lt x y1
H5:bst ll
H6:bst rl

match X.compare y0 y with | LT _ => match X.compare y0 x with | LT _ => find y0 ll | EQ _ => o | GT _ => find y0 rl end | EQ _ => Some d | GT _ => find y0 r end = match X.compare y0 x with | LT _ => find y0 ll | EQ _ => o | GT _ => match X.compare y0 y with | LT _ => find y0 rl | EQ _ => Some d | GT _ => find y0 r end end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
ll:t elt
o:option elt
rl:t elt
e1:split x l = << ll, o, rl >>
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 ll <-> In y1 l /\ X.lt y1 x
H4:forall y1 : key, In y1 rl <-> In y1 l /\ X.lt x y1
H5:bst ll
H6:bst rl
lt_tree y rl
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end = match X.compare y0 x with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:forall y1 : X.t, bst r -> find y1 r = match X.compare y1 x with | LT _ => find y1 rl | EQ _ => o | GT _ => find y1 rr end
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 rl <-> In y1 r /\ X.lt y1 x
H4:forall y1 : key, In y1 rr <-> In y1 r /\ X.lt x y1
H5:bst rl
H6:bst rr
match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end = match X.compare y0 x with | LT _ => find y0 (join l y d rl) | EQ _ => o | GT _ => find y0 rr end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt x y
ll:t elt
o:option elt
rl:t elt
e1:split x l = << ll, o, rl >>
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 ll <-> In y1 l /\ X.lt y1 x
H4:forall y1 : key, In y1 rl <-> In y1 l /\ X.lt x y1
H5:bst ll
H6:bst rl

lt_tree y rl
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end = match X.compare y0 x with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:forall y1 : X.t, bst r -> find y1 r = match X.compare y1 x with | LT _ => find y1 rl | EQ _ => o | GT _ => find y1 rr end
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 rl <-> In y1 r /\ X.lt y1 x
H4:forall y1 : key, In y1 rr <-> In y1 r /\ X.lt x y1
H5:bst rl
H6:bst rr
match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end = match X.compare y0 x with | LT _ => find y0 (join l y d rl) | EQ _ => o | GT _ => find y0 rr end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.eq x y
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r

match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end = match X.compare y0 x with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:forall y1 : X.t, bst r -> find y1 r = match X.compare y1 x with | LT _ => find y1 rl | EQ _ => o | GT _ => find y1 rr end
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 rl <-> In y1 r /\ X.lt y1 x
H4:forall y1 : key, In y1 rr <-> In y1 r /\ X.lt x y1
H5:bst rl
H6:bst rr
match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end = match X.compare y0 x with | LT _ => find y0 (join l y d rl) | EQ _ => o | GT _ => find y0 rr end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
IHt:forall y1 : X.t, bst r -> find y1 r = match X.compare y1 x with | LT _ => find y1 rl | EQ _ => o | GT _ => find y1 rr end
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 rl <-> In y1 r /\ X.lt y1 x
H4:forall y1 : key, In y1 rr <-> In y1 r /\ X.lt x y1
H5:bst rl
H6:bst rr

match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 r end = match X.compare y0 x with | LT _ => find y0 (join l y d rl) | EQ _ => o | GT _ => find y0 rr end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 rl <-> In y1 r /\ X.lt y1 x
H4:forall y1 : key, In y1 rr <-> In y1 r /\ X.lt x y1
H5:bst rl
H6:bst rr

match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => match X.compare y0 x with | LT _ => find y0 rl | EQ _ => o | GT _ => find y0 rr end end = match X.compare y0 x with | LT _ => match X.compare y0 y with | LT _ => find y0 l | EQ _ => Some d | GT _ => find y0 rl end | EQ _ => o | GT _ => find y0 rr end
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 rl <-> In y1 r /\ X.lt y1 x
H4:forall y1 : key, In y1 rr <-> In y1 r /\ X.lt x y1
H5:bst rl
H6:bst rr
gt_tree y rl
elt:Type
x:X.t
l:t elt
y:key
d:elt
r:t elt
_x:int
_x0:X.lt y x
rl:t elt
o:option elt
rr:t elt
e1:split x r = << rl, o, rr >>
y0:X.t
H0:bst l
H1:bst r
H2:lt_tree y l
H3:gt_tree y r
H:forall y1 : key, In y1 rl <-> In y1 r /\ X.lt y1 x
H4:forall y1 : key, In y1 rr <-> In y1 r /\ X.lt x y1
H5:bst rl
H6:bst rr

gt_tree y rl
intros y1; rewrite H; intuition. Qed.

Concatenation

elt:Type

forall (m1 m2 : t elt) (y : key), In y (concat m1 m2) <-> In y m1 \/ In y m2
elt:Type

forall (m1 m2 : t elt) (y : key), In y (concat m1 m2) <-> In y m1 \/ In y m2
elt:Type
m2:t elt
y:key

In y m2 <-> In y (Leaf elt) \/ In y m2
elt:Type
y:key
m1:t elt
In y m1 <-> In y m1 \/ In y (Leaf elt)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:key
m1:t elt
In y (join m1 xd#1 xd#2 m2') <-> In y m1 \/ In y (Node l2 x2 d2 r2 _x4)
elt:Type
y:key
m1:t elt

In y m1 <-> In y m1 \/ In y (Leaf elt)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:key
m1:t elt
In y (join m1 xd#1 xd#2 m2') <-> In y m1 \/ In y (Node l2 x2 d2 r2 _x4)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:key
m1:t elt

In y (join m1 xd#1 xd#2 m2') <-> In y m1 \/ In y (Node l2 x2 d2 r2 _x4)
rewrite join_in, remove_min_in, e1; simpl; intuition. Qed.
elt:Type

forall m1 m2 : t elt, bst m1 -> bst m2 -> (forall y1 y2 : key, In y1 m1 -> In y2 m2 -> X.lt y1 y2) -> bst (concat m1 m2)
elt:Type

forall m1 m2 : t elt, bst m1 -> bst m2 -> (forall y1 y2 : key, In y1 m1 -> In y2 m2 -> X.lt y1 y2) -> bst (concat m1 m2)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2

bst (join m1 xd#1 xd#2 m2')
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2

bst m2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
lt_tree xd#1 m1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
gt_tree xd#1 m2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2

bst (m2', xd)#1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
lt_tree xd#1 m1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
gt_tree xd#1 m2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2

lt_tree xd#1 m1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
gt_tree xd#1 m2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
y:key
Hy:In y m1

X.lt y xd#1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
gt_tree xd#1 m2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
y:key
Hy:In y m1

In xd#1 (Node l2 x2 d2 r2 _x4)
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
gt_tree xd#1 m2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2

gt_tree xd#1 m2'
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2

gt_tree ((m2', xd)#2)#1 (m2', xd)#1
rewrite <-e1; eauto. Qed. Hint Resolve concat_bst : core.
elt:Type

forall (m1 m2 : t elt) (y : X.t), bst m1 -> bst m2 -> (forall y1 y2 : key, In y1 m1 -> In y2 m2 -> X.lt y1 y2) -> find y (concat m1 m2) = match find y m2 with | Some d => Some d | None => find y m1 end
elt:Type

forall (m1 m2 : t elt) (y : X.t), bst m1 -> bst m2 -> (forall y1 y2 : key, In y1 m1 -> In y2 m2 -> X.lt y1 y2) -> find y (concat m1 m2) = match find y m2 with | Some d => Some d | None => find y m1 end
elt:Type
m2:t elt
y:X.t
H:bst (Leaf elt)
H0:bst m2
H1:forall y1 y2 : key, In y1 (Leaf elt) -> In y2 m2 -> X.lt y1 y2

find y m2 = match find y m2 with | Some d => Some d | None => find y (Leaf elt) end
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
find y (join m1 xd#1 xd#2 m2') = match find y (Node l2 x2 d2 r2 _x4) with | Some d => Some d | None => find y m1 end
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2

find y (join m1 xd#1 xd#2 m2') = match find y (Node l2 x2 d2 r2 _x4) with | Some d => Some d | None => find y m1 end
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H0:bst (Node l2 x2 d2 r2 _x4)
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
H2:find y (Node l2 x2 d2 r2 _x4) = match X.compare y xd#1 with | LT _ => None | EQ _ => Some xd#2 | GT _ => find y m2' end
H3:forall y0 : key, In y0 (Node l2 x2 d2 r2 _x4) <-> X.eq y0 xd#1 \/ In y0 m2'
H4:bst m2'
H5:gt_tree xd#1 m2'

find y (join m1 xd#1 xd#2 m2') = match find y (Node l2 x2 d2 r2 _x4) with | Some d => Some d | None => find y m1 end
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
H2:find y (Node l2 x2 d2 r2 _x4) = match X.compare y xd#1 with | LT _ => None | EQ _ => Some xd#2 | GT _ => find y m2' end
H3:forall y0 : key, In y0 (Node l2 x2 d2 r2 _x4) <-> X.eq y0 xd#1 \/ In y0 m2'
H4:bst m2'
H5:gt_tree xd#1 m2'
H6:bst l2
H7:bst r2
H8:lt_tree x2 l2
H9:gt_tree x2 r2

find y (join m1 xd#1 xd#2 m2') = match find y (Node l2 x2 d2 r2 _x4) with | Some d => Some d | None => find y m1 end
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
H3:forall y0 : key, In y0 (Node l2 x2 d2 r2 _x4) <-> X.eq y0 xd#1 \/ In y0 m2'
H4:bst m2'
H5:gt_tree xd#1 m2'
H6:bst l2
H7:bst r2
H8:lt_tree x2 l2
H9:gt_tree x2 r2

find y (create m1 xd#1 xd#2 m2') = match match X.compare y xd#1 with | LT _ => None | EQ _ => Some xd#2 | GT _ => find y m2' end with | Some d => Some d | None => find y m1 end
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
H3:forall y0 : key, In y0 (Node l2 x2 d2 r2 _x4) <-> X.eq y0 xd#1 \/ In y0 m2'
H4:bst m2'
H5:gt_tree xd#1 m2'
H6:bst l2
H7:bst r2
H8:lt_tree x2 l2
H9:gt_tree x2 r2
lt_tree xd#1 m1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
H3:forall y0 : key, In y0 (Node l2 x2 d2 r2 _x4) <-> X.eq y0 xd#1 \/ In y0 m2'
H4:bst m2'
H5:gt_tree xd#1 m2'
H6:bst l2
H7:bst r2
H8:lt_tree x2 l2
H9:gt_tree x2 r2
Hlt:X.lt xd#1 y

find y m2' = match find y m2' with | Some d => Some d | None => find y m1 end
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
H3:forall y0 : key, In y0 (Node l2 x2 d2 r2 _x4) <-> X.eq y0 xd#1 \/ In y0 m2'
H4:bst m2'
H5:gt_tree xd#1 m2'
H6:bst l2
H7:bst r2
H8:lt_tree x2 l2
H9:gt_tree x2 r2
lt_tree xd#1 m1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
H3:forall y0 : key, In y0 (Node l2 x2 d2 r2 _x4) <-> X.eq y0 xd#1 \/ In y0 m2'
H4:bst m2'
H5:gt_tree xd#1 m2'
H6:bst l2
H7:bst r2
H8:lt_tree x2 l2
H9:gt_tree x2 r2
Hlt:X.lt xd#1 y

None = find y m1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
H3:forall y0 : key, In y0 (Node l2 x2 d2 r2 _x4) <-> X.eq y0 xd#1 \/ In y0 m2'
H4:bst m2'
H5:gt_tree xd#1 m2'
H6:bst l2
H7:bst r2
H8:lt_tree x2 l2
H9:gt_tree x2 r2
lt_tree xd#1 m1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
H3:forall y0 : key, In y0 (Node l2 x2 d2 r2 _x4) <-> X.eq y0 xd#1 \/ In y0 m2'
H4:bst m2'
H5:gt_tree xd#1 m2'
H6:bst l2
H7:bst r2
H8:lt_tree x2 l2
H9:gt_tree x2 r2
Hlt:X.lt xd#1 y
H0:In y m1

False
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
H3:forall y0 : key, In y0 (Node l2 x2 d2 r2 _x4) <-> X.eq y0 xd#1 \/ In y0 m2'
H4:bst m2'
H5:gt_tree xd#1 m2'
H6:bst l2
H7:bst r2
H8:lt_tree x2 l2
H9:gt_tree x2 r2
lt_tree xd#1 m1
elt:Type
l2:t elt
x2:key
d2:elt
r2:t elt
_x4:int
m2':t elt
xd:(key * elt)%type
e1:remove_min l2 x2 d2 r2 = (m2', xd)
y:X.t
m1:t elt
H:bst m1
H1:forall y1 y2 : key, In y1 m1 -> In y2 (Node l2 x2 d2 r2 _x4) -> X.lt y1 y2
H3:forall y0 : key, In y0 (Node l2 x2 d2 r2 _x4) <-> X.eq y0 xd#1 \/ In y0 m2'
H4:bst m2'
H5:gt_tree xd#1 m2'
H6:bst l2
H7:bst r2
H8:lt_tree x2 l2
H9:gt_tree x2 r2

lt_tree xd#1 m1
intros z Hz; apply H1; auto; rewrite H3; auto. Qed.

Elements

Notation eqk := (PX.eqk (elt:= elt)).
Notation eqke := (PX.eqke (elt:= elt)).
Notation ltk := (PX.ltk (elt:= elt)).

elt:Type

forall (s : t elt) (acc : list (key * elt)) (x : X.t) (e : elt), InA eqke (x, e) (elements_aux acc s) <-> MapsTo x e s \/ InA eqke (x, e) acc
elt:Type

forall (s : t elt) (acc : list (key * elt)) (x : X.t) (e : elt), InA eqke (x, e) (elements_aux acc s) <-> MapsTo x e s \/ InA eqke (x, e) acc
elt:Type

forall (acc : list (key * elt)) (x : X.t) (e : elt), InA eqke (x, e) acc <-> MapsTo x e (Leaf elt) \/ InA eqke (x, e) acc
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
Hl:forall (acc : list (key * elt)) (x0 : X.t) (e0 : elt), InA eqke (x0, e0) (elements_aux acc l) <-> MapsTo x0 e0 l \/ InA eqke (x0, e0) acc
Hr:forall (acc : list (key * elt)) (x0 : X.t) (e0 : elt), InA eqke (x0, e0) (elements_aux acc r) <-> MapsTo x0 e0 r \/ InA eqke (x0, e0) acc
forall (acc : list (key * elt)) (x0 : X.t) (e0 : elt), InA eqke (x0, e0) (elements_aux ((x, e) :: elements_aux acc r) l) <-> MapsTo x0 e0 (Node l x e r h) \/ InA eqke (x0, e0) acc
elt:Type
acc:list (key * elt)
x:X.t
e:elt
H0:MapsTo x e (Leaf elt)

InA eqke (x, e) acc
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
Hl:forall (acc : list (key * elt)) (x0 : X.t) (e0 : elt), InA eqke (x0, e0) (elements_aux acc l) <-> MapsTo x0 e0 l \/ InA eqke (x0, e0) acc
Hr:forall (acc : list (key * elt)) (x0 : X.t) (e0 : elt), InA eqke (x0, e0) (elements_aux acc r) <-> MapsTo x0 e0 r \/ InA eqke (x0, e0) acc
forall (acc : list (key * elt)) (x0 : X.t) (e0 : elt), InA eqke (x0, e0) (elements_aux ((x, e) :: elements_aux acc r) l) <-> MapsTo x0 e0 (Node l x e r h) \/ InA eqke (x0, e0) acc
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
Hl:forall (acc : list (key * elt)) (x0 : X.t) (e0 : elt), InA eqke (x0, e0) (elements_aux acc l) <-> MapsTo x0 e0 l \/ InA eqke (x0, e0) acc
Hr:forall (acc : list (key * elt)) (x0 : X.t) (e0 : elt), InA eqke (x0, e0) (elements_aux acc r) <-> MapsTo x0 e0 r \/ InA eqke (x0, e0) acc

forall (acc : list (key * elt)) (x0 : X.t) (e0 : elt), InA eqke (x0, e0) (elements_aux ((x, e) :: elements_aux acc r) l) <-> MapsTo x0 e0 (Node l x e r h) \/ InA eqke (x0, e0) acc
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
Hl:forall (acc0 : list (key * elt)) (x1 : X.t) (e1 : elt), InA eqke (x1, e1) (elements_aux acc0 l) <-> MapsTo x1 e1 l \/ InA eqke (x1, e1) acc0
Hr:forall (acc0 : list (key * elt)) (x1 : X.t) (e1 : elt), InA eqke (x1, e1) (elements_aux acc0 r) <-> MapsTo x1 e1 r \/ InA eqke (x1, e1) acc0
acc:list (key * elt)
x0:X.t
e0:elt

InA eqke (x0, e0) (elements_aux ((x, e) :: elements_aux acc r) l) <-> MapsTo x0 e0 (Node l x e r h) \/ InA eqke (x0, e0) acc
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
Hl:forall (acc0 : list (key * elt)) (x1 : X.t) (e1 : elt), InA eqke (x1, e1) (elements_aux acc0 l) <-> MapsTo x1 e1 l \/ InA eqke (x1, e1) acc0
Hr:forall (acc0 : list (key * elt)) (x1 : X.t) (e1 : elt), InA eqke (x1, e1) (elements_aux acc0 r) <-> MapsTo x1 e1 r \/ InA eqke (x1, e1) acc0
acc:list (key * elt)
x0:X.t
e0:elt

MapsTo x0 e0 l \/ InA eqke (x0, e0) ((x, e) :: elements_aux acc r) <-> MapsTo x0 e0 (Node l x e r h) \/ InA eqke (x0, e0) acc
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
acc:list (key * elt)
x0:X.t
e0:elt
H:InA eqke (x0, e0) (elements_aux acc r) -> MapsTo x0 e0 r \/ InA eqke (x0, e0) acc
H0:MapsTo x0 e0 r \/ InA eqke (x0, e0) acc -> InA eqke (x0, e0) (elements_aux acc r)

MapsTo x0 e0 l \/ InA eqke (x0, e0) ((x, e) :: elements_aux acc r) <-> MapsTo x0 e0 (Node l x e r h) \/ InA eqke (x0, e0) acc
elt:Type
l:t elt
x:key
e:elt
r:t elt
h:int
acc:list (key * elt)
x0:X.t
e0:elt
H:InA eqke (x0, e0) (elements_aux acc r) -> MapsTo x0 e0 r \/ InA eqke (x0, e0) acc
H1:MapsTo x0 e0 r -> InA eqke (x0, e0) (elements_aux acc r)
H2:InA eqke (x0, e0) acc -> InA eqke (x0, e0) (elements_aux acc r)
H0:eqke (x0, e0) (x, e)

MapsTo x0 e0 (Node l x e r h) \/ InA eqke (x0, e0) acc
destruct H0; simpl in *; subst; intuition. Qed.
elt:Type

forall (s : t elt) (x : X.t) (e : elt), InA eqke (x, e) (elements s) <-> MapsTo x e s
elt:Type

forall (s : t elt) (x : X.t) (e : elt), InA eqke (x, e) (elements s) <-> MapsTo x e s
elt:Type
s:t elt
x:X.t
e:elt
H:MapsTo x e s -> InA eqke (x, e) (elements_aux nil s)
H1:InA eqke (x, e) (elements s)
H0:InA eqke (x, e) nil
H3:InA eqke (x, e) (elements_aux nil s)

MapsTo x e s
inversion_clear H0. Qed.
elt:Type

forall (s : t elt) (x : X.t), L.PX.In x (elements s) <-> In x s
elt:Type

forall (s : t elt) (x : X.t), L.PX.In x (elements s) <-> In x s
elt:Type
s:t elt
x:X.t

L.PX.In x (elements s) <-> In x s
elt:Type
s:t elt
x:X.t

(exists e : elt, L.PX.MapsTo x e (elements s)) <-> In x s
elt:Type
s:t elt
x:X.t

(exists e : elt, L.PX.MapsTo x e (elements s)) <-> (exists e : elt, MapsTo x e s)
elt:Type
s:t elt
x:X.t
x0:elt
H:L.PX.MapsTo x x0 (elements s)

exists e : elt, MapsTo x e s
elt:Type
s:t elt
x:X.t
x0:elt
H:MapsTo x x0 s
exists e : elt, L.PX.MapsTo x e (elements s)
elt:Type
s:t elt
x:X.t
x0:elt
H:L.PX.MapsTo x x0 (elements s)

MapsTo x x0 s
elt:Type
s:t elt
x:X.t
x0:elt
H:MapsTo x x0 s
exists e : elt, L.PX.MapsTo x e (elements s)
elt:Type
s:t elt
x:X.t
x0:elt
H:MapsTo x x0 s

exists e : elt, L.PX.MapsTo x e (elements s)
elt:Type
s:t elt
x:X.t
x0:elt
H:MapsTo x x0 s

L.PX.MapsTo x x0 (elements s)
unfold L.PX.MapsTo; rewrite elements_mapsto; auto. Qed.
elt:Type

forall (s : t elt) (acc : list (X.t * elt)), bst s -> Sorted ltk acc -> (forall (x : X.t) (e : elt) (y : key), InA eqke (x, e) acc -> In y s -> X.lt y x) -> Sorted ltk (elements_aux acc s)
elt:Type

forall (s : t elt) (acc : list (X.t * elt)), bst s -> Sorted ltk acc -> (forall (x : X.t) (e : elt) (y : key), InA eqke (x, e) acc -> In y s -> X.lt y x) -> Sorted ltk (elements_aux acc s)
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H:bst (Node l y e r h)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x

Sorted ltk (elements_aux ((y, e) :: elements_aux acc r) l)
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r

Sorted ltk (elements_aux ((y, e) :: elements_aux acc r) l)
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r

Sorted ltk ((y, e) :: elements_aux acc r)
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) ((y, e) :: elements_aux acc r) -> In y0 l -> X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r

Sorted ltk (elements_aux acc r)
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
HdRel ltk (y, e) (elements_aux acc r)
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) ((y, e) :: elements_aux acc r) -> In y0 l -> X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r

HdRel ltk (y, e) (elements_aux acc r)
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) ((y, e) :: elements_aux acc r) -> In y0 l -> X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r

forall y0 : X.t * elt, InA eqke y0 (elements_aux acc r) -> ltk (y, e) y0
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) ((y, e) :: elements_aux acc r) -> In y0 l -> X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
y':X.t
e':elt
H6:InA eqke (y', e') (elements_aux acc r)

ltk (y, e) (y', e')
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) ((y, e) :: elements_aux acc r) -> In y0 l -> X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
y':X.t
e':elt
H6:InA eqke (y', e') (elements_aux acc r)
H:MapsTo y' e' r
H9:InA eqke (y', e') acc -> InA eqke (y', e') (elements_aux acc r)
H7:InA eqke (y', e') (elements_aux acc r)

ltk (y, e) (y', e')
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
y':X.t
e':elt
H6:InA eqke (y', e') (elements_aux acc r)
H:InA eqke (y', e') acc
H8:MapsTo y' e' r -> InA eqke (y', e') (elements_aux acc r)
H7:InA eqke (y', e') (elements_aux acc r)
ltk (y, e) (y', e')
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) ((y, e) :: elements_aux acc r) -> In y0 l -> X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
y':X.t
e':elt
H6:InA eqke (y', e') (elements_aux acc r)
H:InA eqke (y', e') acc
H8:MapsTo y' e' r -> InA eqke (y', e') (elements_aux acc r)
H7:InA eqke (y', e') (elements_aux acc r)

ltk (y, e) (y', e')
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) ((y, e) :: elements_aux acc r) -> In y0 l -> X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 l -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc0 -> In y0 r -> X.lt y0 x) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) acc -> In y0 (Node l y e r h) -> X.lt y0 x
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r

forall (x : X.t) (e0 : elt) (y0 : key), InA eqke (x, e0) ((y, e) :: elements_aux acc r) -> In y0 l -> X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 l -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 r -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc -> In y1 (Node l y e r h) -> X.lt y1 x0
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
x:X.t
e0:elt
y0:key
H:InA eqke (x, e0) ((y, e) :: elements_aux acc r)
H6:In y0 l

X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 l -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 r -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc -> In y1 (Node l y e r h) -> X.lt y1 x0
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
x:X.t
e0:elt
y0:key
H6:In y0 l
H7:eqke (x, e0) (y, e)

X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 l -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 r -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc -> In y1 (Node l y e r h) -> X.lt y1 x0
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
x:X.t
e0:elt
y0:key
H6:In y0 l
H7:InA eqke (x, e0) (elements_aux acc r)
X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 l -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 r -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc -> In y1 (Node l y e r h) -> X.lt y1 x0
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
x:X.t
e0:elt
y0:key
H6:In y0 l
H:X.eq x y
H7:e0 = e

X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 l -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 r -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc -> In y1 (Node l y e r h) -> X.lt y1 x0
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
x:X.t
e0:elt
y0:key
H6:In y0 l
H7:InA eqke (x, e0) (elements_aux acc r)
X.lt y0 x
elt:Type
l:t elt
y:key
e:elt
r:t elt
h:int
Hl:forall acc0 : list (X.t * elt), bst l -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 l -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 l)
Hr:forall acc0 : list (X.t * elt), bst r -> Sorted ltk acc0 -> (forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc0 -> In y1 r -> X.lt y1 x0) -> Sorted ltk (elements_aux acc0 r)
acc:list (X.t * elt)
H0:Sorted ltk acc
H1:forall (x0 : X.t) (e1 : elt) (y1 : key), InA eqke (x0, e1) acc -> In y1 (Node l y e r h) -> X.lt y1 x0
H2:bst l
H3:bst r
H4:lt_tree y l
H5:gt_tree y r
x:X.t
e0:elt
y0:key
H6:In y0 l
H7:InA eqke (x, e0) (elements_aux acc r)

X.lt y0 x
destruct (elements_aux_mapsto r acc x e0); intuition eauto. Qed.
elt:Type

forall s : t elt, bst s -> Sorted ltk (elements s)
elt:Type

forall s : t elt, bst s -> Sorted ltk (elements s)
elt:Type
s:t elt
H:bst s

forall (x : X.t) (e : elt) (y : key), InA eqke (x, e) nil -> In y s -> X.lt y x
intros; inversion H0. Qed. Hint Resolve elements_sort : core.
elt:Type

forall s : t elt, bst s -> NoDupA eqk (elements s)
elt:Type

forall s : t elt, bst s -> NoDupA eqk (elements s)
intros; apply PX.Sort_NoDupA; auto. Qed.
elt:Type

forall (m : t elt) (acc : list (key * elt)), (length acc + cardinal m)%nat = length (elements_aux acc m)
elt:Type

forall (m : t elt) (acc : list (key * elt)), (length acc + cardinal m)%nat = length (elements_aux acc m)
elt:Type
m, t:Raw.t elt
H:forall acc0 : list (key * elt), (length acc0 + cardinal t)%nat = length (elements_aux acc0 t)
k:key
e:elt
t0:Raw.t elt
H0:forall acc0 : list (key * elt), (length acc0 + cardinal t0)%nat = length (elements_aux acc0 t0)
t1:int
acc:list (key * elt)

(length acc + S (cardinal t + cardinal t0))%nat = length (elements_aux ((k, e) :: elements_aux acc t0) t)
elt:Type
m, t:Raw.t elt
H:forall acc0 : list (key * elt), (length acc0 + cardinal t)%nat = length (elements_aux acc0 t)
k:key
e:elt
t0:Raw.t elt
H0:forall acc0 : list (key * elt), (length acc0 + cardinal t0)%nat = length (elements_aux acc0 t0)
t1:int
acc:list (key * elt)

(length acc + S (cardinal t + cardinal t0))%nat = S (length (elements_aux acc t0) + cardinal t)
rewrite <- H0; omega. Qed.
elt:Type

forall m : t elt, cardinal m = length (elements m)
elt:Type

forall m : t elt, cardinal m = length (elements m)
exact (fun m => elements_aux_cardinal m nil). Qed.
elt:Type

forall (s : t elt) (acc : list (key * elt)), elements_aux acc s = elements s ++ acc
elt:Type

forall (s : t elt) (acc : list (key * elt)), elements_aux acc s = elements s ++ acc
elt:Type
s1:Raw.t elt
k:key
e:elt
s2:Raw.t elt
t:int
IHs1:forall acc0 : list (key * elt), elements_aux acc0 s1 = elements s1 ++ acc0
IHs2:forall acc0 : list (key * elt), elements_aux acc0 s2 = elements s2 ++ acc0
acc:list (key * elt)

elements_aux ((k, e) :: elements_aux acc s2) s1 = elements (Node s1 k e s2 t) ++ acc
elt:Type
s1:Raw.t elt
k:key
e:elt
s2:Raw.t elt
t:int
IHs1:forall acc0 : list (key * elt), elements_aux acc0 s1 = elements s1 ++ acc0
IHs2:forall acc0 : list (key * elt), elements_aux acc0 s2 = elements s2 ++ acc0
acc:list (key * elt)

elements s1 ++ (k, e) :: elements s2 ++ acc = elements (Node s1 k e s2 t) ++ acc
elt:Type
s1:Raw.t elt
k:key
e:elt
s2:Raw.t elt
t:int
IHs1:forall acc0 : list (key * elt), elements_aux acc0 s1 = elements s1 ++ acc0
IHs2:forall acc0 : list (key * elt), elements_aux acc0 s2 = elements s2 ++ acc0
acc:list (key * elt)

elements_aux nil s1 ++ (k, e) :: elements_aux nil s2 ++ acc = elements_aux ((k, e) :: elements_aux nil s2) s1 ++ acc
rewrite 2 IHs1, IHs2, !app_nil_r, !app_ass; auto. Qed.
elt:Type

forall (t1 t2 : t elt) (x : key) (e : elt) (z : int) (l : list (key * elt)), elements t1 ++ (x, e) :: elements t2 ++ l = elements (Node t1 x e t2 z) ++ l
elt:Type

forall (t1 t2 : t elt) (x : key) (e : elt) (z : int) (l : list (key * elt)), elements t1 ++ (x, e) :: elements t2 ++ l = elements (Node t1 x e t2 z) ++ l
elt:Type
t1, t2:t elt
x:key
e:elt
z:int
l:list (key * elt)

elements_aux nil t1 ++ (x, e) :: elements_aux nil t2 ++ l = elements_aux ((x, e) :: elements_aux nil t2) t1 ++ l
rewrite !elements_app, !app_nil_r, !app_ass; auto. Qed.

Fold

Definition fold' (A : Type) (f : key -> elt -> A -> A)(s : t elt) :=
  L.fold f (elements s).

elt:Type

forall (A : Type) (s : t elt) (f : key -> elt -> A -> A) (a : A) (acc : list (key * elt)), L.fold f (elements_aux acc s) a = L.fold f acc (fold f s a)
elt:Type

forall (A : Type) (s : t elt) (f : key -> elt -> A -> A) (a : A) (acc : list (key * elt)), L.fold f (elements_aux acc s) a = L.fold f acc (fold f s a)
elt, A:Type
s:t elt

forall (f : key -> elt -> A -> A) (a : A) (acc : list (key * elt)), L.fold f (elements_aux acc (Leaf elt)) a = L.fold f acc (fold f (Leaf elt) a)
elt, A:Type
s:t elt
forall t : t elt, (forall (f : key -> elt -> A -> A) (a : A) (acc : list (key * elt)), L.fold f (elements_aux acc t) a = L.fold f acc (fold f t a)) -> forall (k : key) (e : elt) (t0 : Raw.t elt), (forall (f : key -> elt -> A -> A) (a : A) (acc : list (key * elt)), L.fold f (elements_aux acc t0) a = L.fold f acc (fold f t0 a)) -> forall (t1 : int) (f : key -> elt -> A -> A) (a : A) (acc : list (key * elt)), L.fold f (elements_aux acc (Node t k e t0 t1)) a = L.fold f acc (fold f (Node t k e t0 t1) a)
elt, A:Type
s:t elt

forall t : t elt, (forall (f : key -> elt -> A -> A) (a : A) (acc : list (key * elt)), L.fold f (elements_aux acc t) a = L.fold f acc (fold f t a)) -> forall (k : key) (e : elt) (t0 : Raw.t elt), (forall (f : key -> elt -> A -> A) (a : A) (acc : list (key * elt)), L.fold f (elements_aux acc t0) a = L.fold f acc (fold f t0 a)) -> forall (t1 : int) (f : key -> elt -> A -> A) (a : A) (acc : list (key * elt)), L.fold f (elements_aux acc (Node t k e t0 t1)) a = L.fold f acc (fold f (Node t k e t0 t1) a)
elt, A:Type
s, t:Raw.t elt
H:forall (f0 : key -> elt -> A -> A) (a0 : A) (acc0 : list (key * elt)), L.fold f0 (elements_aux acc0 t) a0 = L.fold f0 acc0 (fold f0 t a0)
k:key
e:elt
t0:Raw.t elt
H0:forall (f0 : key -> elt -> A -> A) (a0 : A) (acc0 : list (key * elt)), L.fold f0 (elements_aux acc0 t0) a0 = L.fold f0 acc0 (fold f0 t0 a0)
t1:int
f:key -> elt -> A -> A
a:A
acc:list (key * elt)

L.fold f (elements_aux ((k, e) :: elements_aux acc t0) t) a = L.fold f acc (fold f t0 (f k e (fold f t a)))
elt, A:Type
s, t:Raw.t elt
H:forall (f0 : key -> elt -> A -> A) (a0 : A) (acc0 : list (key * elt)), L.fold f0 (elements_aux acc0 t) a0 = L.fold f0 acc0 (fold f0 t a0)
k:key
e:elt
t0:Raw.t elt
H0:forall (f0 : key -> elt -> A -> A) (a0 : A) (acc0 : list (key * elt)), L.fold f0 (elements_aux acc0 t0) a0 = L.fold f0 acc0 (fold f0 t0 a0)
t1:int
f:key -> elt -> A -> A
a:A
acc:list (key * elt)

L.fold f ((k, e) :: elements_aux acc t0) (fold f t a) = L.fold f acc (fold f t0 (f k e (fold f t a)))
elt, A:Type
s, t:Raw.t elt
H:forall (f0 : key -> elt -> A -> A) (a0 : A) (acc0 : list (key * elt)), L.fold f0 (elements_aux acc0 t) a0 = L.fold f0 acc0 (fold f0 t a0)
k:key
e:elt
t0:Raw.t elt
H0:forall (f0 : key -> elt -> A -> A) (a0 : A) (acc0 : list (key * elt)), L.fold f0 (elements_aux acc0 t0) a0 = L.fold f0 acc0 (fold f0 t0 a0)
t1:int
f:key -> elt -> A -> A
a:A
acc:list (key * elt)

L.fold f (elements_aux acc t0) (f k e (fold f t a)) = L.fold f acc (fold f t0 (f k e (fold f t a)))
apply H0. Qed.
elt:Type

forall (A : Type) (s : t elt) (f : key -> elt -> A -> A) (a : A), fold f s a = fold' f s a
elt:Type

forall (A : Type) (s : t elt) (f : key -> elt -> A -> A) (a : A), fold f s a = fold' f s a
elt:Type

forall (A : Type) (s : t elt) (f : key -> elt -> A -> A) (a : A), fold f s a = L.fold f (elements_aux nil s) a
elt, A:Type
s, t:Raw.t elt
H:forall (f0 : key -> elt -> A -> A) (a0 : A), fold f0 t a0 = L.fold f0 (elements_aux nil t) a0
k:key
e:elt
t0:Raw.t elt
H0:forall (f0 : key -> elt -> A -> A) (a0 : A), fold f0 t0 a0 = L.fold f0 (elements_aux nil t0) a0
t1:int
f:key -> elt -> A -> A
a:A

fold f t0 (f k e (fold f t a)) = L.fold f (elements_aux ((k, e) :: elements_aux nil t0) t) a
elt, A:Type
s, t:Raw.t elt
H:forall (f0 : key -> elt -> A -> A) (a0 : A), fold f0 t a0 = L.fold f0 (elements_aux nil t) a0
k:key
e:elt
t0:Raw.t elt
H0:forall (f0 : key -> elt -> A -> A) (a0 : A), fold f0 t0 a0 = L.fold f0 (elements_aux nil t0) a0
t1:int
f:key -> elt -> A -> A
a:A

fold f t0 (f k e (fold f t a)) = L.fold f ((k, e) :: elements_aux nil t0) (fold f t a)
elt, A:Type
s, t:Raw.t elt
H:forall (f0 : key -> elt -> A -> A) (a0 : A), fold f0 t a0 = L.fold f0 (elements_aux nil t) a0
k:key
e:elt
t0:Raw.t elt
H0:forall (f0 : key -> elt -> A -> A) (a0 : A), fold f0 t0 a0 = L.fold f0 (elements_aux nil t0) a0
t1:int
f:key -> elt -> A -> A
a:A

L.fold f (elements_aux nil t0) (f k e (fold f t a)) = L.fold f ((k, e) :: elements_aux nil t0) (fold f t a)
simpl; auto. Qed.
elt:Type

forall s : t elt, bst s -> forall (A : Type) (i : A) (f : key -> elt -> A -> A), fold f s i = fold_left (fun (a : A) (p : key * elt) => f p#1 p#2 a) (elements s) i
elt:Type

forall s : t elt, bst s -> forall (A : Type) (i : A) (f : key -> elt -> A -> A), fold f s i = fold_left (fun (a : A) (p : key * elt) => f p#1 p#2 a) (elements s) i
elt:Type
s:t elt
Hs:bst s
A:Type
i:A
f:key -> elt -> A -> A

fold f s i = fold_left (fun (a : A) (p : key * elt) => f p#1 p#2 a) (elements s) i
elt:Type
s:t elt
Hs:bst s
A:Type
i:A
f:key -> elt -> A -> A

fold' f s i = fold_left (fun (a : A) (p : key * elt) => f p#1 p#2 a) (elements s) i
elt:Type
s:t elt
Hs:bst s
A:Type
i:A
f:key -> elt -> A -> A

L.fold f (elements s) i = fold_left (fun (a : A) (p : key * elt) => f p#1 p#2 a) (elements s) i
elt:Type
s:t elt
Hs:bst s
A:Type
i:A
f:key -> elt -> A -> A

fold_left (fun (a : A) (p : L.key * elt) => f p#1 p#2 a) (L.elements (elt:=elt) (elements s)) i = fold_left (fun (a : A) (p : key * elt) => f p#1 p#2 a) (elements s) i
unfold L.elements; auto. Qed.

Comparison

flatten_e e returns the list of elements of the enumeration e i.e. the list of elements actually compared
Fixpoint flatten_e (e : enumeration elt) : list (key*elt) := match e with
  | End _ => nil
  | More x e t r => (x,e) :: elements t ++ flatten_e r
 end.

elt:Type

forall (l r : t elt) (x : key) (d : elt) (z : int) (e : enumeration elt), elements l ++ flatten_e (More x d r e) = elements (Node l x d r z) ++ flatten_e e
elt:Type

forall (l r : t elt) (x : key) (d : elt) (z : int) (e : enumeration elt), elements l ++ flatten_e (More x d r e) = elements (Node l x d r z) ++ flatten_e e
intros; apply elements_node. Qed.
elt:Type

forall (s : t elt) (e : enumeration elt), flatten_e (cons s e) = elements s ++ flatten_e e
elt:Type

forall (s : t elt) (e : enumeration elt), flatten_e (cons s e) = elements s ++ flatten_e e
elt:Type
s1:Raw.t elt
k:key
e:elt
s2:Raw.t elt
t:int
IHs1:forall e1 : enumeration elt, flatten_e (cons s1 e1) = elements s1 ++ flatten_e e1
IHs2:forall e1 : enumeration elt, flatten_e (cons s2 e1) = elements s2 ++ flatten_e e1
e0:enumeration elt

flatten_e (cons (Node s1 k e s2 t) e0) = elements (Node s1 k e s2 t) ++ flatten_e e0
simpl flatten_e; rewrite IHs1; apply flatten_e_elements; auto. Qed.
Proof of correction for the comparison
Variable cmp : elt->elt->bool.

Definition IfEq b l1 l2 := L.equal cmp l1 l2 = b.

elt:Type
cmp:elt -> elt -> bool

forall (b : bool) (x1 x2 : X.t) (d1 d2 : elt) (l1 l2 : L.t elt), X.eq x1 x2 -> cmp d1 d2 = true -> IfEq b l1 l2 -> IfEq b ((x1, d1) :: l1) ((x2, d2) :: l2)
elt:Type
cmp:elt -> elt -> bool

forall (b : bool) (x1 x2 : X.t) (d1 d2 : elt) (l1 l2 : L.t elt), X.eq x1 x2 -> cmp d1 d2 = true -> IfEq b l1 l2 -> IfEq b ((x1, d1) :: l1) ((x2, d2) :: l2)
unfold IfEq; destruct b; simpl; intros; destruct X.compare; simpl; try rewrite H0; auto; order. Qed.
elt:Type
cmp:elt -> elt -> bool

forall e2 : enumeration elt, IfEq (equal_end e2) nil (flatten_e e2)
elt:Type
cmp:elt -> elt -> bool

forall e2 : enumeration elt, IfEq (equal_end e2) nil (flatten_e e2)
destruct e2; red; auto. Qed.
elt:Type
cmp:elt -> elt -> bool

forall (x1 : X.t) (d1 : elt) (cont : enumeration elt -> bool) (x2 : key) (d2 : elt) (r2 : t elt) (e2 : enumeration elt) (l : L.t elt), IfEq (cont (cons r2 e2)) l (elements r2 ++ flatten_e e2) -> IfEq (equal_more cmp x1 d1 cont (More x2 d2 r2 e2)) ((x1, d1) :: l) (flatten_e (More x2 d2 r2 e2))
elt:Type
cmp:elt -> elt -> bool

forall (x1 : X.t) (d1 : elt) (cont : enumeration elt -> bool) (x2 : key) (d2 : elt) (r2 : t elt) (e2 : enumeration elt) (l : L.t elt), IfEq (cont (cons r2 e2)) l (elements r2 ++ flatten_e e2) -> IfEq (equal_more cmp x1 d1 cont (More x2 d2 r2 e2)) ((x1, d1) :: l) (flatten_e (More x2 d2 r2 e2))
elt:Type
cmp:elt -> elt -> bool
x1:X.t
d1:elt
cont:enumeration elt -> bool
x2:key
d2:elt
r2:t elt
e2:enumeration elt
l:L.t elt
H:L.equal cmp l (elements r2 ++ flatten_e e2) = cont (cons r2 e2)
e:X.eq x1 x2

cmp d1 d2 && L.equal cmp l (elements r2 ++ flatten_e e2) = cmp d1 d2 &&& cont (cons r2 e2)
rewrite <-andb_lazy_alt; f_equal; auto. Qed.
elt:Type
cmp:elt -> elt -> bool

forall (m1 : t elt) (cont : enumeration elt -> bool) (e2 : enumeration elt) (l : L.t elt), (forall e : enumeration elt, IfEq (cont e) l (flatten_e e)) -> IfEq (equal_cont cmp m1 cont e2) (elements m1 ++ l) (flatten_e e2)
elt:Type
cmp:elt -> elt -> bool

forall (m1 : t elt) (cont : enumeration elt -> bool) (e2 : enumeration elt) (l : L.t elt), (forall e : enumeration elt, IfEq (cont e) l (flatten_e e)) -> IfEq (equal_cont cmp m1 cont e2) (elements m1 ++ l) (flatten_e e2)
elt:Type
cmp:elt -> elt -> bool
l1:t elt
x1:key
d1:elt
r1:t elt
h1:int
Hl1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp l1 cont0 e0) (elements l1 ++ l0) (flatten_e e0)
Hr1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp r1 cont0 e0) (elements r1 ++ l0) (flatten_e e0)
cont:enumeration elt -> bool
e2:enumeration elt
l:L.t elt
H:forall e : enumeration elt, IfEq (cont e) l (flatten_e e)

IfEq (equal_cont cmp (Node l1 x1 d1 r1 h1) cont e2) (elements (Node l1 x1 d1 r1 h1) ++ l) (flatten_e e2)
elt:Type
cmp:elt -> elt -> bool
l1:t elt
x1:key
d1:elt
r1:t elt
h1:int
Hl1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp l1 cont0 e0) (elements l1 ++ l0) (flatten_e e0)
Hr1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp r1 cont0 e0) (elements r1 ++ l0) (flatten_e e0)
cont:enumeration elt -> bool
e2:enumeration elt
l:L.t elt
H:forall e : enumeration elt, IfEq (cont e) l (flatten_e e)

IfEq (equal_cont cmp l1 (equal_more cmp x1 d1 (equal_cont cmp r1 cont)) e2) (elements l1 ++ (x1, d1) :: elements r1 ++ l) (flatten_e e2)
elt:Type
cmp:elt -> elt -> bool
l1:t elt
x1:key
d1:elt
r1:t elt
h1:int
Hl1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp l1 cont0 e0) (elements l1 ++ l0) (flatten_e e0)
Hr1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp r1 cont0 e0) (elements r1 ++ l0) (flatten_e e0)
cont:enumeration elt -> bool
e2:enumeration elt
l:L.t elt
H:forall e : enumeration elt, IfEq (cont e) l (flatten_e e)

forall e : enumeration elt, IfEq (equal_more cmp x1 d1 (equal_cont cmp r1 cont) e) ((x1, d1) :: elements r1 ++ l) (flatten_e e)
elt:Type
cmp:elt -> elt -> bool
l1:t elt
x1:key
d1:elt
r1:t elt
h1:int
Hl1:forall (cont0 : enumeration elt -> bool) (e2 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp l1 cont0 e2) (elements l1 ++ l0) (flatten_e e2)
Hr1:forall (cont0 : enumeration elt -> bool) (e2 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp r1 cont0 e2) (elements r1 ++ l0) (flatten_e e2)
cont:enumeration elt -> bool
l:L.t elt
H:forall e : enumeration elt, IfEq (cont e) l (flatten_e e)

IfEq (equal_more cmp x1 d1 (equal_cont cmp r1 cont) (End elt)) ((x1, d1) :: elements r1 ++ l) (flatten_e (End elt))
elt:Type
cmp:elt -> elt -> bool
l1:t elt
x1:key
d1:elt
r1:t elt
h1:int
Hl1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp l1 cont0 e0) (elements l1 ++ l0) (flatten_e e0)
Hr1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp r1 cont0 e0) (elements r1 ++ l0) (flatten_e e0)
cont:enumeration elt -> bool
l:L.t elt
H:forall e : enumeration elt, IfEq (cont e) l (flatten_e e)
x2:key
d2:elt
r2:t elt
e2:enumeration elt
IfEq (equal_more cmp x1 d1 (equal_cont cmp r1 cont) (More x2 d2 r2 e2)) ((x1, d1) :: elements r1 ++ l) (flatten_e (More x2 d2 r2 e2))
elt:Type
cmp:elt -> elt -> bool
l1:t elt
x1:key
d1:elt
r1:t elt
h1:int
Hl1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp l1 cont0 e0) (elements l1 ++ l0) (flatten_e e0)
Hr1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp r1 cont0 e0) (elements r1 ++ l0) (flatten_e e0)
cont:enumeration elt -> bool
l:L.t elt
H:forall e : enumeration elt, IfEq (cont e) l (flatten_e e)
x2:key
d2:elt
r2:t elt
e2:enumeration elt

IfEq (equal_more cmp x1 d1 (equal_cont cmp r1 cont) (More x2 d2 r2 e2)) ((x1, d1) :: elements r1 ++ l) (flatten_e (More x2 d2 r2 e2))
elt:Type
cmp:elt -> elt -> bool
l1:t elt
x1:key
d1:elt
r1:t elt
h1:int
Hl1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp l1 cont0 e0) (elements l1 ++ l0) (flatten_e e0)
Hr1:forall (cont0 : enumeration elt -> bool) (e0 : enumeration elt) (l0 : L.t elt), (forall e : enumeration elt, IfEq (cont0 e) l0 (flatten_e e)) -> IfEq (equal_cont cmp r1 cont0 e0) (elements r1 ++ l0) (flatten_e e0)
cont:enumeration elt -> bool
l:L.t elt
H:forall e : enumeration elt, IfEq (cont e) l (flatten_e e)
x2:key
d2:elt
r2:t elt
e2:enumeration elt

IfEq (equal_cont cmp r1 cont (cons r2 e2)) (elements r1 ++ l) (elements r2 ++ flatten_e e2)
rewrite <- cons_1; auto. Qed.
elt:Type
cmp:elt -> elt -> bool

forall m1 m2 : t elt, IfEq (equal cmp m1 m2) (elements m1) (elements m2)
elt:Type
cmp:elt -> elt -> bool

forall m1 m2 : t elt, IfEq (equal cmp m1 m2) (elements m1) (elements m2)
elt:Type
cmp:elt -> elt -> bool
m1, m2:t elt

IfEq (equal_cont cmp m1 (equal_end (elt:=elt)) (cons m2 (End elt))) (elements m1) (elements m2)
elt:Type
cmp:elt -> elt -> bool
m1, m2:t elt

IfEq (equal_cont cmp m1 (equal_end (elt:=elt)) (cons m2 (End elt))) (elements m1 ++ nil) (elements m2)
elt:Type
cmp:elt -> elt -> bool
m1, m2:t elt

IfEq (equal_cont cmp m1 (equal_end (elt:=elt)) (cons m2 (End elt))) (elements m1 ++ nil) (flatten_e (cons m2 (End elt)))
elt:Type
cmp:elt -> elt -> bool
m1, m2:t elt

forall e : enumeration elt, IfEq (equal_end e) nil (flatten_e e)
elt:Type
cmp:elt -> elt -> bool
m1, m2:t elt
e:enumeration elt

IfEq (equal_end e) nil (flatten_e e)
apply equal_end_IfEq; auto. Qed. Definition Equivb m m' := (forall k, In k m <-> In k m') /\ (forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).
elt:Type
cmp:elt -> elt -> bool

forall s s' : t elt, Equivb s s' <-> L.Equivb cmp (elements s) (elements s')
elt:Type
cmp:elt -> elt -> bool

forall s s' : t elt, Equivb s s' <-> L.Equivb cmp (elements s) (elements s')
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : key, In k0 s <-> In k0 s') /\ (forall (k0 : key) (e e' : elt), MapsTo k0 e s -> MapsTo k0 e' s' -> cmp e e' = true)
k:X.t

L.PX.In k (elements s) <-> L.PX.In k (elements s')
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : key, In k0 s <-> In k0 s') /\ (forall (k0 : key) (e0 e'0 : elt), MapsTo k0 e0 s -> MapsTo k0 e'0 s' -> cmp e0 e'0 = true)
k:X.t
e, e':elt
H0:L.PX.MapsTo k e (elements s)
H1:L.PX.MapsTo k e' (elements s')
cmp e e' = true
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : X.t, L.PX.In k0 (elements s) <-> L.PX.In k0 (elements s')) /\ (forall (k0 : X.t) (e e' : elt), L.PX.MapsTo k0 e (elements s) -> L.PX.MapsTo k0 e' (elements s') -> cmp e e' = true)
k:key
In k s <-> In k s'
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : X.t, L.PX.In k0 (elements s) <-> L.PX.In k0 (elements s')) /\ (forall (k0 : X.t) (e0 e'0 : elt), L.PX.MapsTo k0 e0 (elements s) -> L.PX.MapsTo k0 e'0 (elements s') -> cmp e0 e'0 = true)
k:key
e, e':elt
H0:MapsTo k e s
H1:MapsTo k e' s'
cmp e e' = true
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : key, In k0 s <-> In k0 s') /\ (forall (k0 : key) (e0 e'0 : elt), MapsTo k0 e0 s -> MapsTo k0 e'0 s' -> cmp e0 e'0 = true)
k:X.t
e, e':elt
H0:L.PX.MapsTo k e (elements s)
H1:L.PX.MapsTo k e' (elements s')

cmp e e' = true
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : X.t, L.PX.In k0 (elements s) <-> L.PX.In k0 (elements s')) /\ (forall (k0 : X.t) (e e' : elt), L.PX.MapsTo k0 e (elements s) -> L.PX.MapsTo k0 e' (elements s') -> cmp e e' = true)
k:key
In k s <-> In k s'
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : X.t, L.PX.In k0 (elements s) <-> L.PX.In k0 (elements s')) /\ (forall (k0 : X.t) (e0 e'0 : elt), L.PX.MapsTo k0 e0 (elements s) -> L.PX.MapsTo k0 e'0 (elements s') -> cmp e0 e'0 = true)
k:key
e, e':elt
H0:MapsTo k e s
H1:MapsTo k e' s'
cmp e e' = true
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:forall k0 : key, In k0 s <-> In k0 s'
H2:forall (k0 : key) (e0 e'0 : elt), MapsTo k0 e0 s -> MapsTo k0 e'0 s' -> cmp e0 e'0 = true
k:X.t
e, e':elt
H0:L.PX.MapsTo k e (elements s)
H1:L.PX.MapsTo k e' (elements s')

cmp e e' = true
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : X.t, L.PX.In k0 (elements s) <-> L.PX.In k0 (elements s')) /\ (forall (k0 : X.t) (e e' : elt), L.PX.MapsTo k0 e (elements s) -> L.PX.MapsTo k0 e' (elements s') -> cmp e e' = true)
k:key
In k s <-> In k s'
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : X.t, L.PX.In k0 (elements s) <-> L.PX.In k0 (elements s')) /\ (forall (k0 : X.t) (e0 e'0 : elt), L.PX.MapsTo k0 e0 (elements s) -> L.PX.MapsTo k0 e'0 (elements s') -> cmp e0 e'0 = true)
k:key
e, e':elt
H0:MapsTo k e s
H1:MapsTo k e' s'
cmp e e' = true
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : X.t, L.PX.In k0 (elements s) <-> L.PX.In k0 (elements s')) /\ (forall (k0 : X.t) (e e' : elt), L.PX.MapsTo k0 e (elements s) -> L.PX.MapsTo k0 e' (elements s') -> cmp e e' = true)
k:key

In k s <-> In k s'
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : X.t, L.PX.In k0 (elements s) <-> L.PX.In k0 (elements s')) /\ (forall (k0 : X.t) (e0 e'0 : elt), L.PX.MapsTo k0 e0 (elements s) -> L.PX.MapsTo k0 e'0 (elements s') -> cmp e0 e'0 = true)
k:key
e, e':elt
H0:MapsTo k e s
H1:MapsTo k e' s'
cmp e e' = true
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:(forall k0 : X.t, L.PX.In k0 (elements s) <-> L.PX.In k0 (elements s')) /\ (forall (k0 : X.t) (e0 e'0 : elt), L.PX.MapsTo k0 e0 (elements s) -> L.PX.MapsTo k0 e'0 (elements s') -> cmp e0 e'0 = true)
k:key
e, e':elt
H0:MapsTo k e s
H1:MapsTo k e' s'

cmp e e' = true
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
H:forall k0 : X.t, L.PX.In k0 (elements s) <-> L.PX.In k0 (elements s')
H2:forall (k0 : X.t) (e0 e'0 : elt), L.PX.MapsTo k0 e0 (elements s) -> L.PX.MapsTo k0 e'0 (elements s') -> cmp e0 e'0 = true
k:key
e, e':elt
H0:MapsTo k e s
H1:MapsTo k e' s'

cmp e e' = true
apply (H2 k); unfold L.PX.MapsTo; rewrite elements_mapsto; auto. Qed.
elt:Type
cmp:elt -> elt -> bool

forall s s' : t elt, bst s -> bst s' -> equal cmp s s' = true <-> Equivb s s'
elt:Type
cmp:elt -> elt -> bool

forall s s' : t elt, bst s -> bst s' -> equal cmp s s' = true <-> Equivb s s'
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
B:bst s
B':bst s'

equal cmp s s' = true <-> Equivb s s'
elt:Type
cmp:elt -> elt -> bool
s, s':t elt
B:bst s
B':bst s'

L.equal cmp (elements s) (elements s') = true <-> L.Equivb cmp (elements s) (elements s')
split; [apply L.equal_2|apply L.equal_1]; auto. Qed. End Elt. Section Map. Variable elt elt' : Type. Variable f : elt -> elt'.
elt, elt':Type
f:elt -> elt'

forall (m : t elt) (x : key) (e : elt), MapsTo x e m -> MapsTo x (f e) (map f m)
elt, elt':Type
f:elt -> elt'

forall (m : t elt) (x : key) (e : elt), MapsTo x e m -> MapsTo x (f e) (map f m)
induction m; simpl; inversion_clear 1; auto. Qed.
elt, elt':Type
f:elt -> elt'

forall (m : t elt) (x : key), In x (map f m) -> In x m
elt, elt':Type
f:elt -> elt'

forall (m : t elt) (x : key), In x (map f m) -> In x m
induction m; simpl; inversion_clear 1; auto. Qed.
elt, elt':Type
f:elt -> elt'

forall m : t elt, bst m -> bst (map f m)
elt, elt':Type
f:elt -> elt'

forall m : t elt, bst m -> bst (map f m)
elt, elt':Type
f:elt -> elt'
m1:Raw.t elt
k:key
e:elt
m2:Raw.t elt
t:int
IHm1:bst m1 -> bst (map f m1)
IHm2:bst m2 -> bst (map f m2)

bst (Node m1 k e m2 t) -> bst (Node (map f m1) k (f e) (map f m2) t)
inversion_clear 1; constructor; auto; red; auto using map_2. Qed. End Map. Section Mapi. Variable elt elt' : Type. Variable f : key -> elt -> elt'.
elt, elt':Type
f:key -> elt -> elt'

forall (m : t elt) (x : key) (e : elt), MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)
elt, elt':Type
f:key -> elt -> elt'

forall (m : t elt) (x : key) (e : elt), MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)
elt, elt':Type
f:key -> elt -> elt'
m1:Raw.t elt
k:key
e:elt
m2:Raw.t elt
t:int
IHm1:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m1 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m1)
IHm2:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m2 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m2)
x:key
e0:elt
H0:X.eq x k

exists y : X.t, X.eq y x /\ MapsTo x (f y e) (Node (mapi f m1) k (f k e) (mapi f m2) t)
elt, elt':Type
f:key -> elt -> elt'
m1:Raw.t elt
k:key
e:elt
m2:Raw.t elt
t:int
IHm1:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m1 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m1)
IHm2:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m2 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m2)
x:key
e0:elt
H0:MapsTo x e0 m1
exists y : X.t, X.eq y x /\ MapsTo x (f y e0) (Node (mapi f m1) k (f k e) (mapi f m2) t)
elt, elt':Type
f:key -> elt -> elt'
m1:Raw.t elt
k:key
e:elt
m2:Raw.t elt
t:int
IHm1:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m1 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m1)
IHm2:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m2 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m2)
x:key
e0:elt
H0:MapsTo x e0 m2
exists y : X.t, X.eq y x /\ MapsTo x (f y e0) (Node (mapi f m1) k (f k e) (mapi f m2) t)
elt, elt':Type
f:key -> elt -> elt'
m1:Raw.t elt
k:key
e:elt
m2:Raw.t elt
t:int
IHm1:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m1 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m1)
IHm2:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m2 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m2)
x:key
e0:elt
H0:MapsTo x e0 m1

exists y : X.t, X.eq y x /\ MapsTo x (f y e0) (Node (mapi f m1) k (f k e) (mapi f m2) t)
elt, elt':Type
f:key -> elt -> elt'
m1:Raw.t elt
k:key
e:elt
m2:Raw.t elt
t:int
IHm1:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m1 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m1)
IHm2:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m2 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m2)
x:key
e0:elt
H0:MapsTo x e0 m2
exists y : X.t, X.eq y x /\ MapsTo x (f y e0) (Node (mapi f m1) k (f k e) (mapi f m2) t)
elt, elt':Type
f:key -> elt -> elt'
m1:Raw.t elt
k:key
e:elt
m2:Raw.t elt
t:int
IHm1:forall (x1 : key) (e1 : elt), MapsTo x1 e1 m1 -> exists y : X.t, X.eq y x1 /\ MapsTo x1 (f y e1) (mapi f m1)
IHm2:forall (x1 : key) (e1 : elt), MapsTo x1 e1 m2 -> exists y : X.t, X.eq y x1 /\ MapsTo x1 (f y e1) (mapi f m2)
x:key
e0:elt
H0:MapsTo x e0 m1
x0:X.t
H:X.eq x0 x /\ MapsTo x (f x0 e0) (mapi f m1)

exists y : X.t, X.eq y x /\ MapsTo x (f y e0) (Node (mapi f m1) k (f k e) (mapi f m2) t)
elt, elt':Type
f:key -> elt -> elt'
m1:Raw.t elt
k:key
e:elt
m2:Raw.t elt
t:int
IHm1:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m1 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m1)
IHm2:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m2 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m2)
x:key
e0:elt
H0:MapsTo x e0 m2
exists y : X.t, X.eq y x /\ MapsTo x (f y e0) (Node (mapi f m1) k (f k e) (mapi f m2) t)
elt, elt':Type
f:key -> elt -> elt'
m1:Raw.t elt
k:key
e:elt
m2:Raw.t elt
t:int
IHm1:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m1 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m1)
IHm2:forall (x0 : key) (e1 : elt), MapsTo x0 e1 m2 -> exists y : X.t, X.eq y x0 /\ MapsTo x0 (f y e1) (mapi f m2)
x:key
e0:elt
H0:MapsTo x e0 m2

exists y : X.t, X.eq y x /\ MapsTo x (f y e0) (Node (mapi f m1) k (f k e) (mapi f m2) t)
elt, elt':Type
f:key -> elt -> elt'
m1:Raw.t elt
k:key
e:elt
m2:Raw.t elt
t:int
IHm1:forall (x1 : key) (e1 : elt), MapsTo x1 e1 m1 -> exists y : X.t, X.eq y x1 /\ MapsTo x1 (f y e1) (mapi f m1)
IHm2:forall (x1 : key) (e1 : elt), MapsTo x1 e1 m2 -> exists y : X.t, X.eq y x1 /\ MapsTo x1 (f y e1) (mapi f m2)
x:key
e0:elt
H0:MapsTo x e0 m2
x0:X.t
H:X.eq x0 x /\ MapsTo x (f x0 e0) (mapi f m2)

exists y : X.t, X.eq y x /\ MapsTo x (f y e0) (Node (mapi f m1) k (f k e) (mapi f m2) t)
exists x0; intuition. Qed.
elt, elt':Type
f:key -> elt -> elt'

forall (m : t elt) (x : key), In x (mapi f m) -> In x m
elt, elt':Type
f:key -> elt -> elt'

forall (m : t elt) (x : key), In x (mapi f m) -> In x m
induction m; simpl; inversion_clear 1; auto. Qed.
elt, elt':Type
f:key -> elt -> elt'

forall m : t elt, bst m -> bst (mapi f m)
elt, elt':Type
f:key -> elt -> elt'

forall m : t elt, bst m -> bst (mapi f m)
elt, elt':Type
f:key -> elt -> elt'
m1:Raw.t elt
k:key
e:elt
m2:Raw.t elt
t:int
IHm1:bst m1 -> bst (mapi f m1)
IHm2:bst m2 -> bst (mapi f m2)

bst (Node m1 k e m2 t) -> bst (Node (mapi f m1) k (f k e) (mapi f m2) t)
inversion_clear 1; constructor; auto; red; auto using mapi_2. Qed. End Mapi. Section Map_option. Variable elt elt' : Type. Variable f : key -> elt -> option elt'. Hypothesis f_compat : forall x x' d, X.eq x x' -> f x d = f x' d.
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x x' : X.t) (d : elt), X.eq x x' -> f x d = f x' d

forall (m : t elt) (x : key), In x (map_option f m) -> exists d : elt, MapsTo x d m /\ f x d <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x x' : X.t) (d : elt), X.eq x x' -> f x d = f x' d

forall (m : t elt) (x : key), In x (map_option f m) -> exists d : elt, MapsTo x d m /\ f x d <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x0 x' : X.t) (d : elt), X.eq x0 x' -> f x0 d = f x' d
x:key
H:In x (Leaf elt')

exists d : elt, MapsTo x d (Leaf elt) /\ f x d <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (join (map_option f l) x d' (map_option f r))
exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (concat (map_option f l) (map_option f r))
exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (join (map_option f l) x d' (map_option f r))

exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (concat (map_option f l) (map_option f r))
exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:X.eq x0 x

exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (map_option f l)
exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (map_option f r)
exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (concat (map_option f l) (map_option f r))
exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (map_option f l)

exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (map_option f r)
exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (concat (map_option f l) (map_option f r))
exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (map_option f r)

exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (concat (map_option f l) (map_option f r))
exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (concat (map_option f l) (map_option f r))

exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (map_option f l)

exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (map_option f r)
exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, In x1 (map_option f l) -> exists d0 : elt, MapsTo x1 d0 l /\ f x1 d0 <> None
IHt0:forall x1 : key, In x1 (map_option f r) -> exists d0 : elt, MapsTo x1 d0 r /\ f x1 d0 <> None
x0:key
H:In x0 (map_option f r)

exists d0 : elt, MapsTo x0 d0 (Node l x d r _x) /\ f x0 d0 <> None
destruct (IHt0 _ H) as (d0 & ? & ?); exists d0; auto. Qed.
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x x' : X.t) (d : elt), X.eq x x' -> f x d = f x' d

forall m : t elt, bst m -> bst (map_option f m)
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x x' : X.t) (d : elt), X.eq x x' -> f x d = f x' d

forall m : t elt, bst m -> bst (map_option f m)
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x0 x' : X.t) (d0 : elt), X.eq x0 x' -> f x0 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:bst l -> bst (map_option f l)
IHt0:bst r -> bst (map_option f r)
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r

bst (join (map_option f l) x d' (map_option f r))
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x0 x' : X.t) (d0 : elt), X.eq x0 x' -> f x0 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:bst l -> bst (map_option f l)
IHt0:bst r -> bst (map_option f r)
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
bst (concat (map_option f l) (map_option f r))
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x0 x' : X.t) (d0 : elt), X.eq x0 x' -> f x0 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:bst l -> bst (map_option f l)
IHt0:bst r -> bst (map_option f r)
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r

bst (concat (map_option f l) (map_option f r))
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x0 x' : X.t) (d0 : elt), X.eq x0 x' -> f x0 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:bst l -> bst (map_option f l)
IHt0:bst r -> bst (map_option f r)
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
y, y':key
H:In y (map_option f l)
H':In y' (map_option f r)

X.lt y y'
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x0 x' : X.t) (d1 : elt), X.eq x0 x' -> f x0 d1 = f x' d1
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:bst l -> bst (map_option f l)
IHt0:bst r -> bst (map_option f r)
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
y, y':key
H:In y (map_option f l)
H':In y' (map_option f r)
d0:elt
H4:MapsTo y d0 l
H5:f y d0 <> None

X.lt y y'
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x0 x' : X.t) (d1 : elt), X.eq x0 x' -> f x0 d1 = f x' d1
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:bst l -> bst (map_option f l)
IHt0:bst r -> bst (map_option f r)
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
y, y':key
H:In y (map_option f l)
H':In y' (map_option f r)
d0:elt
H4:MapsTo y d0 l
H5:f y d0 <> None
d0':elt
H6:MapsTo y' d0' r
H7:f y' d0' <> None

X.lt y y'
eapply X.lt_trans with x; eauto using MapsTo_In. Qed. Hint Resolve map_option_bst : core. Ltac nonify e := replace e with (@None elt) by (symmetry; rewrite not_find_iff; auto; intro; order).
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x x' : X.t) (d : elt), X.eq x x' -> f x d = f x' d

forall (m : t elt) (x : key), bst m -> find x (map_option f m) = match find x m with | Some d => f x d | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x x' : X.t) (d : elt), X.eq x x' -> f x d = f x' d

forall (m : t elt) (x : key), bst m -> find x (map_option f m) = match find x m with | Some d => f x d | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Heq:X.eq x0 x

Some d' = f x0 d
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
lt_tree x (map_option f l)
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
gt_tree x (map_option f r)
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x0 x
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = match find x0 l with | Some d0 => f x0 d0 | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Heq:X.eq x0 x
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = f x0 d
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x x0
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = match find x0 r with | Some d0 => f x0 d0 | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
forall y1 y2 : key, In y1 (map_option f l) -> In y2 (map_option f r) -> X.lt y1 y2
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r

lt_tree x (map_option f l)
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
gt_tree x (map_option f r)
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x0 x
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = match find x0 l with | Some d0 => f x0 d0 | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Heq:X.eq x0 x
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = f x0 d
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x x0
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = match find x0 r with | Some d0 => f x0 d0 | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
forall y1 y2 : key, In y1 (map_option f l) -> In y2 (map_option f r) -> X.lt y1 y2
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
d':elt'
e0:f x d = Some d'
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r

gt_tree x (map_option f r)
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x0 x
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = match find x0 l with | Some d0 => f x0 d0 | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Heq:X.eq x0 x
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = f x0 d
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x x0
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = match find x0 r with | Some d0 => f x0 d0 | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
forall y1 y2 : key, In y1 (map_option f l) -> In y2 (map_option f r) -> X.lt y1 y2
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x0 x

match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = match find x0 l with | Some d0 => f x0 d0 | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Heq:X.eq x0 x
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = f x0 d
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x x0
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = match find x0 r with | Some d0 => f x0 d0 | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
forall y1 y2 : key, In y1 (map_option f l) -> In y2 (map_option f r) -> X.lt y1 y2
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Heq:X.eq x0 x

match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = f x0 d
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x x0
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = match find x0 r with | Some d0 => f x0 d0 | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
forall y1 y2 : key, In y1 (map_option f l) -> In y2 (map_option f r) -> X.lt y1 y2
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Heq:X.eq x0 x

None = f x0 d
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x x0
match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = match find x0 r with | Some d0 => f x0 d0 | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
forall y1 y2 : key, In y1 (map_option f l) -> In y2 (map_option f r) -> X.lt y1 y2
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x x0

match find x0 (map_option f r) with | Some d0 => Some d0 | None => find x0 (map_option f l) end = match find x0 r with | Some d0 => f x0 d0 | None => None end
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
forall y1 y2 : key, In y1 (map_option f l) -> In y2 (map_option f r) -> X.lt y1 y2
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
Hlt:X.lt x x0

match find x0 (map_option f r) with | Some d0 => Some d0 | None => None end = find x0 (map_option f r)
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
forall y1 y2 : key, In y1 (map_option f l) -> In y2 (map_option f r) -> X.lt y1 y2
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r

forall y1 y2 : key, In y1 (map_option f l) -> In y2 (map_option f r) -> X.lt y1 y2
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x1 x' : X.t) (d0 : elt), X.eq x1 x' -> f x1 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x1 : key, bst l -> find x1 (map_option f l) = match find x1 l with | Some d0 => f x1 d0 | None => None end
IHt0:forall x1 : key, bst r -> find x1 (map_option f r) = match find x1 r with | Some d0 => f x1 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
y, y':key
H:In y (map_option f l)
H':In y' (map_option f r)

X.lt y y'
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x2 x' : X.t) (d0 : elt), X.eq x2 x' -> f x2 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x2 : key, bst l -> find x2 (map_option f l) = match find x2 l with | Some d0 => f x2 d0 | None => None end
IHt0:forall x2 : key, bst r -> find x2 (map_option f r) = match find x2 r with | Some d0 => f x2 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
y, y':key
H:In y (map_option f l)
H':In y' (map_option f r)
x1:elt
H4:MapsTo y x1 l
H5:f y x1 <> None

X.lt y y'
elt, elt':Type
f:key -> elt -> option elt'
f_compat:forall (x3 x' : X.t) (d0 : elt), X.eq x3 x' -> f x3 d0 = f x' d0
l:t elt
x:key
d:elt
r:t elt
_x:int
e0:f x d = None
IHt:forall x3 : key, bst l -> find x3 (map_option f l) = match find x3 l with | Some d0 => f x3 d0 | None => None end
IHt0:forall x3 : key, bst r -> find x3 (map_option f r) = match find x3 r with | Some d0 => f x3 d0 | None => None end
x0:key
H0:bst l
H1:bst r
H2:lt_tree x l
H3:gt_tree x r
y, y':key
H:In y (map_option f l)
H':In y' (map_option f r)
x1:elt
H4:MapsTo y x1 l
H5:f y x1 <> None
x2:elt
H6:MapsTo y' x2 r
H7:f y' x2 <> None

X.lt y y'
eapply X.lt_trans with x; eauto using MapsTo_In. Qed. End Map_option. Section Map2_opt. Variable elt elt' elt'' : Type. Variable f0 : key -> option elt -> option elt' -> option elt''. Variable f : key -> elt -> option elt' -> option elt''. Variable mapl : t elt -> t elt''. Variable mapr : t elt' -> t elt''. Hypothesis f0_f : forall x d o, f x d o = f0 x (Some d) o. Hypothesis mapl_bst : forall m, bst m -> bst (mapl m). Hypothesis mapr_bst : forall m', bst m' -> bst (mapr m'). Hypothesis mapl_f0 : forall x m, bst m -> find x (mapl m) = match find x m with Some d => f0 x (Some d) None | None => None end. Hypothesis mapr_f0 : forall x m', bst m' -> find x (mapr m') = match find x m' with Some d' => f0 x None (Some d') | None => None end. Hypothesis f0_compat : forall x x' o o', X.eq x x' -> f0 x o o' = f0 x' o o'. Notation map2_opt := (map2_opt f mapl mapr).
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'

forall (m : t elt) (m' : t elt') (y : key), bst m -> bst m' -> In y (map2_opt m m') -> In y m \/ In y m'
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'

forall (m : t elt) (m' : t elt') (y : key), bst m -> bst m' -> In y (map2_opt m m') -> In y m \/ In y m'
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
m2:t elt'
y:key
H:bst (Leaf elt)
H0:bst m2
H1:In y (mapr m2)

In y (Leaf elt) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
y:key
H:bst (Node l1 x1 d1 r1 _x)
H0:bst (Leaf elt')
H1:In y (mapl (Node l1 x1 d1 r1 _x))
In y (Node l1 x1 d1 r1 _x) \/ In y (Leaf elt')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (join (map2_opt l1 l2') x1 e2 (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (concat (map2_opt l1 l2') (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
m2:t elt'
y:key
H:bst (Leaf elt)
H0:bst m2
H1:In y (mapr m2)

find y m2 <> None
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
y:key
H:bst (Node l1 x1 d1 r1 _x)
H0:bst (Leaf elt')
H1:In y (mapl (Node l1 x1 d1 r1 _x))
In y (Node l1 x1 d1 r1 _x) \/ In y (Leaf elt')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (join (map2_opt l1 l2') x1 e2 (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (concat (map2_opt l1 l2') (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
m2:t elt'
y:key
H:bst (Leaf elt)
H0:bst m2
H1:In y (mapr m2)

match find y m2 with | Some d' => f0 y None (Some d') | None => None end <> None -> find y m2 <> None
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
y:key
H:bst (Node l1 x1 d1 r1 _x)
H0:bst (Leaf elt')
H1:In y (mapl (Node l1 x1 d1 r1 _x))
In y (Node l1 x1 d1 r1 _x) \/ In y (Leaf elt')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (join (map2_opt l1 l2') x1 e2 (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (concat (map2_opt l1 l2') (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
y:key
H:bst (Node l1 x1 d1 r1 _x)
H0:bst (Leaf elt')
H1:In y (mapl (Node l1 x1 d1 r1 _x))

In y (Node l1 x1 d1 r1 _x) \/ In y (Leaf elt')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (join (map2_opt l1 l2') x1 e2 (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (concat (map2_opt l1 l2') (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
y:key
m1:t elt
H:bst m1
H0:bst (Leaf elt')
H1:In y (mapl m1)

In y m1 \/ In y (Leaf elt')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (join (map2_opt l1 l2') x1 e2 (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (concat (map2_opt l1 l2') (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
y:key
m1:t elt
H:bst m1
H0:bst (Leaf elt')
H1:In y (mapl m1)

find y m1 <> None
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (join (map2_opt l1 l2') x1 e2 (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (concat (map2_opt l1 l2') (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
y:key
m1:t elt
H:bst m1
H0:bst (Leaf elt')
H1:In y (mapl m1)

match find y m1 with | Some d => f0 y (Some d) None | None => None end <> None -> find y m1 <> None
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (join (map2_opt l1 l2') x1 e2 (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (concat (map2_opt l1 l2') (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (join (map2_opt l1 l2') x1 e2 (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1

In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (concat (map2_opt l1 l2') (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H':In y (map2_opt l1 l2')
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1

In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H':In y (map2_opt r1 r2')
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (concat (map2_opt l1 l2') (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H':In y (map2_opt r1 r2')
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1

In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (concat (map2_opt l1 l2') (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H1:In y (concat (map2_opt l1 l2') (map2_opt r1 r2'))
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1

In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H':In y (map2_opt l1 l2')
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1

In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H':In y (map2_opt r1 r2')
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1
In y (Node l1 x1 d1 r1 _x) \/ In y m2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 (map2_opt l1 l2') -> In y0 l1 \/ In y0 l2'
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 (map2_opt r1 r2') -> In y0 r1 \/ In y0 r2'
y:key
H0:bst m2
H':In y (map2_opt r1 r2')
H2:In y l2' <-> In y m2 /\ X.lt y x1
H3:In y r2' <-> In y m2 /\ X.lt x1 y
H4:bst l2'
H5:bst r2'
H6:bst l1
H7:bst r1
H8:lt_tree x1 l1
H9:gt_tree x1 r1

In y (Node l1 x1 d1 r1 _x) \/ In y m2
destruct (IHt0 y H7 H5 H'); intuition. Qed.
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'

forall (m : t elt) (m' : t elt'), bst m -> bst m' -> bst (map2_opt m m')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'

forall (m : t elt) (m' : t elt'), bst m -> bst m' -> bst (map2_opt m m')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'

bst (join (map2_opt l1 l2') x1 e2 (map2_opt r1 r2'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'
bst (concat (map2_opt l1 l2') (map2_opt r1 r2'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'

lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'
bst (concat (map2_opt l1 l2') (map2_opt r1 r2'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
y:key
H:In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H6:bst l2'
H7:bst r2'
Hy:In y (map2_opt l1 l2')

X.lt y x1
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'
bst (concat (map2_opt l1 l2') (map2_opt r1 r2'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'

gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'
bst (concat (map2_opt l1 l2') (map2_opt r1 r2'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:elt''
e3:f x1 d1 o2 = Some e2
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
y:key
H5:In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'
Hy:In y (map2_opt r1 r2')

X.lt x1 y
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'
bst (concat (map2_opt l1 l2') (map2_opt r1 r2'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'

bst (concat (map2_opt l1 l2') (map2_opt r1 r2'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
H:forall y : key, In y l2' <-> In y m2 /\ X.lt y x1
H5:forall y : key, In y r2' <-> In y m2 /\ X.lt x1 y
H6:bst l2'
H7:bst r2'

forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
y:key
H:In y l2' <-> In y m2 /\ X.lt y x1
y':key
H5:In y' r2' <-> In y' m2 /\ X.lt x1 y'
H6:bst l2'
H7:bst r2'
Hy:In y (map2_opt l1 l2')
Hy':In y' (map2_opt r1 r2')

X.lt y y'
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
y:key
H:In y l2' <-> In y m2 /\ X.lt y x1
y':key
H5:In y' r2' <-> In y' m2 /\ X.lt x1 y'
H6:bst l2'
H7:bst r2'
Hy:In y (map2_opt l1 l2')
Hy':In y' (map2_opt r1 r2')

X.lt y x1
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
y:key
H:In y l2' <-> In y m2 /\ X.lt y x1
y':key
H5:In y' r2' <-> In y' m2 /\ X.lt x1 y'
H6:bst l2'
H7:bst r2'
Hy:In y (map2_opt l1 l2')
Hy':In y' (map2_opt r1 r2')
X.lt x1 y'
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2':t elt'
o2:option elt'
r2', m2:t elt'
e1:split x1 m2 = << l2', o2, r2' >>
e2:f x1 d1 o2 = None
IHt1:bst l1 -> bst l2' -> bst (map2_opt l1 l2')
IHt0:bst r1 -> bst r2' -> bst (map2_opt r1 r2')
H0:bst m2
H1:bst l1
H2:bst r1
H3:lt_tree x1 l1
H4:gt_tree x1 r1
y:key
H:In y l2' <-> In y m2 /\ X.lt y x1
y':key
H5:In y' r2' <-> In y' m2 /\ X.lt x1 y'
H6:bst l2'
H7:bst r2'
Hy:In y (map2_opt l1 l2')
Hy':In y' (map2_opt r1 r2')

X.lt x1 y'
destruct (map2_opt_2 H2 H7 Hy'); intuition. Qed. Hint Resolve map2_opt_bst : core. Ltac map2_aux := match goal with | H : In ?x _ \/ In ?x ?m, H' : find ?x ?m = find ?x ?m', B:bst ?m, B':bst ?m' |- _ => destruct H; [ intuition_in; order | rewrite <-(find_in_equiv B B' H'); auto ] end. Ltac nonify t := match t with (find ?y (map2_opt ?m ?m')) => replace t with (@None elt''); [ | symmetry; rewrite not_find_iff; auto; intro; destruct (@map2_opt_2 m m' y); auto; order ] end.
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'

forall (m : t elt) (m' : t elt') (y : key), bst m -> bst m' -> In y m \/ In y m' -> find y (map2_opt m m') = f0 y (find y m) (find y m')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'

forall (m : t elt) (m' : t elt') (y : key), bst m -> bst m' -> In y m \/ In y m' -> find y (map2_opt m m') = f0 y (find y m) (find y m')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
m2:t elt'
y:key
H:bst (Leaf elt)
H0:bst m2
H1:In y (Leaf elt) \/ In y m2

find y (mapr m2) = f0 y (find y (Leaf elt)) (find y m2)
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
y:key
H:bst (Node l1 x1 d1 r1 _x)
H0:bst (Leaf elt')
H1:In y (Node l1 x1 d1 r1 _x) \/ In y (Leaf elt')
find y (mapl (Node l1 x1 d1 r1 _x)) = f0 y (find y (Node l1 x1 d1 r1 _x)) (find y (Leaf elt'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
find y (create (map2_opt l1 l2') x1 e2 (map2_opt r1 r2')) = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
m2:t elt'
y:key
H:bst (Leaf elt)
H0:bst m2
H1:In y m2

find y (mapr m2) = f0 y None (find y m2)
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
y:key
H:bst (Node l1 x1 d1 r1 _x)
H0:bst (Leaf elt')
H1:In y (Node l1 x1 d1 r1 _x) \/ In y (Leaf elt')
find y (mapl (Node l1 x1 d1 r1 _x)) = f0 y (find y (Node l1 x1 d1 r1 _x)) (find y (Leaf elt'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
find y (create (map2_opt l1 l2') x1 e2 (map2_opt r1 r2')) = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
m2:t elt'
y:key
H:bst (Leaf elt)
H0:bst m2
H1:In y m2

match find y m2 with | Some d' => f0 y None (Some d') | None => None end = f0 y None (find y m2)
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
y:key
H:bst (Node l1 x1 d1 r1 _x)
H0:bst (Leaf elt')
H1:In y (Node l1 x1 d1 r1 _x) \/ In y (Leaf elt')
find y (mapl (Node l1 x1 d1 r1 _x)) = f0 y (find y (Node l1 x1 d1 r1 _x)) (find y (Leaf elt'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
find y (create (map2_opt l1 l2') x1 e2 (map2_opt r1 r2')) = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
y:key
H:bst (Node l1 x1 d1 r1 _x)
H0:bst (Leaf elt')
H1:In y (Node l1 x1 d1 r1 _x) \/ In y (Leaf elt')

find y (mapl (Node l1 x1 d1 r1 _x)) = f0 y (find y (Node l1 x1 d1 r1 _x)) (find y (Leaf elt'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
find y (create (map2_opt l1 l2') x1 e2 (map2_opt r1 r2')) = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
y:key
m1:t elt
H:bst m1
H0:bst (Leaf elt')
H1:In y m1 \/ In y (Leaf elt')

find y (mapl m1) = f0 y (find y m1) (find y (Leaf elt'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
find y (create (map2_opt l1 l2') x1 e2 (map2_opt r1 r2')) = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
y:key
m1:t elt
H:bst m1
H0:bst (Leaf elt')
H1:In y m1

find y (mapl m1) = f0 y (find y m1) (find y (Leaf elt'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
find y (create (map2_opt l1 l2') x1 e2 (map2_opt r1 r2')) = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
y:key
m1:t elt
H:bst m1
H0:bst (Leaf elt')
H1:In y m1

match find y m1 with | Some d => f0 y (Some d) None | None => None end = f0 y (find y m1) (find y (Leaf elt'))
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
find y (create (map2_opt l1 l2') x1 e2 (map2_opt r1 r2')) = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

find y (create (map2_opt l1 l2') x1 e2 (map2_opt r1 r2')) = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt y x1
H7:find y m2 = find y l2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

find y (map2_opt l1 l2') = f0 y (find y l1) (find y l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
e:X.eq y x1
H7:find y m2 = find x1 m2
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
Some e2 = f0 y (Some d1) (find x1 m2)
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt x1 y
H7:find y m2 = find y r2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
find y (map2_opt r1 r2') = f0 y (find y r1) (find y r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
e:X.eq y x1
H7:find y m2 = find x1 m2
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

Some e2 = f0 y (Some d1) (find x1 m2)
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt x1 y
H7:find y m2 = find y r2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
find y (map2_opt r1 r2') = f0 y (find y r1) (find y r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt x1 y
H7:find y m2 = find y r2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

find y (map2_opt r1 r2') = f0 y (find y r1) (find y r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

lt_tree x1 (map2_opt l1 l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e1:split x1 m2 = << l2', find x1 m2, r2' >>
e2:elt''
e3:f x1 d1 (find x1 m2) = Some e2
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

gt_tree x1 (map2_opt r1 r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y match X.compare y x1 with | LT _ => find y l1 | EQ _ => Some d1 | GT _ => find y r1 end match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt y x1
H7:find y m2 = find y l2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y (find y l1) (find y l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
e:X.eq y x1
H7:find y m2 = find x1 m2
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y (Some d1) (find x1 m2)
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt x1 y
H7:find y m2 = find y r2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y (find y r1) (find y r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt y x1
H7:find y m2 = find y l2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

find y (map2_opt l1 l2') = f0 y (find y l1) (find y l2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
e:X.eq y x1
H7:find y m2 = find x1 m2
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y (Some d1) (find x1 m2)
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt x1 y
H7:find y m2 = find y r2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y (find y r1) (find y r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
e:X.eq y x1
H7:find y m2 = find x1 m2
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y (Some d1) (find x1 m2)
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt x1 y
H7:find y m2 = find y r2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y (find y r1) (find y r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
e:X.eq y x1
H7:find y m2 = find x1 m2
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

find y (map2_opt l1 l2') = f0 y (Some d1) (find x1 m2)
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt x1 y
H7:find y m2 = find y r2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y (find y r1) (find y r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
e:X.eq y x1
H7:find y m2 = find x1 m2
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

None = f0 y (Some d1) (find x1 m2)
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt x1 y
H7:find y m2 = find y r2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y (find y r1) (find y r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt x1 y
H7:find y m2 = find y r2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

match find y (map2_opt r1 r2') with | Some d => Some d | None => find y (map2_opt l1 l2') end = f0 y (find y r1) (find y r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt x1 y
H7:find y m2 = find y r2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

match find y (map2_opt r1 r2') with | Some d => Some d | None => None end = f0 y (find y r1) (find y r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
l:X.lt x1 y
H7:find y m2 = find y r2'
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

match f0 y (find y r1) (find y r2') with | Some d => Some d | None => None end = f0 y (find y r1) (find y r2')
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1

forall y1 y2 : key, In y1 (map2_opt l1 l2') -> In y2 (map2_opt r1 r2') -> X.lt y1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
y1, y2:key
Hy1:In y1 (map2_opt l1 l2')
Hy2:In y2 (map2_opt r1 r2')

X.lt y1 x1
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
y1, y2:key
Hy1:In y1 (map2_opt l1 l2')
Hy2:In y2 (map2_opt r1 r2')
X.lt x1 y2
elt, elt', elt'':Type
f0:key -> option elt -> option elt' -> option elt''
f:key -> elt -> option elt' -> option elt''
mapl:t elt -> t elt''
mapr:t elt' -> t elt''
f0_f:forall (x : key) (d : elt) (o : option elt'), f x d o = f0 x (Some d) o
mapl_bst:forall m : t elt, bst m -> bst (mapl m)
mapr_bst:forall m' : t elt', bst m' -> bst (mapr m')
mapl_f0:forall (x : X.t) (m : t elt), bst m -> find x (mapl m) = match find x m with | Some d => f0 x (Some d) None | None => None end
mapr_f0:forall (x : X.t) (m' : t elt'), bst m' -> find x (mapr m') = match find x m' with | Some d' => f0 x None (Some d') | None => None end
f0_compat:forall (x x' : X.t) (o : option elt) (o' : option elt'), X.eq x x' -> f0 x o o' = f0 x' o o'
l1:t elt
x1:key
d1:elt
r1:t elt
_x:int
l2', r2', m2:t elt'
e2:f x1 d1 (find x1 m2) = None
e1:split x1 m2 = << l2', find x1 m2, r2' >>
IHt1:forall y0 : key, bst l1 -> bst l2' -> In y0 l1 \/ In y0 l2' -> find y0 (map2_opt l1 l2') = f0 y0 (find y0 l1) (find y0 l2')
IHt0:forall y0 : key, bst r1 -> bst r2' -> In y0 r1 \/ In y0 r2' -> find y0 (map2_opt r1 r2') = f0 y0 (find y0 r1) (find y0 r2')
y:key
H0:bst m2
H1:In y (Node l1 x1 d1 r1 _x) \/ In y m2
H2:forall y0 : key, In y0 l2' <-> In y0 m2 /\ X.lt y0 x1
H3:forall y0 : key, In y0 r2' <-> In y0 m2 /\ X.lt x1 y0
H5:bst l2'
H6:bst r2'
H7:find y m2 = match X.compare y x1 with | LT _ => find y l2' | EQ _ => find x1 m2 | GT _ => find y r2' end
H8:lt_tree x1 l2'
H9:gt_tree x1 r2'
H10:bst l1
H11:bst r1
H12:lt_tree x1 l1
H13:gt_tree x1 r1
y1, y2:key
Hy1:In y1 (map2_opt l1 l2')
Hy2:In y2 (map2_opt r1 r2')

X.lt x1 y2
destruct (@map2_opt_2 r1 r2' y2); auto. Qed. End Map2_opt. Section Map2. Variable elt elt' elt'' : Type. Variable f : option elt -> option elt' -> option elt''.
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''

forall (m : t elt) (m' : t elt'), bst m -> bst m' -> bst (map2 f m m')
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''

forall (m : t elt) (m' : t elt'), bst m -> bst m' -> bst (map2 f m m')
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
H:bst m
H0:bst m'

bst (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')
apply map2_opt_bst with (fun _ => f); auto using map_option_bst; intros; rewrite map_option_find; auto. Qed.
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''

forall (m : t elt) (m' : t elt') (y : key), bst m -> bst m' -> In y m \/ In y m' -> find y (map2 f m m') = f (find y m) (find y m')
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''

forall (m : t elt) (m' : t elt') (y : key), bst m -> bst m' -> In y m \/ In y m' -> find y (map2 f m m') = f (find y m) (find y m')
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y m \/ In y m'

find y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m') = f (find y m) (find y m')
rewrite (map2_opt_1 (f0:=fun _ => f)); auto using map_option_bst; intros; rewrite map_option_find; auto. Qed.
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''

forall (m : t elt) (m' : t elt') (y : key), bst m -> bst m' -> In y (map2 f m m') -> In y m \/ In y m'
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''

forall (m : t elt) (m' : t elt') (y : key), bst m -> bst m' -> In y (map2 f m m') -> In y m \/ In y m'
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')

In y m \/ In y m'
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')
m0:t elt
H2:bst m0

bst (map_option (fun (_ : key) (d : elt) => f (Some d) None) m0)
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')
m'0:t elt'
H2:bst m'0
bst (map_option (fun (_ : key) (d' : elt') => f None (Some d')) m'0)
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')
x:X.t
m0:t elt
H2:bst m0
find x (map_option (fun (_ : key) (d : elt) => f (Some d) None) m0) = match find x m0 with | Some d => f (Some d) None | None => None end
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')
x:X.t
m'0:t elt'
H2:bst m'0
find x (map_option (fun (_ : key) (d' : elt') => f None (Some d')) m'0) = match find x m'0 with | Some d' => f None (Some d') | None => None end
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')
m'0:t elt'
H2:bst m'0

bst (map_option (fun (_ : key) (d' : elt') => f None (Some d')) m'0)
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')
x:X.t
m0:t elt
H2:bst m0
find x (map_option (fun (_ : key) (d : elt) => f (Some d) None) m0) = match find x m0 with | Some d => f (Some d) None | None => None end
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')
x:X.t
m'0:t elt'
H2:bst m'0
find x (map_option (fun (_ : key) (d' : elt') => f None (Some d')) m'0) = match find x m'0 with | Some d' => f None (Some d') | None => None end
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')
x:X.t
m0:t elt
H2:bst m0

find x (map_option (fun (_ : key) (d : elt) => f (Some d) None) m0) = match find x m0 with | Some d => f (Some d) None | None => None end
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')
x:X.t
m'0:t elt'
H2:bst m'0
find x (map_option (fun (_ : key) (d' : elt') => f None (Some d')) m'0) = match find x m'0 with | Some d' => f None (Some d') | None => None end
elt, elt', elt'':Type
f:option elt -> option elt' -> option elt''
m:t elt
m':t elt'
y:key
H:bst m
H0:bst m'
H1:In y (map2_opt (fun (_ : key) (d : elt) (o : option elt') => f (Some d) o) (map_option (fun (_ : key) (d : elt) => f (Some d) None)) (map_option (fun (_ : key) (d' : elt') => f None (Some d'))) m m')
x:X.t
m'0:t elt'
H2:bst m'0

find x (map_option (fun (_ : key) (d' : elt') => f None (Some d')) m'0) = match find x m'0 with | Some d' => f None (Some d') | None => None end
rewrite map_option_find; auto. Qed. End Map2. End Proofs. 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.
 Import Raw.Proofs.

 #[universes(template)]
 Record bst (elt:Type) :=
  Bst {this :> Raw.tree elt; is_bst : Raw.bst this}.

 Definition t := bst.
 Definition key := E.t.

 Section Elt.
 Variable elt elt' elt'': Type.

 Implicit Types m : t elt.
 Implicit Types x y : key.
 Implicit Types e : elt.

 Definition empty : t elt := Bst (empty_bst elt).
 Definition is_empty m : bool := Raw.is_empty (this m).
 Definition add x e m : t elt := Bst (add_bst x e (is_bst m)).
 Definition remove x m : t elt := Bst (remove_bst x (is_bst m)).
 Definition mem x m : bool := Raw.mem x (this m).
 Definition find x m : option elt := Raw.find x (this m).
 Definition map f m : t elt' := Bst (map_bst f (is_bst m)).
 Definition mapi (f:key->elt->elt') m : t elt' :=
  Bst (mapi_bst f (is_bst m)).
 Definition map2 f m (m':t elt') : t elt'' :=
  Bst (map2_bst f (is_bst m) (is_bst m')).
 Definition elements m : list (key*elt) := Raw.elements (this m).
 Definition cardinal m := Raw.cardinal (this m).
 Definition fold (A:Type) (f:key->elt->A->A) m i := Raw.fold (A:=A) f (this m) i.
 Definition equal cmp m m' : bool := Raw.equal cmp (this m) (this m').

 Definition MapsTo x e m : Prop := Raw.MapsTo x e (this m).
 Definition In x m : Prop := Raw.In0 x (this m).
 Definition Empty m : Prop := Empty (this m).

 Definition eq_key : (key*elt) -> (key*elt) -> Prop := @PX.eqk elt.
 Definition eq_key_elt : (key*elt) -> (key*elt) -> Prop := @PX.eqke elt.
 Definition lt_key : (key*elt) -> (key*elt) -> Prop := @PX.ltk elt.

 
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e : elt), E.eq x y -> MapsTo x e m -> MapsTo y e m
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e : elt), E.eq x y -> MapsTo x e m -> MapsTo y e m
intros m; exact (@MapsTo_1 _ (this m)). Qed.
elt, elt', elt'':Type

forall (m : t elt) (x : key), In x m -> mem x m = true
elt, elt', elt'':Type

forall (m : t elt) (x : key), In x m -> mem x m = true
elt, elt', elt'':Type
m:t elt
x:key

Raw.bst m
apply (is_bst m). Qed.
elt, elt', elt'':Type

forall (m : t elt) (x : key), mem x m = true -> In x m
elt, elt', elt'':Type

forall (m : t elt) (x : key), mem x m = true -> In x m
unfold In, mem; intros m x; rewrite In_alt; simpl; apply mem_2; auto. Qed.
elt, elt', elt'':Type

Empty empty
elt, elt', elt'':Type

Empty empty
exact (@empty_1 elt). Qed.
elt, elt', elt'':Type

forall m : t elt, Empty m -> is_empty m = true
elt, elt', elt'':Type

forall m : t elt, Empty m -> is_empty m = true
intros m; exact (@is_empty_1 _ (this m)). Qed.
elt, elt', elt'':Type

forall m : t elt, is_empty m = true -> Empty m
elt, elt', elt'':Type

forall m : t elt, is_empty m = true -> Empty m
intros m; exact (@is_empty_2 _ (this m)). Qed.
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e : elt), E.eq x y -> MapsTo y e (add x e m)
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e : elt), E.eq x y -> MapsTo y e (add x e m)
intros m x y e; exact (@add_1 elt _ x y e). Qed.
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e e' : elt), ~ E.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m)
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e e' : elt), ~ E.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m)
intros m x y e e'; exact (@add_2 elt _ x y e e'). Qed.
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e e' : elt), ~ E.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e m
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e e' : elt), ~ E.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e m
intros m x y e e'; exact (@add_3 elt _ x y e e'). Qed.
elt, elt', elt'':Type

forall (m : t elt) (x y : key), E.eq x y -> ~ In y (remove x m)
elt, elt', elt'':Type

forall (m : t elt) (x y : key), E.eq x y -> ~ In y (remove x m)
elt, elt', elt'':Type
m:t elt
x, y:key

Raw.bst m
apply (is_bst m). Qed.
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e : elt), ~ E.eq x y -> MapsTo y e m -> MapsTo y e (remove x m)
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e : elt), ~ E.eq x y -> MapsTo y e m -> MapsTo y e (remove x m)
intros m x y e; exact (@remove_2 elt _ x y e (is_bst m)). Qed.
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e : elt), MapsTo y e (remove x m) -> MapsTo y e m
elt, elt', elt'':Type

forall (m : t elt) (x y : key) (e : elt), MapsTo y e (remove x m) -> MapsTo y e m
intros m x y e; exact (@remove_3 elt _ x y e (is_bst m)). Qed.
elt, elt', elt'':Type

forall (m : t elt) (x : key) (e : elt), MapsTo x e m -> find x m = Some e
elt, elt', elt'':Type

forall (m : t elt) (x : key) (e : elt), MapsTo x e m -> find x m = Some e
intros m x e; exact (@find_1 elt _ x e (is_bst m)). Qed.
elt, elt', elt'':Type

forall (m : t elt) (x : key) (e : elt), find x m = Some e -> MapsTo x e m
elt, elt', elt'':Type

forall (m : t elt) (x : key) (e : elt), find x m = Some e -> MapsTo x e m
intros m; exact (@find_2 elt (this m)). Qed.
elt, elt', elt'':Type

forall (m : t elt) (A : Type) (i : A) (f : key -> elt -> A -> A), fold f m i = fold_left (fun (a : A) (p : key * elt) => f (fst p) (snd p) a) (elements m) i
elt, elt', elt'':Type

forall (m : t elt) (A : Type) (i : A) (f : key -> elt -> A -> A), fold f m i = fold_left (fun (a : A) (p : key * elt) => f (fst p) (snd p) a) (elements m) i
intros m; exact (@fold_1 elt (this m) (is_bst m)). Qed.
elt, elt', elt'':Type

forall (m : t elt) (x : key) (e : elt), MapsTo x e m -> InA eq_key_elt (x, e) (elements m)
elt, elt', elt'':Type

forall (m : t elt) (x : key) (e : elt), MapsTo x e m -> InA eq_key_elt (x, e) (elements m)
intros; unfold elements, MapsTo, eq_key_elt; rewrite elements_mapsto; auto. Qed.
elt, elt', elt'':Type

forall (m : t elt) (x : key) (e : elt), InA eq_key_elt (x, e) (elements m) -> MapsTo x e m
elt, elt', elt'':Type

forall (m : t elt) (x : key) (e : elt), InA eq_key_elt (x, e) (elements m) -> MapsTo x e m
intros; unfold elements, MapsTo, eq_key_elt; rewrite <- elements_mapsto; auto. Qed.
elt, elt', elt'':Type

forall m : t elt, Sorted lt_key (elements m)
elt, elt', elt'':Type

forall m : t elt, Sorted lt_key (elements m)
intros m; exact (@elements_sort elt (this m) (is_bst m)). Qed.
elt, elt', elt'':Type

forall m : t elt, NoDupA eq_key (elements m)
elt, elt', elt'':Type

forall m : t elt, NoDupA eq_key (elements m)
intros m; exact (@elements_nodup elt (this m) (is_bst m)). Qed.
elt, elt', elt'':Type

forall m : t elt, cardinal m = length (elements m)
elt, elt', elt'':Type

forall m : t elt, cardinal m = length (elements m)
intro m; exact (@elements_cardinal elt (this m)). Qed. Definition Equal m m' := forall y, find y m = find y m'. Definition Equiv (eq_elt:elt->elt->Prop) m m' := (forall k, In k m <-> In k m') /\ (forall k e e', MapsTo k e m -> MapsTo k e' m' -> eq_elt e e'). Definition Equivb cmp := Equiv (Cmp cmp).
elt, elt', elt'':Type

forall (cmp : elt -> elt -> bool) (m m' : t elt), Equivb cmp m m' <-> Raw.Proofs.Equivb cmp m m'
elt, elt', elt'':Type

forall (cmp : elt -> elt -> bool) (m m' : t elt), Equivb cmp m m' <-> Raw.Proofs.Equivb cmp m m'
elt, elt', elt'':Type
cmp:elt -> elt -> bool
m, m':t elt

(forall k : key, Raw.In0 k m <-> Raw.In0 k m') /\ (forall (k : key) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> Cmp cmp e e') <-> (forall k : Raw.key, Raw.In k m <-> Raw.In k m') /\ (forall (k : Raw.key) (e e' : elt), Raw.MapsTo k e m -> Raw.MapsTo k e' m' -> cmp e e' = true)
elt, elt', elt'':Type
cmp:elt -> elt -> bool
m, m':t elt
H0:forall k0 : key, Raw.In0 k0 m <-> Raw.In0 k0 m'
H1:forall (k0 : key) (e e' : elt), MapsTo k0 e m -> MapsTo k0 e' m' -> Cmp cmp e e'
k:Raw.key
H:Raw.In k m

Raw.In k m'
elt, elt', elt'':Type
cmp:elt -> elt -> bool
m, m':t elt
H0:forall k0 : key, Raw.In0 k0 m <-> Raw.In0 k0 m'
H1:forall (k0 : key) (e e' : elt), MapsTo k0 e m -> MapsTo k0 e' m' -> Cmp cmp e e'
k:Raw.key
H:Raw.In k m'
Raw.In k m
elt, elt', elt'':Type
cmp:elt -> elt -> bool
m, m':t elt
H0:forall k0 : Raw.key, Raw.In k0 m <-> Raw.In k0 m'
H1:forall (k0 : Raw.key) (e e' : elt), Raw.MapsTo k0 e m -> Raw.MapsTo k0 e' m' -> cmp e e' = true
k:key
H:Raw.In0 k m
Raw.In0 k m'
elt, elt', elt'':Type
cmp:elt -> elt -> bool
m, m':t elt
H0:forall k0 : Raw.key, Raw.In k0 m <-> Raw.In k0 m'
H1:forall (k0 : Raw.key) (e e' : elt), Raw.MapsTo k0 e m -> Raw.MapsTo k0 e' m' -> cmp e e' = true
k:key
H:Raw.In0 k m'
Raw.In0 k m
elt, elt', elt'':Type
cmp:elt -> elt -> bool
m, m':t elt
H0:forall k0 : key, Raw.In0 k0 m <-> Raw.In0 k0 m'
H1:forall (k0 : key) (e e' : elt), MapsTo k0 e m -> MapsTo k0 e' m' -> Cmp cmp e e'
k:Raw.key
H:Raw.In k m'

Raw.In k m
elt, elt', elt'':Type
cmp:elt -> elt -> bool
m, m':t elt
H0:forall k0 : Raw.key, Raw.In k0 m <-> Raw.In k0 m'
H1:forall (k0 : Raw.key) (e e' : elt), Raw.MapsTo k0 e m -> Raw.MapsTo k0 e' m' -> cmp e e' = true
k:key
H:Raw.In0 k m
Raw.In0 k m'
elt, elt', elt'':Type
cmp:elt -> elt -> bool
m, m':t elt
H0:forall k0 : Raw.key, Raw.In k0 m <-> Raw.In k0 m'
H1:forall (k0 : Raw.key) (e e' : elt), Raw.MapsTo k0 e m -> Raw.MapsTo k0 e' m' -> cmp e e' = true
k:key
H:Raw.In0 k m'
Raw.In0 k m
elt, elt', elt'':Type
cmp:elt -> elt -> bool
m, m':t elt
H0:forall k0 : Raw.key, Raw.In k0 m <-> Raw.In k0 m'
H1:forall (k0 : Raw.key) (e e' : elt), Raw.MapsTo k0 e m -> Raw.MapsTo k0 e' m' -> cmp e e' = true
k:key
H:Raw.In0 k m

Raw.In0 k m'
elt, elt', elt'':Type
cmp:elt -> elt -> bool
m, m':t elt
H0:forall k0 : Raw.key, Raw.In k0 m <-> Raw.In k0 m'
H1:forall (k0 : Raw.key) (e e' : elt), Raw.MapsTo k0 e m -> Raw.MapsTo k0 e' m' -> cmp e e' = true
k:key
H:Raw.In0 k m'
Raw.In0 k m
elt, elt', elt'':Type
cmp:elt -> elt -> bool
m, m':t elt
H0:forall k0 : Raw.key, Raw.In k0 m <-> Raw.In k0 m'
H1:forall (k0 : Raw.key) (e e' : elt), Raw.MapsTo k0 e m -> Raw.MapsTo k0 e' m' -> cmp e e' = true
k:key
H:Raw.In0 k m'

Raw.In0 k m
generalize (H0 k); do 2 rewrite <- In_alt; intuition. Qed.
elt, elt', elt'':Type

forall (m m' : t elt) (cmp : elt -> elt -> bool), Equivb cmp m m' -> equal cmp m m' = true
elt, elt', elt'':Type

forall (m m' : t elt) (cmp : elt -> elt -> bool), Equivb cmp m m' -> equal cmp m m' = true
unfold equal; intros (m,b) (m',b') cmp; rewrite Equivb_Equivb; intros; simpl in *; rewrite equal_Equivb; auto. Qed.
elt, elt', elt'':Type

forall (m m' : t elt) (cmp : elt -> elt -> bool), equal cmp m m' = true -> Equivb cmp m m'
elt, elt', elt'':Type

forall (m m' : t elt) (cmp : elt -> elt -> bool), equal cmp m m' = true -> Equivb cmp m m'
unfold equal; intros (m,b) (m',b') cmp; rewrite Equivb_Equivb; intros; simpl in *; rewrite <-equal_Equivb; auto. Qed. End Elt.

forall (elt elt' : Type) (m : t elt) (x : key) (e : elt) (f : elt -> elt'), MapsTo x e m -> MapsTo x (f e) (map f m)

forall (elt elt' : Type) (m : t elt) (x : key) (e : elt) (f : elt -> elt'), MapsTo x e m -> MapsTo x (f e) (map f m)
intros elt elt' m x e f; exact (@map_1 elt elt' f (this m) x e). Qed.

forall (elt elt' : Type) (m : t elt) (x : key) (f : elt -> elt'), In (elt:=elt') x (map f m) -> In (elt:=elt) x m

forall (elt elt' : Type) (m : t elt) (x : key) (f : elt -> elt'), In (elt:=elt') x (map f m) -> In (elt:=elt) x m
elt, elt':Type
m:t elt
x:key
f:elt -> elt'

Raw.In x (Raw.map f m) -> Raw.In x m
apply map_2; auto. Qed.

forall (elt elt' : Type) (m : t elt) (x : key) (e : elt) (f : key -> elt -> elt'), MapsTo x e m -> exists y : E.t, E.eq y x /\ MapsTo x (f y e) (mapi f m)

forall (elt elt' : Type) (m : t elt) (x : key) (e : elt) (f : key -> elt -> elt'), MapsTo x e m -> exists y : E.t, E.eq y x /\ MapsTo x (f y e) (mapi f m)
intros elt elt' m x e f; exact (@mapi_1 elt elt' f (this m) x e). Qed.

forall (elt elt' : Type) (m : t elt) (x : key) (f : key -> elt -> elt'), In (elt:=elt') x (mapi f m) -> In (elt:=elt) x m

forall (elt elt' : Type) (m : t elt) (x : key) (f : key -> elt -> elt'), In (elt:=elt') x (mapi f m) -> In (elt:=elt) x m
intros elt elt' m x f; unfold In in *; do 2 rewrite In_alt; simpl; apply mapi_2; auto. Qed.

forall (elt elt' elt'' : Type) (m : t elt) (m' : t elt') (x : key) (f : option elt -> option elt' -> option elt''), In (elt:=elt) x m \/ In (elt:=elt') x m' -> find (elt:=elt'') x (map2 f m m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')

forall (elt elt' elt'' : Type) (m : t elt) (m' : t elt') (x : key) (f : option elt -> option elt' -> option elt''), In (elt:=elt) x m \/ In (elt:=elt') x m' -> find (elt:=elt'') x (map2 f m m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')
elt, elt', elt'':Type
m:t elt
m':t elt'
x:key
f:option elt -> option elt' -> option elt''

Raw.In0 x m \/ Raw.In0 x m' -> Raw.find x {| this := Raw.map2 f m m'; is_bst := map2_bst f (is_bst m) (is_bst m') |} = f (Raw.find x m) (Raw.find x m')
elt, elt', elt'':Type
m:t elt
m':t elt'
x:key
f:option elt -> option elt' -> option elt''
H:Raw.In x m \/ Raw.In x m'

Raw.bst m
elt, elt', elt'':Type
m:t elt
m':t elt'
x:key
f:option elt -> option elt' -> option elt''
H:Raw.In x m \/ Raw.In x m'
Raw.bst m'
elt, elt', elt'':Type
m:t elt
m':t elt'
x:key
f:option elt -> option elt' -> option elt''
H:Raw.In x m \/ Raw.In x m'

Raw.bst m'
apply (is_bst m'). Qed.

forall (elt elt' elt'' : Type) (m : t elt) (m' : t elt') (x : key) (f : option elt -> option elt' -> option elt''), In (elt:=elt'') x (map2 f m m') -> In (elt:=elt) x m \/ In (elt:=elt') x m'

forall (elt elt' elt'' : Type) (m : t elt) (m' : t elt') (x : key) (f : option elt -> option elt' -> option elt''), In (elt:=elt'') x (map2 f m m') -> In (elt:=elt) x m \/ In (elt:=elt') x m'
elt, elt', elt'':Type
m:t elt
m':t elt'
x:key
f:option elt -> option elt' -> option elt''

Raw.In0 x {| this := Raw.map2 f m m'; is_bst := map2_bst f (is_bst m) (is_bst m') |} -> Raw.In0 x m \/ Raw.In0 x m'
elt, elt', elt'':Type
m:t elt
m':t elt'
x:key
f:option elt -> option elt' -> option elt''
H:Raw.In x (Raw.map2 f m m')

Raw.bst m
elt, elt', elt'':Type
m:t elt
m':t elt'
x:key
f:option elt -> option elt' -> option elt''
H:Raw.In x (Raw.map2 f m m')
Raw.bst m'
elt, elt', elt'':Type
m:t elt
m':t elt'
x:key
f:option elt -> option elt' -> option elt''
H:Raw.In x (Raw.map2 f m m')

Raw.bst m'
apply (is_bst m'). Qed. End IntMake. Module IntMake_ord (I:Int)(X: OrderedType)(D : OrderedType) <: Sord with Module Data := D with Module MapS.E := X. Module Data := D. Module Import MapS := IntMake(I)(X). Module LO := FMapList.Make_ord(X)(D). Module R := Raw. Module P := Raw.Proofs. Definition t := MapS.t D.t. Definition cmp e e' := match D.compare e e' with EQ _ => true | _ => false end.
One step of comparison of elements
  Definition compare_more x1 d1 (cont:R.enumeration D.t -> comparison) e2 :=
   match e2 with
    | R.End _ => Gt
    | R.More x2 d2 r2 e2 =>
       match X.compare x1 x2 with
        | EQ _ => match D.compare d1 d2 with
                   | EQ _ => cont (R.cons r2 e2)
                   | LT _ => Lt
                   | GT _ => Gt
                  end
        | LT _ => Lt
        | GT _ => Gt
       end
   end.
Comparison of left tree, middle element, then right tree
  Fixpoint compare_cont s1 (cont:R.enumeration D.t -> comparison) e2 :=
   match s1 with
    | R.Leaf _ => cont e2
    | R.Node l1 x1 d1 r1 _ =>
       compare_cont l1 (compare_more x1 d1 (compare_cont r1 cont)) e2
   end.
Initial continuation
  Definition compare_end (e2:R.enumeration D.t) :=
   match e2 with R.End _ => Eq | _ => Lt end.
The complete comparison
  Definition compare_pure s1 s2 :=
   compare_cont s1 compare_end (R.cons s2 (Raw.End _)).
Correctness of this comparison
  Definition Cmp c :=
   match c with
    | Eq => LO.eq_list
    | Lt => LO.lt_list
    | Gt => (fun l1 l2 => LO.lt_list l2 l1)
   end.

  

forall (c : comparison) (x1 x2 : X.t) (d1 d2 : D.t) (l1 l2 : list (X.t * D.t)), X.eq x1 x2 -> D.eq d1 d2 -> Cmp c l1 l2 -> Cmp c ((x1, d1) :: l1) ((x2, d2) :: l2)

forall (c : comparison) (x1 x2 : X.t) (d1 d2 : D.t) (l1 l2 : list (X.t * D.t)), X.eq x1 x2 -> D.eq d1 d2 -> Cmp c l1 l2 -> Cmp c ((x1, d1) :: l1) ((x2, d2) :: l2)
destruct c; simpl; intros; P.MX.elim_comp; auto. Qed. Hint Resolve cons_Cmp : core.

forall e2 : R.enumeration D.t, Cmp (compare_end e2) nil (P.flatten_e e2)

forall e2 : R.enumeration D.t, Cmp (compare_end e2) nil (P.flatten_e e2)
destruct e2; simpl; auto. Qed.

forall (x1 : X.t) (d1 : D.t) (cont : R.enumeration D.t -> comparison) (x2 : R.key) (d2 : D.t) (r2 : Raw.t D.t) (e2 : R.enumeration D.t) (l : list (X.t * D.t)), Cmp (cont (R.cons r2 e2)) l (R.elements r2 ++ P.flatten_e e2) -> Cmp (compare_more x1 d1 cont (R.More x2 d2 r2 e2)) ((x1, d1) :: l) (P.flatten_e (R.More x2 d2 r2 e2))

forall (x1 : X.t) (d1 : D.t) (cont : R.enumeration D.t -> comparison) (x2 : R.key) (d2 : D.t) (r2 : Raw.t D.t) (e2 : R.enumeration D.t) (l : list (X.t * D.t)), Cmp (cont (R.cons r2 e2)) l (R.elements r2 ++ P.flatten_e e2) -> Cmp (compare_more x1 d1 cont (R.More x2 d2 r2 e2)) ((x1, d1) :: l) (P.flatten_e (R.More x2 d2 r2 e2))
simpl; intros; destruct X.compare; simpl; try destruct D.compare; simpl; auto; P.MX.elim_comp; auto. Qed.

forall (s1 : Raw.t D.t) (cont : Raw.enumeration D.t -> comparison) (e2 : R.enumeration D.t) (l : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont e) l (P.flatten_e e)) -> Cmp (compare_cont s1 cont e2) (R.elements s1 ++ l) (P.flatten_e e2)

forall (s1 : Raw.t D.t) (cont : Raw.enumeration D.t -> comparison) (e2 : R.enumeration D.t) (l : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont e) l (P.flatten_e e)) -> Cmp (compare_cont s1 cont e2) (R.elements s1 ++ l) (P.flatten_e e2)
l1:Raw.t D.t
x1:Raw.key
d1:D.t
r1:Raw.t D.t
h1:I.t
Hl1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont l1 cont0 e0) (R.elements l1 ++ l0) (P.flatten_e e0)
Hr1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont r1 cont0 e0) (R.elements r1 ++ l0) (P.flatten_e e0)
cont:Raw.enumeration D.t -> comparison
e2:R.enumeration D.t
l:list (X.t * D.t)
H:forall e : Raw.enumeration D.t, Cmp (cont e) l (P.flatten_e e)

Cmp (compare_cont (Raw.Node l1 x1 d1 r1 h1) cont e2) (R.elements (Raw.Node l1 x1 d1 r1 h1) ++ l) (P.flatten_e e2)
l1:Raw.t D.t
x1:Raw.key
d1:D.t
r1:Raw.t D.t
h1:I.t
Hl1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont l1 cont0 e0) (R.elements l1 ++ l0) (P.flatten_e e0)
Hr1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont r1 cont0 e0) (R.elements r1 ++ l0) (P.flatten_e e0)
cont:Raw.enumeration D.t -> comparison
e2:R.enumeration D.t
l:list (X.t * D.t)
H:forall e : Raw.enumeration D.t, Cmp (cont e) l (P.flatten_e e)

Cmp (compare_cont l1 (compare_more x1 d1 (compare_cont r1 cont)) e2) (Raw.elements l1 ++ (x1, d1) :: Raw.elements r1 ++ l) (P.flatten_e e2)
l1:Raw.t D.t
x1:Raw.key
d1:D.t
r1:Raw.t D.t
h1:I.t
Hl1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont l1 cont0 e0) (R.elements l1 ++ l0) (P.flatten_e e0)
Hr1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont r1 cont0 e0) (R.elements r1 ++ l0) (P.flatten_e e0)
cont:Raw.enumeration D.t -> comparison
e2:R.enumeration D.t
l:list (X.t * D.t)
H:forall e : Raw.enumeration D.t, Cmp (cont e) l (P.flatten_e e)

forall e : Raw.enumeration D.t, Cmp (compare_more x1 d1 (compare_cont r1 cont) e) ((x1, d1) :: Raw.elements r1 ++ l) (P.flatten_e e)
l1:Raw.t D.t
x1:Raw.key
d1:D.t
r1:Raw.t D.t
h1:I.t
Hl1:forall (cont0 : Raw.enumeration D.t -> comparison) (e2 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont l1 cont0 e2) (R.elements l1 ++ l0) (P.flatten_e e2)
Hr1:forall (cont0 : Raw.enumeration D.t -> comparison) (e2 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont r1 cont0 e2) (R.elements r1 ++ l0) (P.flatten_e e2)
cont:Raw.enumeration D.t -> comparison
l:list (X.t * D.t)
H:forall e : Raw.enumeration D.t, Cmp (cont e) l (P.flatten_e e)

forall e : Raw.enumeration D.t, Cmp (compare_more x1 d1 (compare_cont r1 cont) e) ((x1, d1) :: Raw.elements r1 ++ l) (P.flatten_e e)
l1:Raw.t D.t
x1:Raw.key
d1:D.t
r1:Raw.t D.t
h1:I.t
Hl1:forall (cont0 : Raw.enumeration D.t -> comparison) (e2 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont l1 cont0 e2) (R.elements l1 ++ l0) (P.flatten_e e2)
Hr1:forall (cont0 : Raw.enumeration D.t -> comparison) (e2 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont r1 cont0 e2) (R.elements r1 ++ l0) (P.flatten_e e2)
cont:Raw.enumeration D.t -> comparison
l:list (X.t * D.t)
H:forall e : Raw.enumeration D.t, Cmp (cont e) l (P.flatten_e e)

Cmp (compare_more x1 d1 (compare_cont r1 cont) (Raw.End D.t)) ((x1, d1) :: Raw.elements r1 ++ l) (P.flatten_e (Raw.End D.t))
l1:Raw.t D.t
x1:Raw.key
d1:D.t
r1:Raw.t D.t
h1:I.t
Hl1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont l1 cont0 e0) (R.elements l1 ++ l0) (P.flatten_e e0)
Hr1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont r1 cont0 e0) (R.elements r1 ++ l0) (P.flatten_e e0)
cont:Raw.enumeration D.t -> comparison
l:list (X.t * D.t)
H:forall e : Raw.enumeration D.t, Cmp (cont e) l (P.flatten_e e)
x2:Raw.key
d2:D.t
r2:Raw.t D.t
e2:Raw.enumeration D.t
Cmp (compare_more x1 d1 (compare_cont r1 cont) (Raw.More x2 d2 r2 e2)) ((x1, d1) :: Raw.elements r1 ++ l) (P.flatten_e (Raw.More x2 d2 r2 e2))
l1:Raw.t D.t
x1:Raw.key
d1:D.t
r1:Raw.t D.t
h1:I.t
Hl1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont l1 cont0 e0) (R.elements l1 ++ l0) (P.flatten_e e0)
Hr1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont r1 cont0 e0) (R.elements r1 ++ l0) (P.flatten_e e0)
cont:Raw.enumeration D.t -> comparison
l:list (X.t * D.t)
H:forall e : Raw.enumeration D.t, Cmp (cont e) l (P.flatten_e e)
x2:Raw.key
d2:D.t
r2:Raw.t D.t
e2:Raw.enumeration D.t

Cmp (compare_more x1 d1 (compare_cont r1 cont) (Raw.More x2 d2 r2 e2)) ((x1, d1) :: Raw.elements r1 ++ l) (P.flatten_e (Raw.More x2 d2 r2 e2))
l1:Raw.t D.t
x1:Raw.key
d1:D.t
r1:Raw.t D.t
h1:I.t
Hl1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont l1 cont0 e0) (R.elements l1 ++ l0) (P.flatten_e e0)
Hr1:forall (cont0 : Raw.enumeration D.t -> comparison) (e0 : R.enumeration D.t) (l0 : list (X.t * D.t)), (forall e : Raw.enumeration D.t, Cmp (cont0 e) l0 (P.flatten_e e)) -> Cmp (compare_cont r1 cont0 e0) (R.elements r1 ++ l0) (P.flatten_e e0)
cont:Raw.enumeration D.t -> comparison
l:list (X.t * D.t)
H:forall e : Raw.enumeration D.t, Cmp (cont e) l (P.flatten_e e)
x2:Raw.key
d2:D.t
r2:Raw.t D.t
e2:Raw.enumeration D.t

Cmp (compare_cont r1 cont (R.cons r2 e2)) (Raw.elements r1 ++ l) (R.elements r2 ++ P.flatten_e e2)
rewrite <- P.cons_1; auto. Qed.

forall s1 s2 : Raw.t D.t, Cmp (compare_pure s1 s2) (R.elements s1) (R.elements s2)

forall s1 s2 : Raw.t D.t, Cmp (compare_pure s1 s2) (R.elements s1) (R.elements s2)
s1, s2:Raw.t D.t

Cmp (compare_cont s1 compare_end (R.cons s2 (Raw.End D.t))) (R.elements s1) (R.elements s2)
s1, s2:Raw.t D.t

Cmp (compare_cont s1 compare_end (R.cons s2 (Raw.End D.t))) (R.elements s1 ++ nil) (R.elements s2)
s1, s2:Raw.t D.t

Cmp (compare_cont s1 compare_end (R.cons s2 (Raw.End D.t))) (R.elements s1 ++ nil) (P.flatten_e (R.cons s2 (R.End D.t)))
auto using compare_cont_Cmp, compare_end_Cmp. Qed.
The dependent-style compare
  Definition eq (m1 m2 : t) := LO.eq_list (elements m1) (elements m2).
  Definition lt (m1 m2 : t) := LO.lt_list (elements m1) (elements m2).

  
s, s':t

Compare lt eq s s'
s, s':t

Compare lt eq s s'
s:Raw.t D.t
b:Raw.bst s
s':Raw.t D.t
b':Raw.bst s'

Compare lt eq {| this := s; is_bst := b |} {| this := s'; is_bst := b' |}
s:Raw.t D.t
b:Raw.bst s
s':Raw.t D.t
b':Raw.bst s'

Cmp (compare_pure s s') (R.elements s) (R.elements s') -> Compare lt eq {| this := s; is_bst := b |} {| this := s'; is_bst := b' |}
destruct compare_pure; intros; [apply EQ|apply LT|apply GT]; red; auto. Defined. (* Proofs about [eq] and [lt] *) Definition selements (m1 : t) := LO.MapS.Build_slist (P.elements_sort (is_bst m1)). Definition seq (m1 m2 : t) := LO.eq (selements m1) (selements m2). Definition slt (m1 m2 : t) := LO.lt (selements m1) (selements m2).

forall m1 m2 : t, eq m1 m2 <-> seq m1 m2

forall m1 m2 : t, eq m1 m2 <-> seq m1 m2
unfold eq, seq, selements, elements, LO.eq; intuition. Qed.

forall m1 m2 : t, lt m1 m2 <-> slt m1 m2

forall m1 m2 : t, lt m1 m2 <-> slt m1 m2
unfold lt, slt, selements, elements, LO.lt; intuition. Qed.

forall m m' : t, Equivb cmp m m' -> eq m m'

forall m m' : t, Equivb cmp m m' -> eq m m'
m, m':t

Equivb cmp m m' -> eq m m'
m, m':t

Equivb cmp m m' -> LO.eq (selements m) (selements m')
m, m':t

Raw.Proofs.Equivb cmp m m' -> LO.eq (selements m) (selements m')
m, m':t

P.L.Equivb cmp (Raw.elements m) (Raw.elements m') -> LO.eq (selements m) (selements m')
auto using LO.eq_1. Qed.

forall m m' : t, eq m m' -> Equivb cmp m m'

forall m m' : t, eq m m' -> Equivb cmp m m'
m, m':t

eq m m' -> Equivb cmp m m'
m, m':t

LO.eq (selements m) (selements m') -> Equivb cmp m m'
m, m':t

LO.eq (selements m) (selements m') -> Raw.Proofs.Equivb cmp m m'
m, m':t

LO.eq (selements m) (selements m') -> P.L.Equivb cmp (Raw.elements m) (Raw.elements m')
m, m':t
H:LO.eq (selements m) (selements m')

P.L.Equivb cmp (Raw.elements m) (Raw.elements m')
m, m':t
H:LO.eq (selements m) (selements m')

LO.MapS.Equivb LO.cmp (selements m) (selements m') -> P.L.Equivb cmp (Raw.elements m) (Raw.elements m')
auto. Qed.

forall m : t, eq m m

forall m : t, eq m m
intros; rewrite eq_seq; unfold seq; intros; apply LO.eq_refl. Qed.

forall m1 m2 : t, eq m1 m2 -> eq m2 m1

forall m1 m2 : t, eq m1 m2 -> eq m2 m1
intros m1 m2; rewrite 2 eq_seq; unfold seq; intros; apply LO.eq_sym; auto. Qed.

forall m1 m2 m3 : t, eq m1 m2 -> eq m2 m3 -> eq m1 m3

forall m1 m2 m3 : t, eq m1 m2 -> eq m2 m3 -> eq m1 m3
m1, m2, M3:t

LO.eq (selements m1) (selements m2) -> LO.eq (selements m2) (selements M3) -> LO.eq (selements m1) (selements M3)
intros; eapply LO.eq_trans; eauto. Qed.

forall m1 m2 m3 : t, lt m1 m2 -> lt m2 m3 -> lt m1 m3

forall m1 m2 m3 : t, lt m1 m2 -> lt m2 m3 -> lt m1 m3
intros m1 m2 m3; rewrite 3 lt_slt; unfold slt; intros; eapply LO.lt_trans; eauto. Qed.

forall m1 m2 : t, lt m1 m2 -> ~ eq m1 m2

forall m1 m2 : t, lt m1 m2 -> ~ eq m1 m2
intros m1 m2; rewrite lt_slt, eq_seq; unfold slt, seq; intros; apply LO.lt_not_eq; auto. Qed. End IntMake_ord. (* For concrete use inside Coq, we propose an instantiation of [Int] by [Z]. *) Module Make (X: OrderedType) <: S with Module E := X :=IntMake(Z_as_Int)(X). Module Make_ord (X: OrderedType)(D: OrderedType) <: Sord with Module Data := D with Module MapS.E := X :=IntMake_ord(Z_as_Int)(X)(D).