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)         *)
(************************************************************************)

MSetRBT : Implementation of MSetInterface via Red-Black trees

Initial author: Andrew W. Appel, 2011. Extra modifications by: Pierre Letouzey
The design decisions behind this implementation are described here:
Additional suggested reading:
Require MSetGenTree.
Require Import Bool List BinPos Pnat Setoid SetoidList PeanoNat.
Local Open Scope list_scope.

(* For nicer extraction, we create induction principles
   only when needed *)
Local Unset Elimination Schemes.
An extra function not (yet?) in MSetInterface.S
Module Type MSetRemoveMin (Import M:MSetInterface.S).

 Parameter remove_min : t -> option (elt * t).

 Axiom remove_min_spec1 : forall s k s',
  remove_min s = Some (k,s') ->
   min_elt s = Some k /\ remove k s [=] s'.

 Axiom remove_min_spec2 : forall s, remove_min s = None -> Empty s.

End MSetRemoveMin.
The type of color annotation.
Inductive color := Red | Black.

Module Color.
 Definition t := color.
End Color.

Ops : the pure functions

Module Ops (X:Orders.OrderedType) <: MSetInterface.Ops X.

Generic trees instantiated with color

We reuse a generic definition of trees where the information parameter is a color. Functions like mem or fold are also provided by this generic functor.
Include MSetGenTree.Ops X Color.

Definition t := tree.
Notation Rd := (Node Red).
Notation Bk := (Node Black).

Basic tree

Definition singleton (k: elt) : tree := Bk Leaf k Leaf.

Changing root color

Definition makeBlack t :=
 match t with
 | Leaf => Leaf
 | Node _ a x b => Bk a x b
 end.

Definition makeRed t :=
 match t with
 | Leaf => Leaf
 | Node _ a x b => Rd a x b
 end.

Balancing

We adapt when one side is not a true red-black tree. Both sides have the same black depth.
Definition lbal l k r :=
 match l with
 | Rd (Rd a x b) y c => Rd (Bk a x b) y (Bk c k r)
 | Rd a x (Rd b y c) => Rd (Bk a x b) y (Bk c k r)
 | _ => Bk l k r
 end.

Definition rbal l k r :=
 match r with
 | Rd (Rd b y c) z d => Rd (Bk l k b) y (Bk c z d)
 | Rd b y (Rd c z d) => Rd (Bk l k b) y (Bk c z d)
 | _ => Bk l k r
 end.
A variant of rbal, with reverse pattern order. Is it really useful ? Should we always use it ?
Definition rbal' l k r :=
 match r with
 | Rd b y (Rd c z d) => Rd (Bk l k b) y (Bk c z d)
 | Rd (Rd b y c) z d => Rd (Bk l k b) y (Bk c z d)
 | _ => Bk l k r
 end.
Balancing with different black depth. One side is almost a red-black tree, while the other is a true red-black tree, but with black depth + 1. Used in deletion.
Definition lbalS l k r :=
 match l with
 | Rd a x b => Rd (Bk a x b) k r
 | _ =>
   match r with
   | Bk a y b => rbal' l k (Rd a y b)
   | Rd (Bk a y b) z c => Rd (Bk l k a) y (rbal' b z (makeRed c))
   | _ => Rd l k r (* impossible *)
   end
 end.

Definition rbalS l k r :=
 match r with
 | Rd b y c => Rd l k (Bk b y c)
 | _ =>
   match l with
   | Bk a x b => lbal (Rd a x b) k r
   | Rd a x (Bk b y c) => Rd (lbal (makeRed a) x b) y (Bk c k r)
   | _ => Rd l k r (* impossible *)
   end
 end.

Insertion

Fixpoint ins x s :=
 match s with
 | Leaf => Rd Leaf x Leaf
 | Node c l y r =>
   match X.compare x y with
   | Eq => s
   | Lt =>
     match c with
     | Red => Rd (ins x l) y r
     | Black => lbal (ins x l) y r
     end
   | Gt =>
     match c with
     | Red => Rd l y (ins x r)
     | Black => rbal l y (ins x r)
     end
   end
 end.

Definition add x s := makeBlack (ins x s).

Deletion

Fixpoint append (l:tree) : tree -> tree :=
 match l with
 | Leaf => fun r => r
 | Node lc ll lx lr =>
   fix append_l (r:tree) : tree :=
   match r with
   | Leaf => l
   | Node rc rl rx rr =>
     match lc, rc with
     | Red, Red =>
       let lrl := append lr rl in
       match lrl with
       | Rd lr' x rl' => Rd (Rd ll lx lr') x (Rd rl' rx rr)
       | _ => Rd ll lx (Rd lrl rx rr)
       end
     | Black, Black =>
       let lrl := append lr rl in
       match lrl with
       | Rd lr' x rl' => Rd (Bk ll lx lr') x (Bk rl' rx rr)
       | _ => lbalS ll lx (Bk lrl rx rr)
       end
     | Black, Red => Rd (append_l rl) rx rr
     | Red, Black => Rd ll lx (append lr r)
     end
   end
 end.

Fixpoint del x t :=
 match t with
 | Leaf => Leaf
 | Node _ a y b =>
   match X.compare x y with
   | Eq => append a b
   | Lt =>
     match a with
     | Bk _ _ _ => lbalS (del x a) y b
     | _ => Rd (del x a) y b
     end
   | Gt =>
     match b with
     | Bk _ _ _ => rbalS a y (del x b)
     | _ => Rd a y (del x b)
     end
   end
 end.

Definition remove x t := makeBlack (del x t).

Removing minimal element

Fixpoint delmin l x r : (elt * tree) :=
 match l with
 | Leaf => (x,r)
 | Node lc ll lx lr =>
   let (k,l') := delmin ll lx lr in
   match lc with
   | Black => (k, lbalS l' x r)
   | Red => (k, Rd l' x r)
   end
 end.

Definition remove_min t : option (elt * tree) :=
 match t with
 | Leaf => None
 | Node _ l x r =>
   let (k,t) := delmin l x r in
   Some (k, makeBlack t)
 end.

Tree-ification

We rebuild a tree of size if pred then n-1 else n as soon as the list l has enough elements
Definition bogus : tree * list elt := (Leaf, nil).

Notation treeify_t := (list elt -> tree * list elt).

Definition treeify_zero : treeify_t :=
 fun acc => (Leaf,acc).

Definition treeify_one : treeify_t :=
 fun acc => match acc with
 | x::acc => (Rd Leaf x Leaf, acc)
 | _ => bogus
 end.

Definition treeify_cont (f g : treeify_t) : treeify_t :=
 fun acc =>
 match f acc with
 | (l, x::acc) =>
   match g acc with
   | (r, acc) => (Bk l x r, acc)
   end
 | _ => bogus
 end.

Fixpoint treeify_aux (pred:bool)(n: positive) : treeify_t :=
 match n with
 | xH => if pred then treeify_zero else treeify_one
 | xO n => treeify_cont (treeify_aux pred n) (treeify_aux true n)
 | xI n => treeify_cont (treeify_aux false n) (treeify_aux pred n)
 end.

Fixpoint plength_aux (l:list elt)(p:positive) := match l with
 | nil => p
 | _::l => plength_aux l (Pos.succ p)
end.

Definition plength l := plength_aux l 1.

Definition treeify (l:list elt) :=
 fst (treeify_aux true (plength l) l).

Filtering

Fixpoint filter_aux (f: elt -> bool) s acc :=
 match s with
 | Leaf => acc
 | Node _ l k r =>
   let acc := filter_aux f r acc in
   if f k then filter_aux f l (k::acc)
   else filter_aux f l acc
 end.

Definition filter (f: elt -> bool) (s: t) : t :=
 treeify (filter_aux f s nil).

Fixpoint partition_aux (f: elt -> bool) s acc1 acc2 :=
 match s with
 | Leaf => (acc1,acc2)
 | Node _ sl k sr =>
   let (acc1, acc2) := partition_aux f sr acc1 acc2 in
   if f k then partition_aux f sl (k::acc1) acc2
   else partition_aux f sl acc1 (k::acc2)
 end.

Definition partition (f: elt -> bool) (s:t) : t*t :=
  let (ok,ko) := partition_aux f s nil nil in
  (treeify ok, treeify ko).

Union, intersection, difference

union of the elements of l1 and l2 into a third acc list.
Fixpoint union_list l1 : list elt -> list elt -> list elt :=
 match l1 with
 | nil => @rev_append _
 | x::l1' =>
    fix union_l1 l2 acc :=
    match l2 with
    | nil => rev_append l1 acc
    | y::l2' =>
       match X.compare x y with
       | Eq => union_list l1' l2' (x::acc)
       | Lt => union_l1 l2' (y::acc)
       | Gt => union_list l1' l2 (x::acc)
       end
    end
 end.

Definition linear_union s1 s2 :=
  treeify (union_list (rev_elements s1) (rev_elements s2) nil).

Fixpoint inter_list l1 : list elt -> list elt -> list elt :=
 match l1 with
 | nil => fun _ acc => acc
 | x::l1' =>
    fix inter_l1 l2 acc :=
    match l2 with
    | nil => acc
    | y::l2' =>
       match X.compare x y with
       | Eq => inter_list l1' l2' (x::acc)
       | Lt => inter_l1 l2' acc
       | Gt => inter_list l1' l2 acc
       end
    end
 end.

Definition linear_inter s1 s2 :=
  treeify (inter_list (rev_elements s1) (rev_elements s2) nil).

Fixpoint diff_list l1 : list elt -> list elt -> list elt :=
 match l1 with
 | nil => fun _ acc => acc
 | x::l1' =>
    fix diff_l1 l2 acc :=
    match l2 with
    | nil => rev_append l1 acc
    | y::l2' =>
       match X.compare x y with
       | Eq => diff_list l1' l2' acc
       | Lt => diff_l1 l2' acc
       | Gt => diff_list l1' l2 (x::acc)
       end
    end
 end.

Definition linear_diff s1 s2 :=
  treeify (diff_list (rev_elements s1) (rev_elements s2) nil).
compare_height returns: Warning: this is not an equivalence relation! but who cares....
Definition skip_red t :=
 match t with
 | Rd t' _ _ => t'
 | _ => t
 end.

Definition skip_black t :=
 match skip_red t with
 | Bk t' _ _ => t'
 | t' => t'
 end.

Fixpoint compare_height (s1x s1 s2 s2x: tree) : comparison :=
 match skip_red s1x, skip_red s1, skip_red s2, skip_red s2x with
 | Node _ s1x' _ _, Node _ s1' _ _, Node _ s2' _ _, Node _ s2x' _ _ =>
   compare_height (skip_black s1x') s1' s2' (skip_black s2x')
 | _, Leaf, _, Node _ _ _ _ => Lt
 | Node _ _ _ _, _, Leaf, _ => Gt
 | Node _ s1x' _ _, Node _ s1' _ _, Node _ s2' _ _, Leaf =>
   compare_height (skip_black s1x') s1' s2' Leaf
 | Leaf, Node _ s1' _ _, Node _ s2' _ _, Node _ s2x' _ _ =>
   compare_height Leaf s1'  s2'  (skip_black s2x')
 | _, _, _, _ => Eq
 end.
When one tree is quite smaller than the other, we simply adds repeatively all its elements in the big one. For trees of comparable height, we rather use linear_union.
Definition union (t1 t2: t) : t :=
 match compare_height t1 t1 t2 t2 with
 | Lt => fold add t1 t2
 | Gt => fold add t2 t1
 | Eq => linear_union t1 t2
 end.

Definition diff (t1 t2: t) : t :=
 match compare_height t1 t1 t2 t2 with
 | Lt => filter (fun k => negb (mem k t2)) t1
 | Gt => fold remove t2 t1
 | Eq => linear_diff t1 t2
 end.

Definition inter (t1 t2: t) : t :=
 match compare_height t1 t1 t2 t2 with
 | Lt => filter (fun k => mem k t2) t1
 | Gt => filter (fun k => mem k t1) t2
 | Eq => linear_inter t1 t2
 end.

End Ops.

MakeRaw : the pure functions and their specifications

Module Type MakeRaw (X:Orders.OrderedType) <: MSetInterface.RawSets X.
Include Ops X.
Generic definition of binary-search-trees and proofs of specifications for generic functions such as mem or fold.
Include MSetGenTree.Props X Color.

Notation Rd := (Node Red).
Notation Bk := (Node Black).

Local Hint Immediate MX.eq_sym : core.
Local Hint Unfold In lt_tree gt_tree Ok : core.
Local Hint Constructors InT bst : core.
Local Hint Resolve MX.eq_refl MX.eq_trans MX.lt_trans ok : core.
Local Hint Resolve lt_leaf gt_leaf lt_tree_node gt_tree_node : core.
Local Hint Resolve lt_tree_not_in lt_tree_trans gt_tree_not_in gt_tree_trans : core.
Local Hint Resolve elements_spec2 : core.

Singleton set

x, y:elt

InT y (singleton x) <-> X.eq y x
x, y:elt

InT y (singleton x) <-> X.eq y x
unfold singleton; intuition_in. Qed.
x:elt

Ok (singleton x)
x:elt

Ok (singleton x)
unfold singleton; auto. Qed.

makeBlack, MakeRed

s:tree
x:elt

InT x (makeBlack s) <-> InT x s
s:tree
x:elt

InT x (makeBlack s) <-> InT x s
destruct s; simpl; intuition_in. Qed.
s:tree
x:elt

InT x (makeRed s) <-> InT x s
s:tree
x:elt

InT x (makeRed s) <-> InT x s
destruct s; simpl; intuition_in. Qed.
s:tree
H:Ok s

Ok (makeBlack s)
s:tree
H:Ok s

Ok (makeBlack s)
destruct s; simpl; ok. Qed.
s:tree
H:Ok s

Ok (makeRed s)
s:tree
H:Ok s

Ok (makeRed s)
destruct s; simpl; ok. Qed.

Generic handling for red-matching and red-red-matching

Definition isblack t :=
 match t with Bk _ _ _ => True | _ => False end.

Definition notblack t :=
 match t with Bk _ _ _ => False | _ => True end.

Definition notred t :=
 match t with Rd _ _ _ => False | _ => True end.

Definition rcase {A} f g t : A :=
 match t with
 | Rd a x b => f a x b
 | _ => g t
 end.

Inductive rspec {A} f g : tree -> A -> Prop :=
 | rred a x b : rspec f g (Rd a x b) (f a x b)
 | relse t : notred t -> rspec f g t (g t).

A:Type
f:tree -> X.t -> tree -> A
g:tree -> A
t:tree

rspec f g t (rcase f g t)
A:Type
f:tree -> X.t -> tree -> A
g:tree -> A
t:tree

rspec f g t (rcase f g t)
destruct t as [|[|] l x r]; simpl; now constructor. Qed. Definition rrcase {A} f g t : A := match t with | Rd (Rd a x b) y c => f a x b y c | Rd a x (Rd b y c) => f a x b y c | _ => g t end. Notation notredred := (rrcase (fun _ _ _ _ _ => False) (fun _ => True)). Inductive rrspec {A} f g : tree -> A -> Prop := | rrleft a x b y c : rrspec f g (Rd (Rd a x b) y c) (f a x b y c) | rrright a x b y c : rrspec f g (Rd a x (Rd b y c)) (f a x b y c) | rrelse t : notredred t -> rrspec f g t (g t).
A:Type
f:tree -> X.t -> tree -> X.t -> tree -> A
g:tree -> A
t:tree

rrspec f g t (rrcase f g t)
A:Type
f:tree -> X.t -> tree -> X.t -> tree -> A
g:tree -> A
t:tree

rrspec f g t (rrcase f g t)
A:Type
f:tree -> X.t -> tree -> X.t -> tree -> A
g:tree -> A
l:tree
x:X.t
r:tree

rrspec f g (Rd l x r) match l with | Leaf => match r with | Rd b y c => f l x b y c | _ => g (Rd l x r) end | Rd a x0 b => f a x0 b x r | Bk _ _ _ => match r with | Rd b0 y c => f l x b0 y c | _ => g (Rd l x r) end end
destruct l as [|[|] ll lx lr], r as [|[|] rl rx rr]; now constructor. Qed. Definition rrcase' {A} f g t : A := match t with | Rd a x (Rd b y c) => f a x b y c | Rd (Rd a x b) y c => f a x b y c | _ => g t end.
A:Type
f:tree -> X.t -> tree -> X.t -> tree -> A
g:tree -> A
t:tree

rrspec f g t (rrcase' f g t)
A:Type
f:tree -> X.t -> tree -> X.t -> tree -> A
g:tree -> A
t:tree

rrspec f g t (rrcase' f g t)
A:Type
f:tree -> X.t -> tree -> X.t -> tree -> A
g:tree -> A
l:tree
x:X.t
r:tree

rrspec f g (Rd l x r) match l with | Leaf => match r with | Rd b y c => f l x b y c | _ => g (Rd l x r) end | Rd a x0 b => match r with | Rd b0 y c => f l x b0 y c | _ => f a x0 b x r end | Bk _ _ _ => match r with | Rd b0 y c => f l x b0 y c | _ => g (Rd l x r) end end
destruct l as [|[|] ll lx lr], r as [|[|] rl rx rr]; now constructor. Qed.
Balancing operations are instances of generic match
l:tree
k:X.t
r:tree

rrspec (fun (a : tree) (x : X.t) (b : tree) (y : X.t) (c : tree) => Rd (Bk a x b) y (Bk c k r)) (fun l0 : tree => Bk l0 k r) l (lbal l k r)
l:tree
k:X.t
r:tree

rrspec (fun (a : tree) (x : X.t) (b : tree) (y : X.t) (c : tree) => Rd (Bk a x b) y (Bk c k r)) (fun l0 : tree => Bk l0 k r) l (lbal l k r)
exact (rrmatch _ _ _). Qed.
l:tree
k:X.t
r:tree

rrspec (fun (a : tree) (x : X.t) (b : tree) (y : X.t) (c : tree) => Rd (Bk l k a) x (Bk b y c)) (fun r0 : tree => Bk l k r0) r (rbal l k r)
l:tree
k:X.t
r:tree

rrspec (fun (a : tree) (x : X.t) (b : tree) (y : X.t) (c : tree) => Rd (Bk l k a) x (Bk b y c)) (fun r0 : tree => Bk l k r0) r (rbal l k r)
exact (rrmatch _ _ _). Qed.
l:tree
k:X.t
r:tree

rrspec (fun (a : tree) (x : X.t) (b : tree) (y : X.t) (c : tree) => Rd (Bk l k a) x (Bk b y c)) (fun r0 : tree => Bk l k r0) r (rbal' l k r)
l:tree
k:X.t
r:tree

rrspec (fun (a : tree) (x : X.t) (b : tree) (y : X.t) (c : tree) => Rd (Bk l k a) x (Bk b y c)) (fun r0 : tree => Bk l k r0) r (rbal' l k r)
exact (rrmatch' _ _ _). Qed.
l:tree
x:X.t
r:tree

rspec (fun (a : tree) (y : X.t) (b : tree) => Rd (Bk a y b) x r) (fun l0 : tree => match r with | Rd (Bk a0 y b) z c => Rd (Bk l0 x a0) y (rbal' b z (makeRed c)) | Bk a z c => rbal' l0 x (Rd a z c) | _ => Rd l0 x r end) l (lbalS l x r)
l:tree
x:X.t
r:tree

rspec (fun (a : tree) (y : X.t) (b : tree) => Rd (Bk a y b) x r) (fun l0 : tree => match r with | Rd (Bk a0 y b) z c => Rd (Bk l0 x a0) y (rbal' b z (makeRed c)) | Bk a z c => rbal' l0 x (Rd a z c) | _ => Rd l0 x r end) l (lbalS l x r)
exact (rmatch _ _ _). Qed.
l:tree
x:X.t
r:tree

rspec (fun (a : tree) (y : X.t) (b : tree) => Rd l x (Bk a y b)) (fun r0 : tree => match l with | Leaf => Rd l x r0 | Rd a y Leaf | Rd a y (Rd _ _ _) => Rd l x r0 | Rd a y (Bk b0 z c) => Rd (lbal (makeRed a) y b0) z (Bk c x r0) | Bk a y b => lbal (Rd a y b) x r0 end) r (rbalS l x r)
l:tree
x:X.t
r:tree

rspec (fun (a : tree) (y : X.t) (b : tree) => Rd l x (Bk a y b)) (fun r0 : tree => match l with | Leaf => Rd l x r0 | Rd a y Leaf | Rd a y (Rd _ _ _) => Rd l x r0 | Rd a y (Bk b0 z c) => Rd (lbal (makeRed a) y b0) z (Bk c x r0) | Bk a y b => lbal (Rd a y b) x r0 end) r (rbalS l x r)
exact (rmatch _ _ _). Qed.

Balancing for insertion

l:tree
x:X.t
r:tree
y:elt

InT y (lbal l x r) <-> X.eq y x \/ InT y l \/ InT y r
l:tree
x:X.t
r:tree
y:elt

InT y (lbal l x r) <-> X.eq y x \/ InT y l \/ InT y r
case lbal_match; intuition_in. Qed.
l:tree
x:X.t
r:tree
H:Ok l
H0:Ok r
lt_tree0:lt_tree x l
gt_tree0:gt_tree x r

Ok (lbal l x r)
l:tree
x:X.t
r:tree
H:Ok l
H0:Ok r
lt_tree0:lt_tree x l
gt_tree0:gt_tree x r

Ok (lbal l x r)
destruct (lbal_match l x r); ok. Qed.
l:tree
x:X.t
r:tree
y:elt

InT y (rbal l x r) <-> X.eq y x \/ InT y l \/ InT y r
l:tree
x:X.t
r:tree
y:elt

InT y (rbal l x r) <-> X.eq y x \/ InT y l \/ InT y r
case rbal_match; intuition_in. Qed.
l:tree
x:X.t
r:tree
H:Ok l
H0:Ok r
lt_tree0:lt_tree x l
gt_tree0:gt_tree x r

Ok (rbal l x r)
l:tree
x:X.t
r:tree
H:Ok l
H0:Ok r
lt_tree0:lt_tree x l
gt_tree0:gt_tree x r

Ok (rbal l x r)
destruct (rbal_match l x r); ok. Qed.
l:tree
x:X.t
r:tree
y:elt

InT y (rbal' l x r) <-> X.eq y x \/ InT y l \/ InT y r
l:tree
x:X.t
r:tree
y:elt

InT y (rbal' l x r) <-> X.eq y x \/ InT y l \/ InT y r
case rbal'_match; intuition_in. Qed.
l:tree
x:X.t
r:tree
H:Ok l
H0:Ok r
lt_tree0:lt_tree x l
gt_tree0:gt_tree x r

Ok (rbal' l x r)
l:tree
x:X.t
r:tree
H:Ok l
H0:Ok r
lt_tree0:lt_tree x l
gt_tree0:gt_tree x r

Ok (rbal' l x r)
destruct (rbal'_match l x r); ok. Qed. Hint Rewrite In_node_iff In_leaf_iff makeRed_spec makeBlack_spec lbal_spec rbal_spec rbal'_spec : rb. Ltac descolor := destruct_all Color.t. Ltac destree t := destruct t as [|[|] ? ? ?]. Ltac autorew := autorewrite with rb. Tactic Notation "autorew" "in" ident(H) := autorewrite with rb in H.

Insertion


forall (s : tree) (x : X.t) (y : elt), InT y (ins x s) <-> X.eq y x \/ InT y s

forall (s : tree) (x : X.t) (y : elt), InT y (ins x s) <-> X.eq y x \/ InT y s
x:X.t
y:elt

InT y (Rd Leaf x Leaf) <-> X.eq y x \/ InT y Leaf
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 l
IHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 r
x:X.t
y:elt
H:X.eq x x'
InT y (Node i l x' r) <-> X.eq y x \/ InT y (Node i l x' r)
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 l
IHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 r
x:X.t
y:elt
H:X.lt x x'
InT y match i with | Red => Rd (ins x l) x' r | Black => lbal (ins x l) x' r end <-> X.eq y x \/ InT y (Node i l x' r)
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 l
IHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 r
x:X.t
y:elt
H:X.lt x' x
InT y match i with | Red => Rd l x' (ins x r) | Black => rbal l x' (ins x r) end <-> X.eq y x \/ InT y (Node i l x' r)
x:X.t
y:elt

InT y (Rd Leaf x Leaf) <-> X.eq y x \/ InT y Leaf
intuition_in.
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 l
IHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 r
x:X.t
y:elt
H:X.eq x x'

InT y (Node i l x' r) <-> X.eq y x \/ InT y (Node i l x' r)
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 l
IHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 r
x:X.t
y:elt
H:X.eq x x'
H1:X.eq y x

InT y (Node i l x' r)
setoid_replace y with x; eauto.
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 l
IHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 r
x:X.t
y:elt
H:X.lt x x'

InT y match i with | Red => Rd (ins x l) x' r | Black => lbal (ins x l) x' r end <-> X.eq y x \/ InT y (Node i l x' r)
descolor; autorew; rewrite IHl; intuition_in.
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 l
IHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 r
x:X.t
y:elt
H:X.lt x' x

InT y match i with | Red => Rd l x' (ins x r) | Black => rbal l x' (ins x r) end <-> X.eq y x \/ InT y (Node i l x' r)
descolor; autorew; rewrite IHr; intuition_in. Qed. Hint Rewrite ins_spec : rb.
s:tree
x:X.t
H:Ok s

Ok (ins x s)
s:tree
x:X.t
H:Ok s

Ok (ins x s)
induct s x; auto; descolor; (apply lbal_ok || apply rbal_ok || ok); auto; intros y; autorew; intuition; order. Qed.
s:tree
x:X.t
y:elt

InT y (add x s) <-> X.eq y x \/ InT y s
s:tree
x:X.t
y:elt

InT y (add x s) <-> X.eq y x \/ InT y s
s:tree
x:X.t
y:elt

InT y (makeBlack (ins x s)) <-> X.eq y x \/ InT y s
now autorew. Qed. Hint Rewrite add_spec' : rb.
s:tree
x:X.t
y:elt
H:Ok s

InT y (add x s) <-> X.eq y x \/ InT y s
s:tree
x:X.t
y:elt
H:Ok s

InT y (add x s) <-> X.eq y x \/ InT y s
apply add_spec'. Qed.
s:tree
x:X.t
H:Ok s

Ok (add x s)
s:tree
x:X.t
H:Ok s

Ok (add x s)
unfold add; auto_tc. Qed.

Balancing for deletion

l:tree
x:X.t
r:tree
y:elt

InT y (lbalS l x r) <-> X.eq y x \/ InT y l \/ InT y r
l:tree
x:X.t
r:tree
y:elt

InT y (lbalS l x r) <-> X.eq y x \/ InT y l \/ InT y r
l:tree
x:X.t
r:tree
y:elt

forall (a : tree) (x0 : X.t) (b : tree), InT y (Rd (Bk a x0 b) x r) <-> X.eq y x \/ InT y (Rd a x0 b) \/ InT y r
l:tree
x:X.t
r:tree
y:elt
forall t0 : tree, notred t0 -> InT y match r with | Rd (Bk a0 y0 b) z c => Rd (Bk t0 x a0) y0 (rbal' b z (makeRed c)) | Bk a z c => rbal' t0 x (Rd a z c) | _ => Rd t0 x r end <-> X.eq y x \/ InT y t0 \/ InT y r
l:tree
x:X.t
r:tree
y:elt

forall (a : tree) (x0 : X.t) (b : tree), InT y (Rd (Bk a x0 b) x r) <-> X.eq y x \/ InT y (Rd a x0 b) \/ InT y r
intros; autorew; intuition_in.
l:tree
x:X.t
r:tree
y:elt

forall t0 : tree, notred t0 -> InT y match r with | Rd (Bk a0 y0 b) z c => Rd (Bk t0 x a0) y0 (rbal' b z (makeRed c)) | Bk a z c => rbal' t0 x (Rd a z c) | _ => Rd t0 x r end <-> X.eq y x \/ InT y t0 \/ InT y r
x:X.t
r:tree
y:elt

forall t0 : tree, notred t0 -> InT y match r with | Rd (Bk a0 y0 b) z c => Rd (Bk t0 x a0) y0 (rbal' b z (makeRed c)) | Bk a z c => rbal' t0 x (Rd a z c) | _ => Rd t0 x r end <-> X.eq y x \/ InT y t0 \/ InT y r
x:X.t
r:tree
y:elt
l:tree

InT y match r with | Rd (Bk a0 y0 b) z c => Rd (Bk l x a0) y0 (rbal' b z (makeRed c)) | Bk a z c => rbal' l x (Rd a z c) | _ => Rd l x r end <-> X.eq y x \/ InT y l \/ InT y r
x:X.t
y:elt
l:tree

InT y (Rd l x Leaf) <-> X.eq y x \/ InT y l \/ InT y Leaf
x:X.t
rl:tree
rx:X.t
rr:tree
y:elt
l:tree
InT y match rl with | Bk a y0 b => Rd (Bk l x a) y0 (rbal' b rx (makeRed rr)) | _ => Rd l x (Rd rl rx rr) end <-> X.eq y x \/ InT y l \/ InT y (Rd rl rx rr)
x:X.t
rl:tree
rx:X.t
rr:tree
y:elt
l:tree
InT y (rbal' l x (Rd rl rx rr)) <-> X.eq y x \/ InT y l \/ InT y (Bk rl rx rr)
x:X.t
y:elt
l:tree

InT y (Rd l x Leaf) <-> X.eq y x \/ InT y l \/ InT y Leaf
x:X.t
y:elt
l:tree

InT y l \/ X.eq y x \/ False <-> X.eq y x \/ InT y l \/ False
intuition_in.
x:X.t
rl:tree
rx:X.t
rr:tree
y:elt
l:tree

InT y match rl with | Bk a y0 b => Rd (Bk l x a) y0 (rbal' b rx (makeRed rr)) | _ => Rd l x (Rd rl rx rr) end <-> X.eq y x \/ InT y l \/ InT y (Rd rl rx rr)
destree rl; autorew; intuition_in.
x:X.t
rl:tree
rx:X.t
rr:tree
y:elt
l:tree

InT y (rbal' l x (Rd rl rx rr)) <-> X.eq y x \/ InT y l \/ InT y (Bk rl rx rr)
x:X.t
rl:tree
rx:X.t
rr:tree
y:elt
l:tree

X.eq y x \/ InT y l \/ InT y rl \/ X.eq y rx \/ InT y rr <-> X.eq y x \/ InT y l \/ InT y rl \/ X.eq y rx \/ InT y rr
intuition_in. Qed.
l:tree
x:X.t
r:tree

Ok l -> Ok r -> lt_tree x l -> gt_tree x r -> Ok (lbalS l x r)
l:tree
x:X.t
r:tree

Ok l -> Ok r -> lt_tree x l -> gt_tree x r -> Ok (lbalS l x r)
l:tree
x:X.t
r, a:tree
x0:X.t
b:tree
H:Ok (Rd a x0 b)
H0:Ok r
lt_tree0:lt_tree x (Rd a x0 b)
gt_tree0:gt_tree x r

Ok (Rd (Bk a x0 b) x r)
l:tree
x:X.t
r, t0:tree
H:notred t0
H0:Ok t0
H1:Ok r
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x r
Ok match r with | Rd (Bk a0 y b) z c => Rd (Bk t0 x a0) y (rbal' b z (makeRed c)) | Bk a z c => rbal' t0 x (Rd a z c) | _ => Rd t0 x r end
l:tree
x:X.t
r, a:tree
x0:X.t
b:tree
H:Ok (Rd a x0 b)
H0:Ok r
lt_tree0:lt_tree x (Rd a x0 b)
gt_tree0:gt_tree x r

Ok (Rd (Bk a x0 b) x r)
ok.
l:tree
x:X.t
r, t0:tree
H:notred t0
H0:Ok t0
H1:Ok r
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x r

Ok match r with | Rd (Bk a0 y b) z c => Rd (Bk t0 x a0) y (rbal' b z (makeRed c)) | Bk a z c => rbal' t0 x (Rd a z c) | _ => Rd t0 x r end
l:tree
x:X.t
t0:tree
H:notred t0
H0:Ok t0
H1:Ok Leaf
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x Leaf

Ok (Rd t0 x Leaf)
l:tree
x:X.t
rl:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
H1:Ok (Rd rl rx rr)
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Rd rl rx rr)
Ok match rl with | Bk a y b => Rd (Bk t0 x a) y (rbal' b rx (makeRed rr)) | _ => Rd t0 x (Rd rl rx rr) end
l:tree
x:X.t
rl:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
H1:Ok (Bk rl rx rr)
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Bk rl rx rr)
Ok (rbal' t0 x (Rd rl rx rr))
l:tree
x:X.t
t0:tree
H:notred t0
H0:Ok t0
H1:Ok Leaf
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x Leaf

Ok (Rd t0 x Leaf)
ok.
l:tree
x:X.t
rl:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
H1:Ok (Rd rl rx rr)
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Rd rl rx rr)

Ok match rl with | Bk a y b => Rd (Bk t0 x a) y (rbal' b rx (makeRed rr)) | _ => Rd t0 x (Rd rl rx rr) end
l:tree
x:X.t
rll:tree
rlx:X.t
rlr:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)
H7:Ok rr
H8:lt_tree rx (Bk rll rlx rlr)
H9:gt_tree rx rr
H5:Ok rll
H10:Ok rlr
H11:lt_tree rlx rll
H12:gt_tree rlx rlr

Ok (rbal' rlr rx (makeRed rr))
l:tree
x:X.t
rll:tree
rlx:X.t
rlr:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)
H7:Ok rr
H8:lt_tree rx (Bk rll rlx rlr)
H9:gt_tree rx rr
H5:Ok rll
H10:Ok rlr
H11:lt_tree rlx rll
H12:gt_tree rlx rlr
gt_tree rlx (rbal' rlr rx (makeRed rr))
l:tree
x:X.t
rll:tree
rlx:X.t
rlr:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)
H7:Ok rr
H8:lt_tree rx (Bk rll rlx rlr)
H9:gt_tree rx rr
H5:Ok rll
H10:Ok rlr
H11:lt_tree rlx rll
H12:gt_tree rlx rlr

Ok (rbal' rlr rx (makeRed rr))
l:tree
x:X.t
rll:tree
rlx:X.t
rlr:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)
H7:Ok rr
H8:lt_tree rx (Bk rll rlx rlr)
H9:gt_tree rx rr
H5:Ok rll
H10:Ok rlr
H11:lt_tree rlx rll
H12:gt_tree rlx rlr

gt_tree rx (makeRed rr)
intros w; autorew; auto.
l:tree
x:X.t
rll:tree
rlx:X.t
rlr:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)
H7:Ok rr
H8:lt_tree rx (Bk rll rlx rlr)
H9:gt_tree rx rr
H5:Ok rll
H10:Ok rlr
H11:lt_tree rlx rll
H12:gt_tree rlx rlr

gt_tree rlx (rbal' rlr rx (makeRed rr))
l:tree
x:X.t
rll:tree
rlx:X.t
rlr:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)
H7:Ok rr
H8:lt_tree rx (Bk rll rlx rlr)
H9:gt_tree rx rr
H5:Ok rll
H10:Ok rlr
H11:lt_tree rlx rll
H12:gt_tree rlx rlr
w:elt

X.eq w rx \/ InT w rlr \/ InT w rr -> X.lt rlx w
destruct 1 as [Hw|[Hw|Hw]]; try rewrite Hw; eauto.
l:tree
x:X.t
rl:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
H1:Ok (Bk rl rx rr)
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Bk rl rx rr)

Ok (rbal' t0 x (Rd rl rx rr))
l:tree
x:X.t
rl:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Bk rl rx rr)
H6:Ok rl
H7:Ok rr
H8:lt_tree rx rl
H9:gt_tree rx rr

Ok (rbal' t0 x (Rd rl rx rr))
l:tree
x:X.t
rl:tree
rx:X.t
rr, t0:tree
H:notred t0
H0:Ok t0
lt_tree0:lt_tree x t0
gt_tree0:gt_tree x (Bk rl rx rr)
H6:Ok rl
H7:Ok rr
H8:lt_tree rx rl
H9:gt_tree rx rr

Ok (rbal' t0 x (Rd rl rx rr))
apply rbal'_ok; ok. Qed.
l:tree
x:X.t
r:tree
y:elt

InT y (rbalS l x r) <-> X.eq y x \/ InT y l \/ InT y r
l:tree
x:X.t
r:tree
y:elt

InT y (rbalS l x r) <-> X.eq y x \/ InT y l \/ InT y r
l:tree
x:X.t
r:tree
y:elt

forall (a : tree) (x0 : X.t) (b : tree), InT y (Rd l x (Bk a x0 b)) <-> X.eq y x \/ InT y l \/ InT y (Rd a x0 b)
l:tree
x:X.t
r:tree
y:elt
forall t0 : tree, notred t0 -> InT y match l with | Leaf => Rd l x t0 | Rd a y0 Leaf | Rd a y0 (Rd _ _ _) => Rd l x t0 | Rd a y0 (Bk b0 z c) => Rd (lbal (makeRed a) y0 b0) z (Bk c x t0) | Bk a y0 b => lbal (Rd a y0 b) x t0 end <-> X.eq y x \/ InT y l \/ InT y t0
l:tree
x:X.t
r:tree
y:elt

forall (a : tree) (x0 : X.t) (b : tree), InT y (Rd l x (Bk a x0 b)) <-> X.eq y x \/ InT y l \/ InT y (Rd a x0 b)
intros; autorew; intuition_in.
l:tree
x:X.t
r:tree
y:elt

forall t0 : tree, notred t0 -> InT y match l with | Leaf => Rd l x t0 | Rd a y0 Leaf | Rd a y0 (Rd _ _ _) => Rd l x t0 | Rd a y0 (Bk b0 z c) => Rd (lbal (makeRed a) y0 b0) z (Bk c x t0) | Bk a y0 b => lbal (Rd a y0 b) x t0 end <-> X.eq y x \/ InT y l \/ InT y t0
l:tree
x:X.t
r:tree
y:elt
t:tree

InT y match l with | Leaf => Rd l x t | Rd a y0 Leaf | Rd a y0 (Rd _ _ _) => Rd l x t | Rd a y0 (Bk b0 z c) => Rd (lbal (makeRed a) y0 b0) z (Bk c x t) | Bk a y0 b => lbal (Rd a y0 b) x t end <-> X.eq y x \/ InT y l \/ InT y t
x:X.t
r:tree
y:elt
t:tree

InT y (Rd Leaf x t) <-> X.eq y x \/ InT y Leaf \/ InT y t
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
y:elt
t:tree
InT y match lr with | Bk b z c => Rd (lbal (makeRed ll) lx b) z (Bk c x t) | _ => Rd (Rd ll lx lr) x t end <-> X.eq y x \/ InT y (Rd ll lx lr) \/ InT y t
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
y:elt
t:tree
InT y (lbal (Rd ll lx lr) x t) <-> X.eq y x \/ InT y (Bk ll lx lr) \/ InT y t
x:X.t
r:tree
y:elt
t:tree

InT y (Rd Leaf x t) <-> X.eq y x \/ InT y Leaf \/ InT y t
x:X.t
r:tree
y:elt
t:tree

False \/ X.eq y x \/ InT y t <-> X.eq y x \/ False \/ InT y t
intuition_in.
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
y:elt
t:tree

InT y match lr with | Bk b z c => Rd (lbal (makeRed ll) lx b) z (Bk c x t) | _ => Rd (Rd ll lx lr) x t end <-> X.eq y x \/ InT y (Rd ll lx lr) \/ InT y t
destruct lr as [|[|] lrl lrx lrr]; autorew; intuition_in.
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
y:elt
t:tree

InT y (lbal (Rd ll lx lr) x t) <-> X.eq y x \/ InT y (Bk ll lx lr) \/ InT y t
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
y:elt
t:tree

X.eq y x \/ (InT y ll \/ X.eq y lx \/ InT y lr) \/ InT y t <-> X.eq y x \/ (InT y ll \/ X.eq y lx \/ InT y lr) \/ InT y t
intuition_in. Qed.
l:tree
x:X.t
r:tree

Ok l -> Ok r -> lt_tree x l -> gt_tree x r -> Ok (rbalS l x r)
l:tree
x:X.t
r:tree

Ok l -> Ok r -> lt_tree x l -> gt_tree x r -> Ok (rbalS l x r)
l:tree
x:X.t
r, a:tree
x0:X.t
b:tree
H:Ok l
H0:Ok (Rd a x0 b)
lt_tree0:lt_tree x l
gt_tree0:gt_tree x (Rd a x0 b)

Ok (Rd l x (Bk a x0 b))
l:tree
x:X.t
r, t0:tree
H:notred t0
H0:Ok l
H1:Ok t0
lt_tree0:lt_tree x l
gt_tree0:gt_tree x t0
Ok match l with | Leaf => Rd l x t0 | Rd a y Leaf | Rd a y (Rd _ _ _) => Rd l x t0 | Rd a y (Bk b0 z c) => Rd (lbal (makeRed a) y b0) z (Bk c x t0) | Bk a y b => lbal (Rd a y b) x t0 end
l:tree
x:X.t
r, a:tree
x0:X.t
b:tree
H:Ok l
H0:Ok (Rd a x0 b)
lt_tree0:lt_tree x l
gt_tree0:gt_tree x (Rd a x0 b)

Ok (Rd l x (Bk a x0 b))
ok.
l:tree
x:X.t
r, t0:tree
H:notred t0
H0:Ok l
H1:Ok t0
lt_tree0:lt_tree x l
gt_tree0:gt_tree x t0

Ok match l with | Leaf => Rd l x t0 | Rd a y Leaf | Rd a y (Rd _ _ _) => Rd l x t0 | Rd a y (Bk b0 z c) => Rd (lbal (makeRed a) y b0) z (Bk c x t0) | Bk a y b => lbal (Rd a y b) x t0 end
x:X.t
r, t0:tree
H:notred t0
H0:Ok Leaf
H1:Ok t0
lt_tree0:lt_tree x Leaf
gt_tree0:gt_tree x t0

Ok (Rd Leaf x t0)
ll:tree
lx:X.t
lr:tree
x:X.t
r, t0:tree
H:notred t0
H0:Ok (Rd ll lx lr)
H1:Ok t0
lt_tree0:lt_tree x (Rd ll lx lr)
gt_tree0:gt_tree x t0
Ok match lr with | Bk b z c => Rd (lbal (makeRed ll) lx b) z (Bk c x t0) | _ => Rd (Rd ll lx lr) x t0 end
ll:tree
lx:X.t
lr:tree
x:X.t
r, t0:tree
H:notred t0
H0:Ok (Bk ll lx lr)
H1:Ok t0
lt_tree0:lt_tree x (Bk ll lx lr)
gt_tree0:gt_tree x t0
Ok (lbal (Rd ll lx lr) x t0)
x:X.t
r, t0:tree
H:notred t0
H0:Ok Leaf
H1:Ok t0
lt_tree0:lt_tree x Leaf
gt_tree0:gt_tree x t0

Ok (Rd Leaf x t0)
ok.
ll:tree
lx:X.t
lr:tree
x:X.t
r, t0:tree
H:notred t0
H0:Ok (Rd ll lx lr)
H1:Ok t0
lt_tree0:lt_tree x (Rd ll lx lr)
gt_tree0:gt_tree x t0

Ok match lr with | Bk b z c => Rd (lbal (makeRed ll) lx b) z (Bk c x t0) | _ => Rd (Rd ll lx lr) x t0 end
ll:tree
lx:X.t
lrl:tree
lrx:X.t
lrr:tree
x:X.t
r, t0:tree
H:notred t0
H1:Ok t0
lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))
gt_tree0:gt_tree x t0
H6:Ok ll
H8:lt_tree lx ll
H9:gt_tree lx (Bk lrl lrx lrr)
H5:Ok lrl
H10:Ok lrr
H11:lt_tree lrx lrl
H12:gt_tree lrx lrr

Ok (lbal (makeRed ll) lx lrl)
ll:tree
lx:X.t
lrl:tree
lrx:X.t
lrr:tree
x:X.t
r, t0:tree
H:notred t0
H1:Ok t0
lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))
gt_tree0:gt_tree x t0
H6:Ok ll
H8:lt_tree lx ll
H9:gt_tree lx (Bk lrl lrx lrr)
H5:Ok lrl
H10:Ok lrr
H11:lt_tree lrx lrl
H12:gt_tree lrx lrr
lt_tree lrx (lbal (makeRed ll) lx lrl)
ll:tree
lx:X.t
lrl:tree
lrx:X.t
lrr:tree
x:X.t
r, t0:tree
H:notred t0
H1:Ok t0
lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))
gt_tree0:gt_tree x t0
H6:Ok ll
H8:lt_tree lx ll
H9:gt_tree lx (Bk lrl lrx lrr)
H5:Ok lrl
H10:Ok lrr
H11:lt_tree lrx lrl
H12:gt_tree lrx lrr

Ok (lbal (makeRed ll) lx lrl)
ll:tree
lx:X.t
lrl:tree
lrx:X.t
lrr:tree
x:X.t
r, t0:tree
H:notred t0
H1:Ok t0
lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))
gt_tree0:gt_tree x t0
H6:Ok ll
H8:lt_tree lx ll
H9:gt_tree lx (Bk lrl lrx lrr)
H5:Ok lrl
H10:Ok lrr
H11:lt_tree lrx lrl
H12:gt_tree lrx lrr

lt_tree lx (makeRed ll)
intros w; autorew; auto.
ll:tree
lx:X.t
lrl:tree
lrx:X.t
lrr:tree
x:X.t
r, t0:tree
H:notred t0
H1:Ok t0
lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))
gt_tree0:gt_tree x t0
H6:Ok ll
H8:lt_tree lx ll
H9:gt_tree lx (Bk lrl lrx lrr)
H5:Ok lrl
H10:Ok lrr
H11:lt_tree lrx lrl
H12:gt_tree lrx lrr

lt_tree lrx (lbal (makeRed ll) lx lrl)
ll:tree
lx:X.t
lrl:tree
lrx:X.t
lrr:tree
x:X.t
r, t0:tree
H:notred t0
H1:Ok t0
lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))
gt_tree0:gt_tree x t0
H6:Ok ll
H8:lt_tree lx ll
H9:gt_tree lx (Bk lrl lrx lrr)
H5:Ok lrl
H10:Ok lrr
H11:lt_tree lrx lrl
H12:gt_tree lrx lrr
w:elt

X.eq w lx \/ InT w ll \/ InT w lrl -> X.lt w lrx
destruct 1 as [Hw|[Hw|Hw]]; try rewrite Hw; eauto.
ll:tree
lx:X.t
lr:tree
x:X.t
r, t0:tree
H:notred t0
H0:Ok (Bk ll lx lr)
H1:Ok t0
lt_tree0:lt_tree x (Bk ll lx lr)
gt_tree0:gt_tree x t0

Ok (lbal (Rd ll lx lr) x t0)
ll:tree
lx:X.t
lr:tree
x:X.t
r, t0:tree
H:notred t0
H1:Ok t0
lt_tree0:lt_tree x (Bk ll lx lr)
gt_tree0:gt_tree x t0
H6:Ok ll
H7:Ok lr
H8:lt_tree lx ll
H9:gt_tree lx lr

Ok (lbal (Rd ll lx lr) x t0)
apply lbal_ok; ok. Qed. Hint Rewrite lbalS_spec rbalS_spec : rb.

Append for deletion

Ltac append_tac l r :=
 induction l as [| lc ll _ lx lr IHlr];
 [intro r; simpl
 |induction r as [| rc rl IHrl rx rr _];
   [simpl
   |destruct lc, rc;
     [specialize (IHlr rl); clear IHrl
     |simpl;
      assert (Hr:notred (Bk rl rx rr)) by (simpl; trivial);
      set (r:=Bk rl rx rr) in *; clearbody r; clear IHrl rl rx rr;
      specialize (IHlr r)
     |change (append _ _) with (Rd (append (Bk ll lx lr) rl) rx rr);
      assert (Hl:notred (Bk ll lx lr)) by (simpl; trivial);
      set (l:=Bk ll lx lr) in *; clearbody l; clear IHlr ll lx lr
     |specialize (IHlr rl); clear IHrl]]].

ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree

rspec (fun (a : tree) (x : X.t) (b : tree) => Rd (Rd ll lx a) x (Rd b rx rr)) (fun t0 : tree => Rd ll lx (Rd t0 rx rr)) (append lr rl) (append (Rd ll lx lr) (Rd rl rx rr))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree

rspec (fun (a : tree) (x : X.t) (b : tree) => Rd (Rd ll lx a) x (Rd b rx rr)) (fun t0 : tree => Rd ll lx (Rd t0 rx rr)) (append lr rl) (append (Rd ll lx lr) (Rd rl rx rr))
exact (rmatch _ _ _). Qed.
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree

rspec (fun (a : tree) (x : X.t) (b : tree) => Rd (Bk ll lx a) x (Bk b rx rr)) (fun t0 : tree => lbalS ll lx (Bk t0 rx rr)) (append lr rl) (append (Bk ll lx lr) (Bk rl rx rr))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree

rspec (fun (a : tree) (x : X.t) (b : tree) => Rd (Bk ll lx a) x (Bk b rx rr)) (fun t0 : tree => lbalS ll lx (Bk t0 rx rr)) (append lr rl) (append (Bk ll lx lr) (Bk rl rx rr))
exact (rmatch _ _ _). Qed.
l, r:tree
x:elt

InT x (append l r) <-> InT x l \/ InT x r
l, r:tree
x:elt

InT x (append l r) <-> InT x l \/ InT x r
l:tree
x:elt

forall r : tree, InT x (append l r) <-> InT x l \/ InT x r
ll:tree
lx:X.t
lr:tree
x:elt
rl:tree
IHlr:InT x (append lr rl) <-> InT x lr \/ InT x rl
rx:X.t
rr:tree

InT x (append (Rd ll lx lr) (Rd rl rx rr)) <-> (InT x ll \/ X.eq x lx \/ InT x lr) \/ InT x rl \/ X.eq x rx \/ InT x rr
ll:tree
lx:X.t
lr:tree
x:elt
rl:tree
IHlr:InT x (append lr rl) <-> InT x lr \/ InT x rl
rx:X.t
rr:tree
InT x (append (Bk ll lx lr) (Bk rl rx rr)) <-> (InT x ll \/ X.eq x lx \/ InT x lr) \/ InT x rl \/ X.eq x rx \/ InT x rr
ll:tree
lx:X.t
lr:tree
x:elt
rl:tree
IHlr:InT x (append lr rl) <-> InT x lr \/ InT x rl
rx:X.t
rr:tree

InT x (append (Rd ll lx lr) (Rd rl rx rr)) <-> (InT x ll \/ X.eq x lx \/ InT x lr) \/ InT x rl \/ X.eq x rx \/ InT x rr
revert IHlr; case append_rr_match; [intros a y b | intros t Ht]; autorew; tauto.
ll:tree
lx:X.t
lr:tree
x:elt
rl:tree
IHlr:InT x (append lr rl) <-> InT x lr \/ InT x rl
rx:X.t
rr:tree

InT x (append (Bk ll lx lr) (Bk rl rx rr)) <-> (InT x ll \/ X.eq x lx \/ InT x lr) \/ InT x rl \/ X.eq x rx \/ InT x rr
revert IHlr; case append_bb_match; [intros a y b | intros t Ht]; autorew; tauto. Qed. Hint Rewrite append_spec : rb.

forall (x : X.t) (l r : tree), Ok l -> Ok r -> lt_tree x l -> gt_tree x r -> Ok (append l r)

forall (x : X.t) (l r : tree), Ok l -> Ok r -> lt_tree x l -> gt_tree x r -> Ok (append l r)
x:X.t
r:tree

Ok Leaf -> Ok r -> lt_tree x Leaf -> gt_tree x r -> Ok r
x:X.t
lc:Color.t
ll:tree
lx:X.t
lr:tree
IHlr:forall r : tree, Ok lr -> Ok r -> lt_tree x lr -> gt_tree x r -> Ok (append lr r)
Ok (Node lc ll lx lr) -> Ok Leaf -> lt_tree x (Node lc ll lx lr) -> gt_tree x Leaf -> Ok (Node lc ll lx lr)
x:X.t
ll:tree
lx:X.t
lr, rl:tree
IHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)
rx:X.t
rr:tree
Ok (Rd ll lx lr) -> Ok (Rd rl rx rr) -> lt_tree x (Rd ll lx lr) -> gt_tree x (Rd rl rx rr) -> Ok (append (Rd ll lx lr) (Rd rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, r:tree
IHlr:Ok lr -> Ok r -> lt_tree x lr -> gt_tree x r -> Ok (append lr r)
Hr:notred r
Ok (Rd ll lx lr) -> Ok r -> lt_tree x (Rd ll lx lr) -> gt_tree x r -> Ok (Rd ll lx (append lr r))
x:X.t
rl:tree
rx:X.t
rr, l:tree
IHrl:Ok l -> Ok rl -> lt_tree x l -> gt_tree x rl -> Ok (append l rl)
Hl:notred l
Ok l -> Ok (Rd rl rx rr) -> lt_tree x l -> gt_tree x (Rd rl rx rr) -> Ok (Rd (append l rl) rx rr)
x:X.t
ll:tree
lx:X.t
lr, rl:tree
IHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)
rx:X.t
rr:tree
Ok (Bk ll lx lr) -> Ok (Bk rl rx rr) -> lt_tree x (Bk ll lx lr) -> gt_tree x (Bk rl rx rr) -> Ok (append (Bk ll lx lr) (Bk rl rx rr))
x:X.t
r:tree

Ok Leaf -> Ok r -> lt_tree x Leaf -> gt_tree x r -> Ok r
trivial.
x:X.t
lc:Color.t
ll:tree
lx:X.t
lr:tree
IHlr:forall r : tree, Ok lr -> Ok r -> lt_tree x lr -> gt_tree x r -> Ok (append lr r)

Ok (Node lc ll lx lr) -> Ok Leaf -> lt_tree x (Node lc ll lx lr) -> gt_tree x Leaf -> Ok (Node lc ll lx lr)
trivial.
x:X.t
ll:tree
lx:X.t
lr, rl:tree
IHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)
rx:X.t
rr:tree

Ok (Rd ll lx lr) -> Ok (Rd rl rx rr) -> lt_tree x (Rd ll lx lr) -> gt_tree x (Rd rl rx rr) -> Ok (append (Rd ll lx lr) (Rd rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
IHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr

Ok (append (Rd ll lx lr) (Rd rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
IHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)

Ok (append (Rd ll lx lr) (Rd rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)

Ok (append (Rd ll lx lr) (Rd rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx

Ok (append (Rd ll lx lr) (Rd rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx

gt_tree lx (append lr rl)
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)
Ok (append (Rd ll lx lr) (Rd rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx

gt_tree lx (append lr rl)
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
w:elt

InT w (append lr rl) -> X.lt lx w
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
w:elt

InT w lr \/ InT w rl -> X.lt lx w
destruct 1; [|transitivity x]; eauto.
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)

Ok (append (Rd ll lx lr) (Rd rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)

lt_tree rx (append lr rl)
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)
L:lt_tree rx (append lr rl)
Ok (append (Rd ll lx lr) (Rd rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)

lt_tree rx (append lr rl)
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)
w:elt

InT w (append lr rl) -> X.lt w rx
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)
w:elt

InT w lr \/ InT w rl -> X.lt w rx
destruct 1; [transitivity x|]; eauto.
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)
L:lt_tree rx (append lr rl)

Ok (append (Rd ll lx lr) (Rd rl rx rr))
revert IH G L; case append_rr_match; intros; ok.
x:X.t
ll:tree
lx:X.t
lr, r:tree
IHlr:Ok lr -> Ok r -> lt_tree x lr -> gt_tree x r -> Ok (append lr r)
Hr:notred r

Ok (Rd ll lx lr) -> Ok r -> lt_tree x (Rd ll lx lr) -> gt_tree x r -> Ok (Rd ll lx (append lr r))
x:X.t
ll:tree
lx:X.t
lr, r:tree
IHlr:Ok lr -> Ok r -> lt_tree x lr -> gt_tree x r -> Ok (append lr r)
Hr:notred r
H0:Ok r
H1:lt_tree x (Rd ll lx lr)
H2:gt_tree x r
H7:Ok ll
H8:Ok lr
H9:lt_tree lx ll
H10:gt_tree lx lr

gt_tree lx (append lr r)
intros w; autorew; destruct 1; eauto.
x:X.t
rl:tree
rx:X.t
rr, l:tree
IHrl:Ok l -> Ok rl -> lt_tree x l -> gt_tree x rl -> Ok (append l rl)
Hl:notred l

Ok l -> Ok (Rd rl rx rr) -> lt_tree x l -> gt_tree x (Rd rl rx rr) -> Ok (Rd (append l rl) rx rr)
x:X.t
rl:tree
rx:X.t
rr, l:tree
IHrl:Ok l -> Ok rl -> lt_tree x l -> gt_tree x rl -> Ok (append l rl)
Hl:notred l
H:Ok l
H1:lt_tree x l
H2:gt_tree x (Rd rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr

lt_tree rx (append l rl)
intros w; autorew; destruct 1; eauto.
x:X.t
ll:tree
lx:X.t
lr, rl:tree
IHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)
rx:X.t
rr:tree

Ok (Bk ll lx lr) -> Ok (Bk rl rx rr) -> lt_tree x (Bk ll lx lr) -> gt_tree x (Bk rl rx rr) -> Ok (append (Bk ll lx lr) (Bk rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
IHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr

Ok (append (Bk ll lx lr) (Bk rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
IHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)

Ok (append (Bk ll lx lr) (Bk rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)

Ok (append (Bk ll lx lr) (Bk rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx

Ok (append (Bk ll lx lr) (Bk rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx

gt_tree lx (append lr rl)
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)
Ok (append (Bk ll lx lr) (Bk rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx

gt_tree lx (append lr rl)
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
w:elt

InT w (append lr rl) -> X.lt lx w
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
w:elt

InT w lr \/ InT w rl -> X.lt lx w
destruct 1; [|transitivity x]; eauto.
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)

Ok (append (Bk ll lx lr) (Bk rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)

lt_tree rx (append lr rl)
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)
L:lt_tree rx (append lr rl)
Ok (append (Bk ll lx lr) (Bk rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)

lt_tree rx (append lr rl)
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)
w:elt

InT w (append lr rl) -> X.lt w rx
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)
w:elt

InT w lr \/ InT w rl -> X.lt w rx
destruct 1; [transitivity x|]; eauto.
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
IH:Ok (append lr rl)
H:X.lt lx rx
G:gt_tree lx (append lr rl)
L:lt_tree rx (append lr rl)

Ok (append (Bk ll lx lr) (Bk rl rx rr))
x:X.t
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
H1:lt_tree x (Bk ll lx lr)
H2:gt_tree x (Bk rl rx rr)
H7:Ok rl
H8:Ok rr
H9:lt_tree rx rl
H10:gt_tree rx rr
H6:Ok ll
H11:Ok lr
H12:lt_tree lx ll
H13:gt_tree lx lr
H:X.lt lx rx
t0:tree
H0:notred t0
IH:Ok t0
G:gt_tree lx t0
L:lt_tree rx t0

Ok (lbalS ll lx (Bk t0 rx rr))
apply lbalS_ok; ok. Qed.

Deletion


forall (s : tree) (x : X.t) (y : elt), Ok s -> InT y (del x s) <-> InT y s /\ ~ X.eq y x

forall (s : tree) (x : X.t) (y : elt), Ok s -> InT y (del x s) <-> InT y s /\ ~ X.eq y x
x:X.t
y:elt
H:Ok Leaf

InT y Leaf <-> InT y Leaf /\ ~ X.eq y x
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
InT y (append l r) <-> InT y (Node i l x' r) /\ ~ X.eq y x
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0
x:X.t
y:elt
H0:X.lt x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
InT y match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end <-> InT y (Node i l x' r) /\ ~ X.eq y x
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0
x:X.t
y:elt
H0:X.lt x' x
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
InT y match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end <-> InT y (Node i l x' r) /\ ~ X.eq y x
x:X.t
y:elt
H:Ok Leaf

InT y Leaf <-> InT y Leaf /\ ~ X.eq y x
intuition_in.
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r

InT y (append l r) <-> InT y (Node i l x' r) /\ ~ X.eq y x
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ (X.eq y0 x0 -> False)
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ (X.eq y0 x0 -> False)
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H1:InT y l
H:X.eq y x

False
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ (X.eq y0 x0 -> False)
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ (X.eq y0 x0 -> False)
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H1:InT y r
H:X.eq y x
False
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ (X.eq y0 x0 -> False)
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ (X.eq y0 x0 -> False)
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H2:X.eq y x -> False
H1:X.eq y x'
InT y l \/ InT y r
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ (X.eq y0 x0 -> False)
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ (X.eq y0 x0 -> False)
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H1:InT y l
H:X.eq y x
H2:X.lt y x'

False
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ (X.eq y0 x0 -> False)
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ (X.eq y0 x0 -> False)
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H1:InT y r
H:X.eq y x
False
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ (X.eq y0 x0 -> False)
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ (X.eq y0 x0 -> False)
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H2:X.eq y x -> False
H1:X.eq y x'
InT y l \/ InT y r
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ (X.eq y0 x0 -> False)
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ (X.eq y0 x0 -> False)
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H1:InT y r
H:X.eq y x

False
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ (X.eq y0 x0 -> False)
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ (X.eq y0 x0 -> False)
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H2:X.eq y x -> False
H1:X.eq y x'
InT y l \/ InT y r
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ (X.eq y0 x0 -> False)
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ (X.eq y0 x0 -> False)
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H1:InT y r
H:X.eq y x
H2:X.lt x' y

False
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ (X.eq y0 x0 -> False)
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ (X.eq y0 x0 -> False)
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H2:X.eq y x -> False
H1:X.eq y x'
InT y l \/ InT y r
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ (X.eq y0 x0 -> False)
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ (X.eq y0 x0 -> False)
x:X.t
y:elt
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H2:X.eq y x -> False
H1:X.eq y x'

InT y l \/ InT y r
order.
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0
x:X.t
y:elt
H0:X.lt x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r

InT y match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end <-> InT y (Node i l x' r) /\ ~ X.eq y x
destruct l as [|[|] ll lx lr]; autorew; rewrite ?IHl by trivial; intuition_in; order.
i:Color.t
l:tree
x':X.t
r:tree
IHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0
IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0
x:X.t
y:elt
H0:X.lt x' x
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r

InT y match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end <-> InT y (Node i l x' r) /\ ~ X.eq y x
destruct r as [|[|] rl rx rr]; autorew; rewrite ?IHr by trivial; intuition_in; order. Qed. Hint Rewrite del_spec : rb.
s:tree
x:X.t
H:Ok s

Ok (del x s)
s:tree
x:X.t
H:Ok s

Ok (del x s)
x:X.t
H:Ok Leaf

Ok Leaf
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
Ok (append l r)
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
Ok match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x' x
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
Ok match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
x:X.t
H:Ok Leaf

Ok Leaf
trivial.
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.eq x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r

Ok (append l r)
eapply append_ok; eauto.
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r

Ok match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r

lt_tree x' (del x l)
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H:lt_tree x' (del x l)
Ok match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r

lt_tree x' (del x l)
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
w:elt

InT w (del x l) -> X.lt w x'
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
w:elt

InT w l /\ ~ X.eq w x -> X.lt w x'
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
w:elt
H:InT w l
H1:~ X.eq w x

X.lt w x'
eauto.
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x x'
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H:lt_tree x' (del x l)

Ok match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
destruct l as [|[|] ll lx lr]; auto_tc.
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x' x
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r

Ok match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x' x
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r

gt_tree x' (del x r)
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x' x
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H:gt_tree x' (del x r)
Ok match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x' x
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r

gt_tree x' (del x r)
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x' x
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
w:elt

InT w (del x r) -> X.lt x' w
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x' x
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
w:elt

InT w r /\ ~ X.eq w x -> X.lt x' w
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x' x
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
w:elt
H:InT w r
H1:~ X.eq w x

X.lt x' w
eauto.
i:Color.t
l:tree
x':X.t
r:tree
x:X.t
IHl:Ok l -> Ok (del x l)
IHr:Ok r -> Ok (del x r)
H0:X.lt x' x
H5:Ok l
H6:Ok r
H7:lt_tree x' l
H8:gt_tree x' r
H:gt_tree x' (del x r)

Ok match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
destruct r as [|[|] rl rx rr]; auto_tc. Qed.
s:tree
x:X.t
y:elt
H:Ok s

InT y (remove x s) <-> InT y s /\ ~ X.eq y x
s:tree
x:X.t
y:elt
H:Ok s

InT y (remove x s) <-> InT y s /\ ~ X.eq y x
s:tree
x:X.t
y:elt
H:Ok s

InT y (makeBlack (del x s)) <-> InT y s /\ ~ X.eq y x
now autorew. Qed. Hint Rewrite remove_spec : rb.
s:tree
x:X.t
H:Ok s

Ok (remove x s)
s:tree
x:X.t
H:Ok s

Ok (remove x s)
unfold remove; auto_tc. Qed.

Removing the minimal element

l:tree
y:X.t
r:tree
c:Color.t
x:elt
s':tree
O:Ok (Node c l y r)

delmin l y r = (x, s') -> min_elt (Node c l y r) = Some x /\ del x (Node c l y r) = s'
l:tree
y:X.t
r:tree
c:Color.t
x:elt
s':tree
O:Ok (Node c l y r)

delmin l y r = (x, s') -> min_elt (Node c l y r) = Some x /\ del x (Node c l y r) = s'
l:tree

forall (y : X.t) (r : tree) (c : Color.t) (x : elt) (s' : tree), Ok (Node c l y r) -> delmin l y r = (x, s') -> min_elt (Node c l y r) = Some x /\ del x (Node c l y r) = s'

forall (y : X.t) (r : tree) (c : Color.t) (x : elt) (s' : tree), Ok (Node c Leaf y r) -> delmin Leaf y r = (x, s') -> min_elt (Node c Leaf y r) = Some x /\ del x (Node c Leaf y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
IH:forall (y : X.t) (r : tree) (c : Color.t) (x : elt) (s' : tree), Ok (Node c ll y r) -> delmin ll y r = (x, s') -> min_elt (Node c ll y r) = Some x /\ del x (Node c ll y r) = s'
forall (y : X.t) (r : tree) (c : Color.t) (x : elt) (s' : tree), Ok (Node c (Node lc ll ly lr) y r) -> delmin (Node lc ll ly lr) y r = (x, s') -> min_elt (Node c (Node lc ll ly lr) y r) = Some x /\ del x (Node c (Node lc ll ly lr) y r) = s'

forall (y : X.t) (r : tree) (c : Color.t) (x : elt) (s' : tree), Ok (Node c Leaf y r) -> delmin Leaf y r = (x, s') -> min_elt (Node c Leaf y r) = Some x /\ del x (Node c Leaf y r) = s'

forall (y : X.t) (r : tree) (c : Color.t) (x : elt) (s' : tree), Ok (Node c Leaf y r) -> (y, r) = (x, s') -> Some y = Some x /\ match X.compare x y with | Eq => r | Lt => Rd Leaf y r | Gt => match r with | Bk _ _ _ => rbalS Leaf y (del x r) | _ => Rd Leaf y (del x r) end end = s'
y:X.t
r:tree
x:elt
s':tree

(y, r) = (x, s') -> Some y = Some x /\ match X.compare x y with | Eq => r | Lt => Rd Leaf y r | Gt => match r with | Bk _ _ _ => rbalS Leaf y (del x r) | _ => Rd Leaf y (del x r) end end = s'
x:elt
s':tree
H:(x, s') = (x, s')

Some x = Some x /\ match X.compare x x with | Eq => s' | Lt => Rd Leaf x s' | Gt => match s' with | Bk _ _ _ => rbalS Leaf x (del x s') | _ => Rd Leaf x (del x s') end end = s'
now rewrite MX.compare_refl.
lc:Color.t
ll:tree
ly:X.t
lr:tree
IH:forall (y : X.t) (r : tree) (c : Color.t) (x : elt) (s' : tree), Ok (Node c ll y r) -> delmin ll y r = (x, s') -> min_elt (Node c ll y r) = Some x /\ del x (Node c ll y r) = s'

forall (y : X.t) (r : tree) (c : Color.t) (x : elt) (s' : tree), Ok (Node c (Node lc ll ly lr) y r) -> delmin (Node lc ll ly lr) y r = (x, s') -> min_elt (Node c (Node lc ll ly lr) y r) = Some x /\ del x (Node c (Node lc ll ly lr) y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
IH:forall (y0 : X.t) (r0 : tree) (c0 : Color.t) (x0 : elt) (s'0 : tree), Ok (Node c0 ll y0 r0) -> delmin ll y0 r0 = (x0, s'0) -> min_elt (Node c0 ll y0 r0) = Some x0 /\ del x0 (Node c0 ll y0 r0) = s'0
y:X.t
r:tree
c:Color.t
x:elt
s':tree
O:Ok (Node c (Node lc ll ly lr) y r)

delmin (Node lc ll ly lr) y r = (x, s') -> min_elt (Node c (Node lc ll ly lr) y r) = Some x /\ del x (Node c (Node lc ll ly lr) y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
IH:forall (y0 : X.t) (r0 : tree) (c0 : Color.t) (x0 : elt) (s'0 : tree), Ok (Node c0 ll y0 r0) -> delmin ll y0 r0 = (x0, s'0) -> min_elt (Node c0 ll y0 r0) = Some x0 /\ del x0 (Node c0 ll y0 r0) = s'0
y:X.t
r:tree
c:Color.t
x:elt
s':tree
O:Ok (Node c (Node lc ll ly lr) y r)

(let (k, l') := delmin ll ly lr in match lc with | Red => (k, Rd l' y r) | Black => (k, lbalS l' y r) end) = (x, s') -> min_elt (Node c (Node lc ll ly lr) y r) = Some x /\ del x (Node c (Node lc ll ly lr) y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
IH:forall (c0 : Color.t) (x0 : elt) (s'0 : tree), Ok (Node c0 ll ly lr) -> delmin ll ly lr = (x0, s'0) -> min_elt (Node c0 ll ly lr) = Some x0 /\ del x0 (Node c0 ll ly lr) = s'0
y:X.t
r:tree
c:Color.t
x:elt
s':tree
O:Ok (Node c (Node lc ll ly lr) y r)

(let (k, l') := delmin ll ly lr in match lc with | Red => (k, Rd l' y r) | Black => (k, lbalS l' y r) end) = (x, s') -> min_elt (Node c (Node lc ll ly lr) y r) = Some x /\ del x (Node c (Node lc ll ly lr) y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
IH:forall (c0 : Color.t) (x1 : elt) (s'0 : tree), Ok (Node c0 ll ly lr) -> (x0, s0) = (x1, s'0) -> min_elt (Node c0 ll ly lr) = Some x1 /\ del x1 (Node c0 ll ly lr) = s'0
y:X.t
r:tree
c:Color.t
x:elt
s':tree
O:Ok (Node c (Node lc ll ly lr) y r)

match lc with | Red => (x0, Rd s0 y r) | Black => (x0, lbalS s0 y r) end = (x, s') -> min_elt (Node c (Node lc ll ly lr) y r) = Some x /\ del x (Node c (Node lc ll ly lr) y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s':tree
O:Ok (Node c (Node lc ll ly lr) y r)
H:min_elt (Node lc ll ly lr) = Some x0
H0:del x0 (Node lc ll ly lr) = s0

match lc with | Red => (x0, Rd s0 y r) | Black => (x0, lbalS s0 y r) end = (x, s') -> min_elt (Node c (Node lc ll ly lr) y r) = Some x /\ del x (Node c (Node lc ll ly lr) y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x0
H0:del x0 l = s0

match lc with | Red => (x0, Rd s0 y r) | Black => (x0, lbalS s0 y r) end = (x, s') -> min_elt (Node c l y r) = Some x /\ del x (Node c l y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x0
H0:del x0 l = s0

match lc with | Red => (x0, Rd s0 y r) | Black => (x0, lbalS s0 y r) end = (x, s') -> match l with | Leaf => Some y | Node _ _ _ _ => min_elt l end = Some x /\ del x (Node c l y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x0
H0:del x0 l = s0
E:match lc with | Red => (x0, Rd s0 y r) | Black => (x0, lbalS s0 y r) end = (x, s')

match l with | Leaf => Some y | Node _ _ _ _ => min_elt l end = Some x /\ del x (Node c l y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')

match l with | Leaf => Some y | Node _ _ _ _ => min_elt l end = Some x /\ del x (Node c l y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')

match l with | Leaf => Some y | Node _ _ _ _ => min_elt l end = Some x
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')
del x (Node c l y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')

match l with | Leaf => Some y | Node _ _ _ _ => min_elt l end = Some x
subst l; intuition.
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')

del x (Node c l y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')

X.lt x y
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')
H1:X.lt x y
del x (Node c l y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')

X.lt x y
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')
H1:bst l
H2:bst r
H3:lt_tree y l
H4:gt_tree y r

X.lt x y
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')
H1:bst l
H2:bst r
H3:lt_tree y l
H4:gt_tree y r
H5:InT x l

X.lt x y
auto.
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')
H1:X.lt x y

del x (Node c l y r) = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')
H1:X.lt x y

match X.compare x y with | Eq => append l r | Lt => match l with | Bk _ _ _ => lbalS (del x l) y r | _ => Rd (del x l) y r end | Gt => match r with | Bk _ _ _ => rbalS l y (del x r) | _ => Rd l y (del x r) end end = s'
lc:Color.t
ll:tree
ly:X.t
lr:tree
x0:elt
s0:tree
y:X.t
r:tree
c:Color.t
x:elt
s', l:tree
Heql:l = Node lc ll ly lr
O:Ok (Node c l y r)
H:min_elt l = Some x
H0:del x l = s0
E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')
H1:X.lt x y

X.lt x y -> match l with | Bk _ _ _ => lbalS (del x l) y r | _ => Rd (del x l) y r end = s'
destruct lc; injection E; subst l s0; auto. Qed.
s:tree
x:elt
s':tree
H:Ok s

remove_min s = Some (x, s') -> min_elt s = Some x /\ remove x s = s'
s:tree
x:elt
s':tree
H:Ok s

remove_min s = Some (x, s') -> min_elt s = Some x /\ remove x s = s'
s:tree
x:elt
s':tree
H:Ok s

match s with | Leaf => None | Node _ l x0 r => let (k, t) := delmin l x0 r in Some (k, makeBlack t) end = Some (x, s') -> min_elt s = Some x /\ remove x s = s'
c:Color.t
l:tree
y:X.t
r:tree
x:elt
s':tree
H:Ok (Node c l y r)

(let (k, t) := delmin l y r in Some (k, makeBlack t)) = Some (x, s') -> min_elt (Node c l y r) = Some x /\ remove x (Node c l y r) = s'
c:Color.t
l:tree
y:X.t
r:tree
x:elt
s':tree
H:Ok (Node c l y r)

(forall (x0 : elt) (s'0 : tree), Ok (Node c l y r) -> delmin l y r = (x0, s'0) -> min_elt (Node c l y r) = Some x0 /\ del x0 (Node c l y r) = s'0) -> (let (k, t) := delmin l y r in Some (k, makeBlack t)) = Some (x, s') -> min_elt (Node c l y r) = Some x /\ remove x (Node c l y r) = s'
c:Color.t
l:tree
y:X.t
r:tree
x:elt
s':tree
H:Ok (Node c l y r)
x0:elt
s0:tree

(forall (x1 : elt) (s'0 : tree), Ok (Node c l y r) -> (x0, s0) = (x1, s'0) -> min_elt (Node c l y r) = Some x1 /\ del x1 (Node c l y r) = s'0) -> Some (x0, makeBlack s0) = Some (x, s') -> min_elt (Node c l y r) = Some x /\ remove x (Node c l y r) = s'
c:Color.t
l:tree
y:X.t
r:tree
x:elt
s':tree
H:Ok (Node c l y r)
x0:elt
s0:tree
D:forall (x1 : elt) (s'0 : tree), Ok (Node c l y r) -> (x0, s0) = (x1, s'0) -> min_elt (Node c l y r) = Some x1 /\ del x1 (Node c l y r) = s'0

Some (x0, makeBlack s0) = Some (x, s') -> min_elt (Node c l y r) = Some x /\ remove x (Node c l y r) = s'
c:Color.t
l:tree
y:X.t
r:tree
x:elt
s':tree
H:Ok (Node c l y r)
x0:elt
D:forall (x1 : elt) (s'0 : tree), Ok (Node c l y r) -> (x0, del x0 (Node c l y r)) = (x1, s'0) -> min_elt (Node c l y r) = Some x1 /\ del x1 (Node c l y r) = s'0

Some (x0, makeBlack (del x0 (Node c l y r))) = Some (x, s') -> Some x0 = Some x /\ remove x (Node c l y r) = s'
c:Color.t
l:tree
y:X.t
r:tree
x:elt
s':tree
H:Ok (Node c l y r)
x0:elt
D:forall (x1 : elt) (s'0 : tree), Ok (Node c l y r) -> (x0, del x0 (Node c l y r)) = (x1, s'0) -> min_elt (Node c l y r) = Some x1 /\ del x1 (Node c l y r) = s'0

Some (x0, remove x0 (Node c l y r)) = Some (x, s') -> Some x0 = Some x /\ remove x (Node c l y r) = s'
inversion_clear 1; auto. Qed.
s:tree

remove_min s = None -> Empty s
s:tree

remove_min s = None -> Empty s
s:tree

match s with | Leaf => None | Node _ l x r => let (k, t) := delmin l x r in Some (k, makeBlack t) end = None -> Empty s

None = None -> Empty Leaf
c:Color.t
l:tree
y:X.t
r:tree
(let (k, t) := delmin l y r in Some (k, makeBlack t)) = None -> Empty (Node c l y r)

None = None -> Empty Leaf
easy.
c:Color.t
l:tree
y:X.t
r:tree

(let (k, t) := delmin l y r in Some (k, makeBlack t)) = None -> Empty (Node c l y r)
now destruct delmin. Qed.
s:t
H:Ok s

match remove_min s with | Some (_, s') => Ok s' | None => True end
s:t
H:Ok s

match remove_min s with | Some (_, s') => Ok s' | None => True end
s:t
H:Ok s

(forall (x : elt) (s' : tree), Ok s -> remove_min s = Some (x, s') -> min_elt s = Some x /\ remove x s = s') -> match remove_min s with | Some (_, s') => Ok s' | None => True end
s:t
H:Ok s
x0:elt
s0:tree

(forall (x : elt) (s' : tree), Ok s -> Some (x0, s0) = Some (x, s') -> min_elt s = Some x /\ remove x s = s') -> Ok s0
s:t
H:Ok s
x0:elt
s0:tree
R:forall (x : elt) (s' : tree), Ok s -> Some (x0, s0) = Some (x, s') -> min_elt s = Some x /\ remove x s = s'

Ok s0
s:t
H:Ok s
x0:elt
s0:tree
R:forall (x : elt) (s' : tree), Ok s -> Some (x0, s0) = Some (x, s') -> min_elt s = Some x /\ remove x s = s'
H0:min_elt s = Some x0
H1:remove x0 s = s0

Ok s0
s:t
H:Ok s
x0:elt
R:forall (x : elt) (s' : tree), Ok s -> Some (x0, remove x0 s) = Some (x, s') -> min_elt s = Some x /\ remove x s = s'
H0:min_elt s = Some x0

Ok (remove x0 s)
auto_tc. Qed.

Treeify

Notation ifpred p n := (if p then pred n else n%nat).

Definition treeify_invariant size (f:treeify_t) :=
 forall acc,
 size <= length acc ->
 let (t,acc') := f acc in
 cardinal t = size /\ acc = elements t ++ acc'.


treeify_invariant 0 treeify_zero

treeify_invariant 0 treeify_zero
acc:list elt

0 <= length acc -> let (t0, acc') := treeify_zero acc in cardinal t0 = 0 /\ acc = elements t0 ++ acc'
acc:list elt

0 <= length acc -> 0 = 0 /\ acc = acc
auto. Qed.

treeify_invariant 1 treeify_one

treeify_invariant 1 treeify_one
intros [|x acc]; simpl; auto; inversion 1. Qed.
f, g:treeify_t
size1, size2, size:nat

treeify_invariant size1 f -> treeify_invariant size2 g -> size = S (size1 + size2) -> treeify_invariant size (treeify_cont f g)
f, g:treeify_t
size1, size2, size:nat

treeify_invariant size1 f -> treeify_invariant size2 g -> size = S (size1 + size2) -> treeify_invariant size (treeify_cont f g)
f, g:treeify_t
size1, size2, size:nat
Hf:treeify_invariant size1 f
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
acc:list elt
LE:size <= length acc

let (t0, acc') := treeify_cont f g acc in cardinal t0 = size /\ acc = elements t0 ++ acc'
f, g:treeify_t
size1, size2, size:nat
Hf:treeify_invariant size1 f
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
acc:list elt
LE:size <= length acc

let (t0, acc') := let (l, l0) := f acc in match l0 with | nil => bogus | x :: acc0 => let (r, acc1) := g acc0 in (Bk l x r, acc1) end in cardinal t0 = size /\ acc = elements t0 ++ acc'
f, g:treeify_t
size1, size2, size:nat
acc:list elt
Hf:size1 <= length acc -> let (t0, acc') := f acc in cardinal t0 = size1 /\ acc = elements t0 ++ acc'
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc

let (t0, acc') := let (l, l0) := f acc in match l0 with | nil => bogus | x :: acc0 => let (r, acc1) := g acc0 in (Bk l x r, acc1) end in cardinal t0 = size /\ acc = elements t0 ++ acc'
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
acc1:list elt
Hf:size1 <= length acc -> cardinal t1 = size1 /\ acc = elements t1 ++ acc1
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc

let (t0, acc') := match acc1 with | nil => bogus | x :: acc0 => let (r, acc2) := g acc0 in (Bk t1 x r, acc2) end in cardinal t0 = size /\ acc = elements t0 ++ acc'
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
acc1:list elt
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc

size1 <= length acc
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
acc1:list elt
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ acc1
let (t0, acc') := match acc1 with | nil => bogus | x :: acc0 => let (r, acc2) := g acc0 in (Bk t1 x r, acc2) end in cardinal t0 = size /\ acc = elements t0 ++ acc'
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
acc1:list elt
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc

size1 <= length acc
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
acc1:list elt
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc

size1 <= size
f, g:treeify_t
size1, size2:nat
acc:list elt
t1:tree
acc1:list elt
Hg:treeify_invariant size2 g
LE:S (size1 + size2) <= length acc

size1 <= S (size1 + size2)
auto with arith.
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
acc1:list elt
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ acc1

let (t0, acc') := match acc1 with | nil => bogus | x :: acc0 => let (r, acc2) := g acc0 in (Bk t1 x r, acc2) end in cardinal t0 = size /\ acc = elements t0 ++ acc'
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ nil

let (t0, acc') := bogus in cardinal t0 = size /\ acc = elements t0 ++ acc'
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
x:elt
acc1:list elt
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ x :: acc1
let (t0, acc') := let (r, acc0) := g acc1 in (Bk t1 x r, acc0) in cardinal t0 = size /\ acc = elements t0 ++ acc'
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ nil

let (t0, acc') := bogus in cardinal t0 = size /\ acc = elements t0 ++ acc'
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ nil

False
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ nil

size <= length acc -> False
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ nil

length acc < size
f, g:treeify_t
size2:nat
t1:tree
Hg:treeify_invariant size2 g

length (elements t1 ++ nil) < S (cardinal t1 + size2)
rewrite app_nil_r, <- elements_cardinal; auto with arith.
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
x:elt
acc1:list elt
Hg:treeify_invariant size2 g
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ x :: acc1

let (t0, acc') := let (r, acc0) := g acc1 in (Bk t1 x r, acc0) in cardinal t0 = size /\ acc = elements t0 ++ acc'
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
x:elt
acc1:list elt
Hg:size2 <= length acc1 -> let (t0, acc') := g acc1 in cardinal t0 = size2 /\ acc1 = elements t0 ++ acc'
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ x :: acc1

let (t0, acc') := let (r, acc0) := g acc1 in (Bk t1 x r, acc0) in cardinal t0 = size /\ acc = elements t0 ++ acc'
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
x:elt
acc1:list elt
t2:tree
acc2:list elt
Hg:size2 <= length acc1 -> cardinal t2 = size2 /\ acc1 = elements t2 ++ acc2
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ x :: acc1

cardinal (Bk t1 x t2) = size /\ acc = elements (Bk t1 x t2) ++ acc2
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
x:elt
acc1:list elt
t2:tree
acc2:list elt
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ x :: acc1

size2 <= length acc1
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
x:elt
acc1:list elt
t2:tree
acc2:list elt
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ x :: acc1
Hg1:cardinal t2 = size2
Hg2:acc1 = elements t2 ++ acc2
cardinal (Bk t1 x t2) = size /\ acc = elements (Bk t1 x t2) ++ acc2
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
x:elt
acc1:list elt
t2:tree
acc2:list elt
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ x :: acc1

size2 <= length acc1
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
x:elt
acc1:list elt
t2:tree
acc2:list elt
EQ:size = S (size1 + size2)
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ x :: acc1

size <= length acc -> size2 <= length acc1
f, g:treeify_t
size2:nat
t1:tree
x:elt
acc1:list elt
t2:tree
acc2:list elt

S (cardinal t1 + size2) <= length (elements t1 ++ x :: acc1) -> size2 <= length acc1
f, g:treeify_t
size2:nat
t1:tree
x:elt
acc1:list elt
t2:tree
acc2:list elt

S (cardinal t1 + size2) <= cardinal t1 + length (x :: acc1) -> size2 <= length acc1
f, g:treeify_t
size2:nat
t1:tree
x:elt
acc1:list elt
t2:tree
acc2:list elt

S (cardinal t1 + size2) <= cardinal t1 + S (length acc1) -> size2 <= length acc1
f, g:treeify_t
size2:nat
t1:tree
x:elt
acc1:list elt
t2:tree
acc2:list elt

cardinal t1 + size2 <= cardinal t1 + length acc1 -> size2 <= length acc1
apply Nat.add_le_mono_l.
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
x:elt
acc1:list elt
t2:tree
acc2:list elt
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ x :: acc1
Hg1:cardinal t2 = size2
Hg2:acc1 = elements t2 ++ acc2

cardinal (Bk t1 x t2) = size /\ acc = elements (Bk t1 x t2) ++ acc2
f, g:treeify_t
size1, size2, size:nat
acc:list elt
t1:tree
x:elt
acc1:list elt
t2:tree
acc2:list elt
EQ:size = S (size1 + size2)
LE:size <= length acc
Hf1:cardinal t1 = size1
Hf2:acc = elements t1 ++ x :: acc1
Hg1:cardinal t2 = size2
Hg2:acc1 = elements t2 ++ acc2

cardinal (Bk t1 x t2) = size /\ acc = elements t1 ++ (x :: elements t2) ++ acc2
now subst. Qed.
n:positive
p:bool

treeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)
n:positive
p:bool

treeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)
n:positive

forall p : bool, treeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool

treeify_invariant (ifpred p (Pos.to_nat n~1)) (treeify_cont (treeify_aux false n) (treeify_aux p n))
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool
treeify_invariant (ifpred p (Pos.to_nat n~0)) (treeify_cont (treeify_aux p n) (treeify_aux true n))
p:bool
treeify_invariant (ifpred p (Pos.to_nat 1)) (if p then treeify_zero else treeify_one)
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool

treeify_invariant (ifpred p (Pos.to_nat n~1)) (treeify_cont (treeify_aux false n) (treeify_aux p n))
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool

ifpred p (Pos.to_nat n~1) = S (Pos.to_nat n + ifpred p (Pos.to_nat n))
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool

ifpred p (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred p (Pos.to_nat n))
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool
H:0 < Pos.to_nat n

ifpred p (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred p (Pos.to_nat n))
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool
H:Pos.to_nat n <> 0

ifpred p (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred p (Pos.to_nat n))
n:positive
IHn:forall p : bool, treeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)
H:Pos.to_nat n <> 0

Pos.to_nat n + Pos.to_nat n = S (Pos.to_nat n + Init.Nat.pred (Pos.to_nat n))
now rewrite <- Nat.add_succ_r, Nat.succ_pred; trivial.
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool

treeify_invariant (ifpred p (Pos.to_nat n~0)) (treeify_cont (treeify_aux p n) (treeify_aux true n))
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool

ifpred p (Pos.to_nat n~0) = S (ifpred p (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool

ifpred p (2 * Pos.to_nat n) = S (ifpred p (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool
H:0 < Pos.to_nat n

ifpred p (2 * Pos.to_nat n) = S (ifpred p (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool
H:Pos.to_nat n <> 0

ifpred p (2 * Pos.to_nat n) = S (ifpred p (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))
n:positive
IHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)
p:bool
H:Pos.to_nat n <> 0

ifpred p (2 * Pos.to_nat n) = ifpred p (Pos.to_nat n) + Pos.to_nat n
n:positive
IHn:forall p : bool, treeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)
H:Pos.to_nat n <> 0

Init.Nat.pred (Pos.to_nat n + Pos.to_nat n) = Init.Nat.pred (Pos.to_nat n) + Pos.to_nat n
n:positive
IHn:forall p : bool, treeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)
H:Pos.to_nat n <> 0

Init.Nat.pred (Pos.to_nat n) + Pos.to_nat n = Init.Nat.pred (Pos.to_nat n + Pos.to_nat n)
now apply Nat.add_pred_l.
p:bool

treeify_invariant (ifpred p (Pos.to_nat 1)) (if p then treeify_zero else treeify_one)
destruct p; [ apply treeify_zero_spec | apply treeify_one_spec ]. Qed.
l:list elt
p:positive

Pos.to_nat (plength_aux l p) = length l + Pos.to_nat p
l:list elt
p:positive

Pos.to_nat (plength_aux l p) = length l + Pos.to_nat p
l:list elt

forall p : positive, Pos.to_nat (plength_aux l p) = length l + Pos.to_nat p
a:elt
l:list elt
IHl:forall p : positive, Pos.to_nat (plength_aux l p) = length l + Pos.to_nat p

forall p : positive, Pos.to_nat (plength_aux (a :: l) p) = length (a :: l) + Pos.to_nat p
a:elt
l:list elt
IHl:forall p : positive, Pos.to_nat (plength_aux l p) = length l + Pos.to_nat p

forall p : positive, Pos.to_nat (plength_aux l (Pos.succ p)) = length (a :: l) + Pos.to_nat p
a:elt
l:list elt
IHl:forall p0 : positive, Pos.to_nat (plength_aux l p0) = length l + Pos.to_nat p0
p:positive

Pos.to_nat (plength_aux l (Pos.succ p)) = length (a :: l) + Pos.to_nat p
now rewrite IHl, Pos2Nat.inj_succ, Nat.add_succ_r. Qed.
l:list elt

Pos.to_nat (plength l) = S (length l)
l:list elt

Pos.to_nat (plength l) = S (length l)
l:list elt

Pos.to_nat (plength_aux l 1) = S (length l)
l:list elt

length l + Pos.to_nat 1 = S (length l)
apply Nat.add_1_r. Qed.
l:list elt

elements (treeify l) = l
l:list elt

elements (treeify l) = l
l:list elt
H:ifpred true (Pos.to_nat (plength l)) <= length l -> let (t0, acc') := treeify_aux true (plength l) l in cardinal t0 = ifpred true (Pos.to_nat (plength l)) /\ l = elements t0 ++ acc'

elements (treeify l) = l
l:list elt
H:ifpred true (Pos.to_nat (plength l)) <= length l -> let (t0, acc') := treeify_aux true (plength l) l in cardinal t0 = ifpred true (Pos.to_nat (plength l)) /\ l = elements t0 ++ acc'

elements (fst (treeify_aux true (plength l) l)) = l
l:list elt
t:tree
acc:list elt
H:Init.Nat.pred (Pos.to_nat (plength l)) <= length l -> cardinal t = Init.Nat.pred (Pos.to_nat (plength l)) /\ l = elements t ++ acc

elements t = l
l:list elt
t:tree
acc:list elt

Init.Nat.pred (Pos.to_nat (plength l)) <= length l
l:list elt
t:tree
acc:list elt
H:cardinal t = Init.Nat.pred (Pos.to_nat (plength l))
H':l = elements t ++ acc
elements t = l
l:list elt
t:tree
acc:list elt

Init.Nat.pred (Pos.to_nat (plength l)) <= length l
now rewrite plength_spec.
l:list elt
t:tree
acc:list elt
H:cardinal t = Init.Nat.pred (Pos.to_nat (plength l))
H':l = elements t ++ acc

elements t = l
t:tree
acc:list elt
H:cardinal t = Init.Nat.pred (Pos.to_nat (plength (elements t ++ acc)))

elements t = elements t ++ acc
t:tree
acc:list elt
H:cardinal t = Init.Nat.pred (S (cardinal t + length acc))

elements t = elements t ++ acc
t:tree
H:cardinal t = Init.Nat.pred (S (cardinal t + length nil))

elements t = elements t ++ nil
t:tree
e:elt
acc:list elt
H:cardinal t = Init.Nat.pred (S (cardinal t + length (e :: acc)))
elements t = elements t ++ e :: acc
t:tree
H:cardinal t = Init.Nat.pred (S (cardinal t + length nil))

elements t = elements t ++ nil
now rewrite app_nil_r.
t:tree
e:elt
acc:list elt
H:cardinal t = Init.Nat.pred (S (cardinal t + length (e :: acc)))

elements t = elements t ++ e :: acc
t:tree
e:elt
acc:list elt
H:cardinal t = Init.Nat.pred (S (cardinal t + length (e :: acc)))

False
t:tree
e:elt
acc:list elt

cardinal t = Init.Nat.pred (S (cardinal t + length (e :: acc))) -> False
t:tree
e:elt
acc:list elt

cardinal t = cardinal t + S (length acc) -> False
t:tree
e:elt
acc:list elt

cardinal t = S (length acc + cardinal t) -> False
apply Nat.succ_add_discr. Qed.
x:elt
l:list elt

InT x (treeify l) <-> InA X.eq x l
x:elt
l:list elt

InT x (treeify l) <-> InA X.eq x l
x:elt
l:list elt

InT x (treeify l) <-> InA X.eq x l
now rewrite <- elements_spec1, treeify_elements. Qed.
l:list X.t

Sorted X.lt l -> Ok (treeify l)
l:list X.t

Sorted X.lt l -> Ok (treeify l)
l:list X.t
H:Sorted X.lt l

Ok (treeify l)
l:list X.t
H:Sorted X.lt l

Sorted X.lt (elements (treeify l))
rewrite treeify_elements; auto. Qed.

Filter

A:Type
f:A -> bool
l, l':list A

List.filter f (l ++ l') = List.filter f l ++ List.filter f l'
A:Type
f:A -> bool
l, l':list A

List.filter f (l ++ l') = List.filter f l ++ List.filter f l'
A:Type
f:A -> bool
x:A
l, l':list A
IH:List.filter f (l ++ l') = List.filter f l ++ List.filter f l'

(if f x then x :: List.filter f (l ++ l') else List.filter f (l ++ l')) = (if f x then x :: List.filter f l else List.filter f l) ++ List.filter f l'
destruct (f x); simpl; now rewrite IH. Qed.
s:tree
f:elt -> bool
acc:list X.t

filter_aux f s acc = List.filter f (elements s) ++ acc
s:tree
f:elt -> bool
acc:list X.t

filter_aux f s acc = List.filter f (elements s) ++ acc
s:tree
f:elt -> bool

forall acc : list X.t, filter_aux f s acc = List.filter f (elements s) ++ acc
c:Color.t
l:tree
x:X.t
r:tree
f:elt -> bool
IHl:forall acc : list X.t, filter_aux f l acc = List.filter f (elements l) ++ acc
IHr:forall acc : list X.t, filter_aux f r acc = List.filter f (elements r) ++ acc

forall acc : list X.t, filter_aux f (Node c l x r) acc = List.filter f (elements (Node c l x r)) ++ acc
c:Color.t
l:tree
x:X.t
r:tree
f:elt -> bool
IHl:forall acc0 : list X.t, filter_aux f l acc0 = List.filter f (elements l) ++ acc0
IHr:forall acc0 : list X.t, filter_aux f r acc0 = List.filter f (elements r) ++ acc0
acc:list X.t

filter_aux f (Node c l x r) acc = List.filter f (elements (Node c l x r)) ++ acc
c:Color.t
l:tree
x:X.t
r:tree
f:elt -> bool
IHl:forall acc0 : list X.t, filter_aux f l acc0 = List.filter f (elements l) ++ acc0
IHr:forall acc0 : list X.t, filter_aux f r acc0 = List.filter f (elements r) ++ acc0
acc:list X.t

filter_aux f (Node c l x r) acc = (List.filter f (elements l) ++ List.filter f (x :: elements r)) ++ acc
c:Color.t
l:tree
x:X.t
r:tree
f:elt -> bool
IHl:forall acc0 : list X.t, filter_aux f l acc0 = List.filter f (elements l) ++ acc0
IHr:forall acc0 : list X.t, filter_aux f r acc0 = List.filter f (elements r) ++ acc0
acc:list X.t

(if f x then filter_aux f l (x :: filter_aux f r acc) else filter_aux f l (filter_aux f r acc)) = (List.filter f (elements l) ++ (if f x then x :: List.filter f (elements r) else List.filter f (elements r))) ++ acc
destruct (f x); now rewrite IHl, IHr, app_ass. Qed.
s:t
f:elt -> bool

elements (filter f s) = List.filter f (elements s)
s:t
f:elt -> bool

elements (filter f s) = List.filter f (elements s)
s:t
f:elt -> bool

elements (treeify (filter_aux f s nil)) = List.filter f (elements s)
now rewrite treeify_elements, filter_aux_elements, app_nil_r. Qed.
s:t
x:elt
f:X.t -> bool

Proper (X.eq ==> Logic.eq) f -> InT x (filter f s) <-> InT x s /\ f x = true
s:t
x:elt
f:X.t -> bool

Proper (X.eq ==> Logic.eq) f -> InT x (filter f s) <-> InT x s /\ f x = true
s:t
x:elt
f:X.t -> bool
Hf:Proper (X.eq ==> Logic.eq) f

InT x (filter f s) <-> InT x s /\ f x = true
rewrite <- elements_spec1, filter_elements, filter_InA, elements_spec1; now auto_tc. Qed.
s:tree
f:elt -> bool
H:Ok s

Ok (filter f s)
s:tree
f:elt -> bool
H:Ok s

Ok (filter f s)
s:tree
f:elt -> bool
H:Ok s

Sorted X.lt (elements (filter f s))
s:tree
f:elt -> bool
H:Ok s

Sorted X.lt (List.filter f (elements s))
apply filter_sort with X.eq; auto_tc. Qed.

Partition

s:tree
f:elt -> bool
acc1, acc2:list X.t

partition_aux f s acc1 acc2 = (filter_aux f s acc1, filter_aux (fun x : elt => negb (f x)) s acc2)
s:tree
f:elt -> bool
acc1, acc2:list X.t

partition_aux f s acc1 acc2 = (filter_aux f s acc1, filter_aux (fun x : elt => negb (f x)) s acc2)
s:tree
f:elt -> bool

forall acc1 acc2 : list X.t, partition_aux f s acc1 acc2 = (filter_aux f s acc1, filter_aux (fun x : elt => negb (f x)) s acc2)
f:elt -> bool

forall acc1 acc2 : list X.t, (acc1, acc2) = (acc1, acc2)
c:Color.t
l:tree
x:X.t
r:tree
f:elt -> bool
Hl:forall acc1 acc2 : list X.t, partition_aux f l acc1 acc2 = (filter_aux f l acc1, filter_aux (fun x0 : elt => negb (f x0)) l acc2)
Hr:forall acc1 acc2 : list X.t, partition_aux f r acc1 acc2 = (filter_aux f r acc1, filter_aux (fun x0 : elt => negb (f x0)) r acc2)
forall acc1 acc2 : list X.t, (let (acc0, acc3) := partition_aux f r acc1 acc2 in if f x then partition_aux f l (x :: acc0) acc3 else partition_aux f l acc0 (x :: acc3)) = (if f x then filter_aux f l (x :: filter_aux f r acc1) else filter_aux f l (filter_aux f r acc1), if negb (f x) then filter_aux (fun x0 : elt => negb (f x0)) l (x :: filter_aux (fun x0 : elt => negb (f x0)) r acc2) else filter_aux (fun x0 : elt => negb (f x0)) l (filter_aux (fun x0 : elt => negb (f x0)) r acc2))
f:elt -> bool

forall acc1 acc2 : list X.t, (acc1, acc2) = (acc1, acc2)
trivial.
c:Color.t
l:tree
x:X.t
r:tree
f:elt -> bool
Hl:forall acc1 acc2 : list X.t, partition_aux f l acc1 acc2 = (filter_aux f l acc1, filter_aux (fun x0 : elt => negb (f x0)) l acc2)
Hr:forall acc1 acc2 : list X.t, partition_aux f r acc1 acc2 = (filter_aux f r acc1, filter_aux (fun x0 : elt => negb (f x0)) r acc2)

forall acc1 acc2 : list X.t, (let (acc0, acc3) := partition_aux f r acc1 acc2 in if f x then partition_aux f l (x :: acc0) acc3 else partition_aux f l acc0 (x :: acc3)) = (if f x then filter_aux f l (x :: filter_aux f r acc1) else filter_aux f l (filter_aux f r acc1), if negb (f x) then filter_aux (fun x0 : elt => negb (f x0)) l (x :: filter_aux (fun x0 : elt => negb (f x0)) r acc2) else filter_aux (fun x0 : elt => negb (f x0)) l (filter_aux (fun x0 : elt => negb (f x0)) r acc2))
c:Color.t
l:tree
x:X.t
r:tree
f:elt -> bool
Hl:forall acc0 acc3 : list X.t, partition_aux f l acc0 acc3 = (filter_aux f l acc0, filter_aux (fun x0 : elt => negb (f x0)) l acc3)
Hr:forall acc0 acc3 : list X.t, partition_aux f r acc0 acc3 = (filter_aux f r acc0, filter_aux (fun x0 : elt => negb (f x0)) r acc3)
acc1, acc2:list X.t

(let (acc0, acc3) := partition_aux f r acc1 acc2 in if f x then partition_aux f l (x :: acc0) acc3 else partition_aux f l acc0 (x :: acc3)) = (if f x then filter_aux f l (x :: filter_aux f r acc1) else filter_aux f l (filter_aux f r acc1), if negb (f x) then filter_aux (fun x0 : elt => negb (f x0)) l (x :: filter_aux (fun x0 : elt => negb (f x0)) r acc2) else filter_aux (fun x0 : elt => negb (f x0)) l (filter_aux (fun x0 : elt => negb (f x0)) r acc2))
destruct (f x); simpl; now rewrite Hr, Hl. Qed.
s:t
f:elt -> bool

partition f s = (filter f s, filter (fun x : elt => negb (f x)) s)
s:t
f:elt -> bool

partition f s = (filter f s, filter (fun x : elt => negb (f x)) s)
s:t
f:elt -> bool

(let (ok0, ko) := partition_aux f s nil nil in (treeify ok0, treeify ko)) = (treeify (filter_aux f s nil), treeify (filter_aux (fun x : elt => negb (f x)) s nil))
now rewrite partition_aux_spec. Qed.
s:t
f:X.t -> bool

Proper (X.eq ==> Logic.eq) f -> Equal (fst (partition f s)) (filter f s)
s:t
f:X.t -> bool

Proper (X.eq ==> Logic.eq) f -> Equal (fst (partition f s)) (filter f s)
now rewrite partition_spec. Qed.
s:t
f:X.t -> bool

Proper (X.eq ==> Logic.eq) f -> Equal (snd (partition f s)) (filter (fun x : elt => negb (f x)) s)
s:t
f:X.t -> bool

Proper (X.eq ==> Logic.eq) f -> Equal (snd (partition f s)) (filter (fun x : elt => negb (f x)) s)
now rewrite partition_spec. Qed.
s:tree
f:elt -> bool
H:Ok s

Ok (fst (partition f s))
s:tree
f:elt -> bool
H:Ok s

Ok (fst (partition f s))
rewrite partition_spec; now apply filter_ok. Qed.
s:tree
f:elt -> bool
H:Ok s

Ok (snd (partition f s))
s:tree
f:elt -> bool
H:Ok s

Ok (snd (partition f s))
rewrite partition_spec; now apply filter_ok. Qed.

An invariant for binary list functions with accumulator.

Ltac inA :=
 rewrite ?InA_app_iff, ?InA_cons, ?InA_nil, ?InA_rev in *; auto_tc.

Record INV l1 l2 acc : Prop := {
 l1_sorted : sort X.lt (rev l1);
 l2_sorted : sort X.lt (rev l2);
 acc_sorted : sort X.lt acc;
 l1_lt_acc x y : InA X.eq x l1 -> InA X.eq y acc -> X.lt x y;
 l2_lt_acc x y : InA X.eq x l2 -> InA X.eq y acc -> X.lt x y}.
Local Hint Resolve l1_sorted l2_sorted acc_sorted : core.

s1, s2:tree
H:Ok s1
H0:Ok s2

INV (rev_elements s1) (rev_elements s2) nil
s1, s2:tree
H:Ok s1
H0:Ok s2

INV (rev_elements s1) (rev_elements s2) nil
s1, s2:tree
H:Ok s1
H0:Ok s2

INV (rev (elements s1)) (rev (elements s2)) nil
split; rewrite ?rev_involutive; auto; intros; now inA. Qed.
l1, l2, acc:list X.t

INV l1 l2 acc -> INV l2 l1 acc
l1, l2, acc:list X.t

INV l1 l2 acc -> INV l2 l1 acc
destruct 1; now split. Qed.
x1:X.t
l1, l2, acc:list X.t

INV (x1 :: l1) l2 acc -> INV l1 l2 acc
x1:X.t
l1, l2, acc:list X.t

INV (x1 :: l1) l2 acc -> INV l1 l2 acc
x1:X.t
l1, l2, acc:list X.t
l1s:Sorted X.lt (rev (x1 :: l1))
l2s:Sorted X.lt (rev l2)
accs:Sorted X.lt acc
l1a:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
l2a:forall x y : X.t, InA X.eq x l2 -> InA X.eq y acc -> X.lt x y

INV l1 l2 acc
x1:X.t
l1, l2, acc:list X.t
l1s:Sorted X.lt (rev l1 ++ x1 :: nil)
l2s:Sorted X.lt (rev l2)
accs:Sorted X.lt acc
l1a:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
l2a:forall x y : X.t, InA X.eq x l2 -> InA X.eq y acc -> X.lt x y

INV l1 l2 acc
x1:X.t
l1, l2, acc:list X.t
l1s:Sorted X.lt (rev l1 ++ x1 :: nil)
l2s:Sorted X.lt (rev l2)
accs:Sorted X.lt acc
l1a:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
l2a:forall x y : X.t, InA X.eq x l2 -> InA X.eq y acc -> X.lt x y
U:Sorted X.lt (rev l1)
V:Sorted X.lt (x1 :: nil)
W:forall x0 x2 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x2 (x1 :: nil) -> X.lt x0 x2

INV l1 l2 acc
split; auto. Qed.
x1, x2:X.t
l1, l2, acc:list X.t

INV (x1 :: l1) (x2 :: l2) acc -> X.eq x1 x2 -> INV l1 l2 (x1 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t

INV (x1 :: l1) (x2 :: l2) acc -> X.eq x1 x2 -> INV l1 l2 (x1 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev (x1 :: l1))
V:Sorted X.lt (rev (x2 :: l2))
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.eq x1 x2

INV l1 l2 (x1 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.eq x1 x2

INV l1 l2 (x1 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3

INV l1 l2 (x1 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

INV l1 l2 (x1 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

Sorted X.lt (x1 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
forall x y : X.t, InA X.eq x l1 -> InA X.eq y (x1 :: acc) -> X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
forall x y : X.t, InA X.eq x l2 -> InA X.eq y (x1 :: acc) -> X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

Sorted X.lt (x1 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

HdRel X.lt x1 acc
apply InA_InfA with X.eq; auto_tc.
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

forall x y : X.t, InA X.eq x l1 -> InA X.eq y (x1 :: acc) -> X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l1
Hy:X.eq y x1

X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l1
Hy:InA X.eq y acc
X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l1
Hy:X.eq y x1

X.lt x y
apply U3; inA.
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l1
Hy:InA X.eq y acc

X.lt x y
apply X; inA.
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

forall x y : X.t, InA X.eq x l2 -> InA X.eq y (x1 :: acc) -> X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l2
Hy:X.eq y x1

X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l2
Hy:InA X.eq y acc
X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l2
Hy:X.eq y x1

X.lt x y
rewrite Hy, EQ; apply V3; inA.
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.eq x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l2
Hy:InA X.eq y acc

X.lt x y
apply Y; inA. Qed.
x1, x2:X.t
l1, l2, acc:list X.t

INV (x1 :: l1) (x2 :: l2) acc -> X.lt x1 x2 -> INV (x1 :: l1) l2 (x2 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t

INV (x1 :: l1) (x2 :: l2) acc -> X.lt x1 x2 -> INV (x1 :: l1) l2 (x2 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev (x1 :: l1))
V:Sorted X.lt (rev (x2 :: l2))
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.lt x1 x2

INV (x1 :: l1) l2 (x2 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.lt x1 x2

INV (x1 :: l1) l2 (x2 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3

INV (x1 :: l1) l2 (x2 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

INV (x1 :: l1) l2 (x2 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

Sorted X.lt (x2 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y (x2 :: acc) -> X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
forall x y : X.t, InA X.eq x l2 -> InA X.eq y (x2 :: acc) -> X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

Sorted X.lt (x2 :: acc)
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

HdRel X.lt x2 acc
apply InA_InfA with X.eq; auto_tc.
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y (x2 :: acc) -> X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:X.eq x x1 \/ InA X.eq x l1
Hy:X.eq y x2

X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:X.eq x x1 \/ InA X.eq x l1
Hy:InA X.eq y acc
X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:X.eq x x1 \/ InA X.eq x l1
Hy:X.eq y x2

X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:X.eq x x1 \/ InA X.eq x l1

X.lt x x2
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
H:InA X.eq x l1

X.lt x x2
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
H:InA X.eq x l1

X.lt x x1
apply U3; inA.
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:X.eq x x1 \/ InA X.eq x l1
Hy:InA X.eq y acc

X.lt x y
apply X; inA.
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x y
Y:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x y
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

forall x y : X.t, InA X.eq x l2 -> InA X.eq y (x2 :: acc) -> X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l2
Hy:X.eq y x2

X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l2
Hy:InA X.eq y acc
X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l2
Hy:X.eq y x2

X.lt x y
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l2
Hy:X.eq y x2

X.lt x x2
apply V3; inA.
x1, x2:X.t
l1, l2, acc:list X.t
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
W:Sorted X.lt acc
X:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0
Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0
EQ:X.lt x1 x2
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
x, y:X.t
Hx:InA X.eq x l2
Hy:InA X.eq y acc

X.lt x y
apply Y; inA. Qed.
l1, l2, acc:list X.t

INV l1 l2 acc -> Sorted X.lt (rev_append l1 acc)
l1, l2, acc:list X.t

INV l1 l2 acc -> Sorted X.lt (rev_append l1 acc)
l1, l2, acc:list X.t
H:INV l1 l2 acc

Sorted X.lt (rev_append l1 acc)
l1, l2, acc:list X.t
H:INV l1 l2 acc

Sorted X.lt (rev l1 ++ acc)
l1, l2, acc:list X.t
H:INV l1 l2 acc

forall x y : X.t, InA X.eq x (rev l1) -> InA X.eq y acc -> X.lt x y
l1, l2, acc:list X.t
H:INV l1 l2 acc
x, y:X.t

InA X.eq x (rev l1) -> InA X.eq y acc -> X.lt x y
l1, l2, acc:list X.t
H:INV l1 l2 acc
x, y:X.t

InA X.eq x l1 -> InA X.eq y acc -> X.lt x y
eapply @l1_lt_acc; eauto. Qed.

union

l1, l2, acc:list X.t

INV l1 l2 acc -> Sorted X.lt (union_list l1 l2 acc)
l1, l2, acc:list X.t

INV l1 l2 acc -> Sorted X.lt (union_list l1 l2 acc)
l1:list X.t

forall l2 acc : list X.t, INV l1 l2 acc -> Sorted X.lt (union_list l1 l2 acc)
l2, acc:list X.t
inv:INV nil l2 acc

Sorted X.lt (union_list nil l2 acc)
x1:X.t
l1:list X.t
IH1:forall l2 acc0 : list X.t, INV l1 l2 acc0 -> Sorted X.lt (union_list l1 l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) nil acc
Sorted X.lt (union_list (x1 :: l1) nil acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
Sorted X.lt (union_list (x1 :: l1) (x2 :: l2) acc)
l2, acc:list X.t
inv:INV nil l2 acc

Sorted X.lt (union_list nil l2 acc)
eapply INV_rev, INV_sym; eauto.
x1:X.t
l1:list X.t
IH1:forall l2 acc0 : list X.t, INV l1 l2 acc0 -> Sorted X.lt (union_list l1 l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) nil acc

Sorted X.lt (union_list (x1 :: l1) nil acc)
eapply INV_rev; eauto.
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc

Sorted X.lt (union_list (x1 :: l1) (x2 :: l2) acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc

Sorted X.lt match X.compare x1 x2 with | Eq => union_list l1 l2 (x1 :: acc) | Lt => (fix union_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => union_list l1 l2' (x1 :: acc0) | Lt => union_l1 l2' (y :: acc0) | Gt => union_list l1 l0 (x1 :: acc0) end end) l2 (x2 :: acc) | Gt => union_list l1 (x2 :: l2) (x1 :: acc) end
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.eq x1 x2

Sorted X.lt (union_list l1 l2 (x1 :: acc))
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x1 x2
Sorted X.lt ((fix union_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => union_list l1 l2' (x1 :: acc0) | Lt => union_l1 l2' (y :: acc0) | Gt => union_list l1 l0 (x1 :: acc0) end end) l2 (x2 :: acc))
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x2 x1
Sorted X.lt (union_list l1 (x2 :: l2) (x1 :: acc))
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.eq x1 x2

Sorted X.lt (union_list l1 l2 (x1 :: acc))
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.eq x1 x2

INV l1 l2 (x1 :: acc)
eapply INV_eq; eauto.
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x1 x2

Sorted X.lt ((fix union_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => union_list l1 l2' (x1 :: acc0) | Lt => union_l1 l2' (y :: acc0) | Gt => union_list l1 l0 (x1 :: acc0) end end) l2 (x2 :: acc))
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x1 x2

INV (x1 :: l1) l2 (x2 :: acc)
eapply INV_lt; eauto.
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x2 x1

Sorted X.lt (union_list l1 (x2 :: l2) (x1 :: acc))
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x2 x1

INV l1 (x2 :: l2) (x1 :: acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x2 x1

INV (x2 :: l2) (x1 :: l1) acc
now apply INV_sym. Qed.
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (linear_union s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (linear_union s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (treeify (union_list (rev_elements s1) (rev_elements s2) nil))
now apply treeify_ok, union_list_ok, INV_init. Qed.
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (fold add s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (fold add s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (fold_right (fun (y : elt) (x : tree) => flip add x y) s2 (rev (elements s1)))
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (fold_right (fun (y : X.t) (x : tree) => flip add x y) s2 (rev (elements s1)))
induction (rev (elements s1)); simpl; unfold flip in *; auto_tc. Qed.
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (union s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (union s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok match compare_height s1 s1 s2 s2 with | Eq => linear_union s1 s2 | Lt => fold add s1 s2 | Gt => fold add s2 s1 end
destruct compare_height; auto_tc. Qed.
x:X.t
l1, l2, acc:list elt

InA X.eq x (union_list l1 l2 acc) <-> InA X.eq x l1 \/ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l1, l2, acc:list elt

InA X.eq x (union_list l1 l2 acc) <-> InA X.eq x l1 \/ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l1:list elt

forall l2 acc : list elt, InA X.eq x (union_list l1 l2 acc) <-> InA X.eq x l1 \/ InA X.eq x l2 \/ InA X.eq x acc
x:X.t

forall l2 acc : list elt, InA X.eq x (union_list nil l2 acc) <-> InA X.eq x nil \/ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
x1:elt
l1:list elt
IH1:forall l2 acc : list elt, InA X.eq x (union_list l1 l2 acc) <-> InA X.eq x l1 \/ InA X.eq x l2 \/ InA X.eq x acc
forall l2 acc : list elt, InA X.eq x (union_list (x1 :: l1) l2 acc) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x l2 \/ InA X.eq x acc
x:X.t

forall l2 acc : list elt, InA X.eq x (union_list nil l2 acc) <-> InA X.eq x nil \/ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l2, acc:list elt

InA X.eq x (rev_append l2 acc) <-> InA X.eq x nil \/ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l2, acc:list elt

InA X.eq x (rev l2 ++ acc) <-> InA X.eq x nil \/ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l2, acc:list elt

InA X.eq x l2 \/ InA X.eq x acc <-> False \/ InA X.eq x l2 \/ InA X.eq x acc
tauto.
x:X.t
x1:elt
l1:list elt
IH1:forall l2 acc : list elt, InA X.eq x (union_list l1 l2 acc) <-> InA X.eq x l1 \/ InA X.eq x l2 \/ InA X.eq x acc

forall l2 acc : list elt, InA X.eq x (union_list (x1 :: l1) l2 acc) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
x1:elt
l1:list elt
IH1:forall l2 acc0 : list elt, InA X.eq x (union_list l1 l2 acc0) <-> InA X.eq x l1 \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

InA X.eq x (rev_append l1 (x1 :: acc)) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x nil \/ InA X.eq x acc
x:X.t
x1:elt
l1:list elt
IH1:forall l0 acc0 : list elt, InA X.eq x (union_list l1 l0 acc0) <-> InA X.eq x l1 \/ InA X.eq x l0 \/ InA X.eq x acc0
x2:elt
l2:list elt
IH2:forall acc0 : list elt, InA X.eq x (union_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
InA X.eq x match X.compare x1 x2 with | Eq => union_list l1 l2 (x1 :: acc) | Lt => (fix union_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => union_list l1 l2' (x1 :: acc0) | Lt => union_l1 l2' (y :: acc0) | Gt => union_list l1 l0 (x1 :: acc0) end end) l2 (x2 :: acc) | Gt => union_list l1 (x2 :: l2) (x1 :: acc) end <-> InA X.eq x (x1 :: l1) \/ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x:X.t
x1:elt
l1:list elt
IH1:forall l2 acc0 : list elt, InA X.eq x (union_list l1 l2 acc0) <-> InA X.eq x l1 \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

InA X.eq x (rev_append l1 (x1 :: acc)) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x nil \/ InA X.eq x acc
x:X.t
x1:elt
l1:list elt
IH1:forall l2 acc0 : list elt, InA X.eq x (union_list l1 l2 acc0) <-> InA X.eq x l1 \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

InA X.eq x (rev l1 ++ x1 :: acc) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x nil \/ InA X.eq x acc
x:X.t
x1:elt
l1:list elt
IH1:forall l2 acc0 : list elt, InA X.eq x (union_list l1 l2 acc0) <-> InA X.eq x l1 \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

InA X.eq x l1 \/ X.eq x x1 \/ InA X.eq x acc <-> (X.eq x x1 \/ InA X.eq x l1) \/ False \/ InA X.eq x acc
tauto.
x:X.t
x1:elt
l1:list elt
IH1:forall l0 acc0 : list elt, InA X.eq x (union_list l1 l0 acc0) <-> InA X.eq x l1 \/ InA X.eq x l0 \/ InA X.eq x acc0
x2:elt
l2:list elt
IH2:forall acc0 : list elt, InA X.eq x (union_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

InA X.eq x match X.compare x1 x2 with | Eq => union_list l1 l2 (x1 :: acc) | Lt => (fix union_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => union_list l1 l2' (x1 :: acc0) | Lt => union_l1 l2' (y :: acc0) | Gt => union_list l1 l0 (x1 :: acc0) end end) l2 (x2 :: acc) | Gt => union_list l1 (x2 :: l2) (x1 :: acc) end <-> InA X.eq x (x1 :: l1) \/ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x:X.t
x1:elt
l1:list elt
IH1:forall l0 acc0 : list elt, InA X.eq x (union_list l1 l0 acc0) <-> InA X.eq x l1 \/ InA X.eq x l0 \/ InA X.eq x acc0
x2:elt
l2:list elt
IH2:forall acc0 : list elt, InA X.eq x (union_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
C:X.eq x1 x2

InA X.eq x (union_list l1 l2 (x1 :: acc)) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x:X.t
x1:elt
l1:list elt
IH1:forall l0 acc0 : list elt, InA X.eq x (union_list l1 l0 acc0) <-> InA X.eq x l1 \/ InA X.eq x l0 \/ InA X.eq x acc0
x2:elt
l2:list elt
IH2:forall acc0 : list elt, InA X.eq x (union_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
C:X.lt x1 x2
InA X.eq x ((fix union_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => union_list l1 l2' (x1 :: acc0) | Lt => union_l1 l2' (y :: acc0) | Gt => union_list l1 l0 (x1 :: acc0) end end) l2 (x2 :: acc)) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x:X.t
x1:elt
l1:list elt
IH1:forall l0 acc0 : list elt, InA X.eq x (union_list l1 l0 acc0) <-> InA X.eq x l1 \/ InA X.eq x l0 \/ InA X.eq x acc0
x2:elt
l2:list elt
IH2:forall acc0 : list elt, InA X.eq x (union_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
C:X.lt x2 x1
InA X.eq x (union_list l1 (x2 :: l2) (x1 :: acc)) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x:X.t
x1:elt
l1:list elt
IH1:forall l0 acc0 : list elt, InA X.eq x (union_list l1 l0 acc0) <-> InA X.eq x l1 \/ InA X.eq x l0 \/ InA X.eq x acc0
x2:elt
l2:list elt
IH2:forall acc0 : list elt, InA X.eq x (union_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
C:X.eq x1 x2

InA X.eq x (union_list l1 l2 (x1 :: acc)) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
rewrite IH1, !InA_cons, C; tauto.
x:X.t
x1:elt
l1:list elt
IH1:forall l0 acc0 : list elt, InA X.eq x (union_list l1 l0 acc0) <-> InA X.eq x l1 \/ InA X.eq x l0 \/ InA X.eq x acc0
x2:elt
l2:list elt
IH2:forall acc0 : list elt, InA X.eq x (union_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
C:X.lt x1 x2

InA X.eq x ((fix union_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => union_list l1 l2' (x1 :: acc0) | Lt => union_l1 l2' (y :: acc0) | Gt => union_list l1 l0 (x1 :: acc0) end end) l2 (x2 :: acc)) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x:X.t
x1:elt
l1:list elt
IH1:forall l0 acc0 : list elt, InA X.eq x (union_list l1 l0 acc0) <-> InA X.eq x l1 \/ InA X.eq x l0 \/ InA X.eq x acc0
x2:elt
l2:list elt
IH2:forall acc0 : list elt, InA X.eq x (union_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
C:X.lt x1 x2

(X.eq x x1 \/ InA X.eq x l1) \/ InA X.eq x l2 \/ X.eq x x2 \/ InA X.eq x acc <-> (X.eq x x1 \/ InA X.eq x l1) \/ (X.eq x x2 \/ InA X.eq x l2) \/ InA X.eq x acc
tauto.
x:X.t
x1:elt
l1:list elt
IH1:forall l0 acc0 : list elt, InA X.eq x (union_list l1 l0 acc0) <-> InA X.eq x l1 \/ InA X.eq x l0 \/ InA X.eq x acc0
x2:elt
l2:list elt
IH2:forall acc0 : list elt, InA X.eq x (union_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
C:X.lt x2 x1

InA X.eq x (union_list l1 (x2 :: l2) (x1 :: acc)) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
rewrite IH1, !InA_cons; tauto. Qed.
s1, s2:tree
x:elt

InT x (linear_union s1 s2) <-> InT x s1 \/ InT x s2
s1, s2:tree
x:elt

InT x (linear_union s1 s2) <-> InT x s1 \/ InT x s2
s1, s2:tree
x:elt

InT x (treeify (union_list (rev_elements s1) (rev_elements s2) nil)) <-> InT x s1 \/ InT x s2
s1, s2:tree
x:elt

InA X.eq x (rev (elements s1)) \/ InA X.eq x (rev (elements s2)) \/ InA X.eq x nil <-> InT x s1 \/ InT x s2
s1, s2:tree
x:elt

InT x s1 \/ InT x s2 \/ False <-> InT x s1 \/ InT x s2
tauto. Qed.
s1, s2:tree
x:elt

InT x (fold add s1 s2) <-> InT x s1 \/ InT x s2
s1, s2:tree
x:elt

InT x (fold add s1 s2) <-> InT x s1 \/ InT x s2
s1, s2:tree
x:elt

InT x (fold_right (fun (y : elt) (x0 : tree) => flip add x0 y) s2 (rev (elements s1))) <-> InT x s1 \/ InT x s2
s1, s2:tree
x:elt

InT x (fold_right (fun (y : elt) (x0 : tree) => flip add x0 y) s2 (rev (elements s1))) <-> InA X.eq x (rev (elements s1)) \/ InT x s2
s1, s2:tree
x:X.t

InT x (fold_right (fun (y : X.t) (x0 : tree) => flip add x0 y) s2 (rev (elements s1))) <-> InA X.eq x (rev (elements s1)) \/ InT x s2
s1, s2:tree
x:X.t

InT x s2 <-> InA X.eq x nil \/ InT x s2
s1, s2:tree
x, a:X.t
l:list X.t
IHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip add x0 y) s2 l) <-> InA X.eq x l \/ InT x s2
InT x (flip add (fold_right (fun (y : X.t) (x0 : tree) => flip add x0 y) s2 l) a) <-> InA X.eq x (a :: l) \/ InT x s2
s1, s2:tree
x:X.t

InT x s2 <-> InA X.eq x nil \/ InT x s2
s1, s2:tree
x:X.t

InT x s2 <-> False \/ InT x s2
tauto.
s1, s2:tree
x, a:X.t
l:list X.t
IHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip add x0 y) s2 l) <-> InA X.eq x l \/ InT x s2

InT x (flip add (fold_right (fun (y : X.t) (x0 : tree) => flip add x0 y) s2 l) a) <-> InA X.eq x (a :: l) \/ InT x s2
s1, s2:tree
x, a:X.t
l:list X.t
IHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip add x0 y) s2 l) <-> InA X.eq x l \/ InT x s2

InT x (add a (fold_right (fun (y : X.t) (x0 : tree) => add y x0) s2 l)) <-> InA X.eq x (a :: l) \/ InT x s2
s1, s2:tree
x, a:X.t
l:list X.t
IHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip add x0 y) s2 l) <-> InA X.eq x l \/ InT x s2

X.eq x a \/ InA X.eq x l \/ InT x s2 <-> (X.eq x a \/ InA X.eq x l) \/ InT x s2
tauto. Qed.
s1, s2:t
x:elt

InT x (union s1 s2) <-> InT x s1 \/ InT x s2
s1, s2:t
x:elt

InT x (union s1 s2) <-> InT x s1 \/ InT x s2
s1, s2:t
x:elt

InT x match compare_height s1 s1 s2 s2 with | Eq => linear_union s1 s2 | Lt => fold add s1 s2 | Gt => fold add s2 s1 end <-> InT x s1 \/ InT x s2
s1, s2:t
x:elt

InT x (linear_union s1 s2) <-> InT x s1 \/ InT x s2
s1, s2:t
x:elt
InT x (fold add s1 s2) <-> InT x s1 \/ InT x s2
s1, s2:t
x:elt
InT x (fold add s2 s1) <-> InT x s1 \/ InT x s2
s1, s2:t
x:elt

InT x (linear_union s1 s2) <-> InT x s1 \/ InT x s2
apply linear_union_spec.
s1, s2:t
x:elt

InT x (fold add s1 s2) <-> InT x s1 \/ InT x s2
apply fold_add_spec.
s1, s2:t
x:elt

InT x (fold add s2 s1) <-> InT x s1 \/ InT x s2
s1, s2:t
x:elt

InT x s2 \/ InT x s1 <-> InT x s1 \/ InT x s2
tauto. Qed.

forall (s1 s2 : tree) (y : elt), Ok s1 -> Ok s2 -> InT y (union s1 s2) <-> InT y s1 \/ InT y s2

forall (s1 s2 : tree) (y : elt), Ok s1 -> Ok s2 -> InT y (union s1 s2) <-> InT y s1 \/ InT y s2
intros; apply union_spec'. Qed.

inter

l1, l2, acc:list X.t

INV l1 l2 acc -> Sorted X.lt (inter_list l1 l2 acc)
l1, l2, acc:list X.t

INV l1 l2 acc -> Sorted X.lt (inter_list l1 l2 acc)
l1:list X.t

forall l2 acc : list X.t, INV l1 l2 acc -> Sorted X.lt (inter_list l1 l2 acc)

forall l2 acc : list X.t, INV nil l2 acc -> Sorted X.lt acc
x1:X.t
l1:list X.t
IH1:forall l2 acc : list X.t, INV l1 l2 acc -> Sorted X.lt (inter_list l1 l2 acc)
forall acc : list X.t, INV (x1 :: l1) nil acc -> Sorted X.lt acc
x1:X.t
l1:list X.t
IH1:forall l0 acc : list X.t, INV l1 l0 acc -> Sorted X.lt (inter_list l1 l0 acc)
x2:X.t
l2:list X.t
IH2:forall acc : list X.t, INV (x1 :: l1) l2 acc -> Sorted X.lt (inter_list (x1 :: l1) l2 acc)
forall acc : list X.t, INV (x1 :: l1) (x2 :: l2) acc -> Sorted X.lt match X.compare x1 x2 with | Eq => inter_list l1 l2 (x1 :: acc) | Lt => (fix inter_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => acc0 | y :: l2' => match X.compare x1 y with | Eq => inter_list l1 l2' (x1 :: acc0) | Lt => inter_l1 l2' acc0 | Gt => inter_list l1 l0 acc0 end end) l2 acc | Gt => inter_list l1 (x2 :: l2) acc end

forall l2 acc : list X.t, INV nil l2 acc -> Sorted X.lt acc
eauto.
x1:X.t
l1:list X.t
IH1:forall l2 acc : list X.t, INV l1 l2 acc -> Sorted X.lt (inter_list l1 l2 acc)

forall acc : list X.t, INV (x1 :: l1) nil acc -> Sorted X.lt acc
eauto.
x1:X.t
l1:list X.t
IH1:forall l0 acc : list X.t, INV l1 l0 acc -> Sorted X.lt (inter_list l1 l0 acc)
x2:X.t
l2:list X.t
IH2:forall acc : list X.t, INV (x1 :: l1) l2 acc -> Sorted X.lt (inter_list (x1 :: l1) l2 acc)

forall acc : list X.t, INV (x1 :: l1) (x2 :: l2) acc -> Sorted X.lt match X.compare x1 x2 with | Eq => inter_list l1 l2 (x1 :: acc) | Lt => (fix inter_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => acc0 | y :: l2' => match X.compare x1 y with | Eq => inter_list l1 l2' (x1 :: acc0) | Lt => inter_l1 l2' acc0 | Gt => inter_list l1 l0 acc0 end end) l2 acc | Gt => inter_list l1 (x2 :: l2) acc end
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc

Sorted X.lt match X.compare x1 x2 with | Eq => inter_list l1 l2 (x1 :: acc) | Lt => (fix inter_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => acc0 | y :: l2' => match X.compare x1 y with | Eq => inter_list l1 l2' (x1 :: acc0) | Lt => inter_l1 l2' acc0 | Gt => inter_list l1 l0 acc0 end end) l2 acc | Gt => inter_list l1 (x2 :: l2) acc end
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.eq x1 x2

Sorted X.lt (inter_list l1 l2 (x1 :: acc))
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x1 x2
Sorted X.lt ((fix inter_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => acc0 | y :: l2' => match X.compare x1 y with | Eq => inter_list l1 l2' (x1 :: acc0) | Lt => inter_l1 l2' acc0 | Gt => inter_list l1 l0 acc0 end end) l2 acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x2 x1
Sorted X.lt (inter_list l1 (x2 :: l2) acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.eq x1 x2

Sorted X.lt (inter_list l1 l2 (x1 :: acc))
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.eq x1 x2

INV l1 l2 (x1 :: acc)
eapply INV_eq; eauto.
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x1 x2

Sorted X.lt ((fix inter_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => acc0 | y :: l2' => match X.compare x1 y with | Eq => inter_list l1 l2' (x1 :: acc0) | Lt => inter_l1 l2' acc0 | Gt => inter_list l1 l0 acc0 end end) l2 acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x1 x2

INV (x1 :: l1) l2 acc
eapply INV_sym, INV_drop, INV_sym; eauto.
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x2 x1

Sorted X.lt (inter_list l1 (x2 :: l2) acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x2 x1

INV l1 (x2 :: l2) acc
eapply INV_drop; eauto. Qed.
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (linear_inter s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (linear_inter s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (treeify (inter_list (rev_elements s1) (rev_elements s2) nil))
now apply treeify_ok, inter_list_ok, INV_init. Qed.
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (inter s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (inter s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok match compare_height s1 s1 s2 s2 with | Eq => linear_inter s1 s2 | Lt => filter (fun k : elt => mem k s2) s1 | Gt => filter (fun k : elt => mem k s1) s2 end
destruct compare_height; auto_tc. Qed.
x:X.t
l1, l2:list X.t
acc:list elt

Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list l1 l2 acc) <-> InA X.eq x l1 /\ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l1, l2:list X.t
acc:list elt

Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list l1 l2 acc) <-> InA X.eq x l1 /\ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l1:list X.t

forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list l1 l2 acc) <-> InA X.eq x l1 /\ InA X.eq x l2 \/ InA X.eq x acc
x:X.t

forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev nil) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list nil l2 acc) <-> InA X.eq x nil /\ InA X.eq x l2 \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list l1 l2 acc) <-> InA X.eq x l1 /\ InA X.eq x l2 \/ InA X.eq x acc
forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc
x:X.t

forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev nil) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list nil l2 acc) <-> InA X.eq x nil /\ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l2:list X.t
acc:list elt

Sorted X.lt nil -> Sorted X.lt (rev l2) -> InA X.eq x acc <-> InA X.eq x nil /\ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l2:list X.t
acc:list elt

Sorted X.lt nil -> Sorted X.lt (rev l2) -> InA X.eq x acc <-> False /\ InA X.eq x l2 \/ InA X.eq x acc
tauto.
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list l1 l2 acc) <-> InA X.eq x l1 /\ InA X.eq x l2 \/ InA X.eq x acc

forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list l1 l2 acc0) <-> InA X.eq x l1 /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev nil) -> InA X.eq x (inter_list (x1 :: l1) nil acc) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x nil \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev (x2 :: l2)) -> InA X.eq x (inter_list (x1 :: l1) (x2 :: l2) acc) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list l1 l2 acc0) <-> InA X.eq x l1 /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev nil) -> InA X.eq x (inter_list (x1 :: l1) nil acc) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x nil \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list l1 l2 acc0) <-> InA X.eq x l1 /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

Sorted X.lt (rev l1 ++ x1 :: nil) -> Sorted X.lt nil -> InA X.eq x acc <-> InA X.eq x (x1 :: l1) /\ InA X.eq x nil \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list l1 l2 acc0) <-> InA X.eq x l1 /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

Sorted X.lt (rev l1 ++ x1 :: nil) -> Sorted X.lt nil -> InA X.eq x acc <-> (X.eq x x1 \/ InA X.eq x l1) /\ False \/ InA X.eq x acc
tauto.
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev (x2 :: l2)) -> InA X.eq x (inter_list (x1 :: l1) (x2 :: l2) acc) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

Sorted X.lt (rev l1 ++ x1 :: nil) -> Sorted X.lt (rev l2 ++ x2 :: nil) -> InA X.eq x match X.compare x1 x2 with | Eq => inter_list l1 l2 (x1 :: acc) | Lt => (fix inter_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => acc0 | y :: l2' => match X.compare x1 y with | Eq => inter_list l1 l2' (x1 :: acc0) | Lt => inter_l1 l2' acc0 | Gt => inter_list l1 l0 acc0 end end) l2 acc | Gt => inter_list l1 (x2 :: l2) acc end <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)

InA X.eq x match X.compare x1 x2 with | Eq => inter_list l1 l2 (x1 :: acc) | Lt => (fix inter_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => acc0 | y :: l2' => match X.compare x1 y with | Eq => inter_list l1 l2' (x1 :: acc0) | Lt => inter_l1 l2' acc0 | Gt => inter_list l1 l0 acc0 end end) l2 acc | Gt => inter_list l1 (x2 :: l2) acc end <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3

InA X.eq x match X.compare x1 x2 with | Eq => inter_list l1 l2 (x1 :: acc) | Lt => (fix inter_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => acc0 | y :: l2' => match X.compare x1 y with | Eq => inter_list l1 l2' (x1 :: acc0) | Lt => inter_l1 l2' acc0 | Gt => inter_list l1 l0 acc0 end end) l2 acc | Gt => inter_list l1 (x2 :: l2) acc end <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

InA X.eq x match X.compare x1 x2 with | Eq => inter_list l1 l2 (x1 :: acc) | Lt => (fix inter_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => acc0 | y :: l2' => match X.compare x1 y with | Eq => inter_list l1 l2' (x1 :: acc0) | Lt => inter_l1 l2' acc0 | Gt => inter_list l1 l0 acc0 end end) l2 acc | Gt => inter_list l1 (x2 :: l2) acc end <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.eq x1 x2

InA X.eq x (inter_list l1 l2 (x1 :: acc)) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2
InA X.eq x ((fix inter_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => acc0 | y :: l2' => match X.compare x1 y with | Eq => inter_list l1 l2' (x1 :: acc0) | Lt => inter_l1 l2' acc0 | Gt => inter_list l1 l0 acc0 end end) l2 acc) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1
InA X.eq x (inter_list l1 (x2 :: l2) acc) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.eq x1 x2

InA X.eq x (inter_list l1 l2 (x1 :: acc)) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
rewrite IH1, !InA_cons, C; tauto.
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2

InA X.eq x ((fix inter_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => acc0 | y :: l2' => match X.compare x1 y with | Eq => inter_list l1 l2' (x1 :: acc0) | Lt => inter_l1 l2' acc0 | Gt => inter_list l1 l0 acc0 end end) l2 acc) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2

InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2

(X.eq x x1 \/ InA X.eq x l1) /\ InA X.eq x l2 \/ InA X.eq x acc <-> (X.eq x x1 \/ InA X.eq x l1) /\ (X.eq x x2 \/ InA X.eq x l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2
H0:InA X.eq x l1
H:X.eq x x2

(X.eq x x1 \/ InA X.eq x l1) /\ InA X.eq x l2 \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2
H0:InA X.eq x l1
H:X.eq x x2
H1:X.lt x x1

(X.eq x x1 \/ InA X.eq x l1) /\ InA X.eq x l2 \/ InA X.eq x acc
order.
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1

InA X.eq x (inter_list l1 (x2 :: l2) acc) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1

InA X.eq x l1 /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc <-> InA X.eq x (x1 :: l1) /\ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1

InA X.eq x l1 /\ (X.eq x x2 \/ InA X.eq x l2) \/ InA X.eq x acc <-> (X.eq x x1 \/ InA X.eq x l1) /\ (X.eq x x2 \/ InA X.eq x l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1
H0:X.eq x x1
H:InA X.eq x l2

InA X.eq x l1 /\ (X.eq x x2 \/ InA X.eq x l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (inter_list l1 l0 acc0) <-> InA X.eq x l1 /\ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (inter_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1
H0:X.eq x x1
H:InA X.eq x l2
H1:X.lt x x2

InA X.eq x l1 /\ (X.eq x x2 \/ InA X.eq x l2) \/ InA X.eq x acc
order. Qed.
s1, s2:tree
x:elt
H:Ok s1
H0:Ok s2

InT x (linear_inter s1 s2) <-> InT x s1 /\ InT x s2
s1, s2:tree
x:elt
H:Ok s1
H0:Ok s2

InT x (linear_inter s1 s2) <-> InT x s1 /\ InT x s2
s1, s2:tree
x:elt
H:Ok s1
H0:Ok s2

InT x (treeify (inter_list (rev_elements s1) (rev_elements s2) nil)) <-> InT x s1 /\ InT x s2
s1, s2:tree
x:elt
H:Ok s1
H0:Ok s2

InA X.eq x (rev (elements s1)) /\ InA X.eq x (rev (elements s2)) \/ InA X.eq x nil <-> InT x s1 /\ InT x s2
s1, s2:tree
x:elt
H:Ok s1
H0:Ok s2

InT x s1 /\ InT x s2 \/ False <-> InT x s1 /\ InT x s2
tauto. Qed.
s:tree
H:Ok s

Proper (X.eq ==> Logic.eq) (fun k : X.t => mem k s)
s:tree
H:Ok s

Proper (X.eq ==> Logic.eq) (fun k : X.t => mem k s)
s:tree
H:Ok s
x, y:X.t
EQ:X.eq x y

mem x s = mem y s
s:tree
H:Ok s
x, y:X.t
EQ:X.eq x y

InT x s <-> InT y s
now rewrite EQ. Qed.
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (inter s1 s2) <-> InT y s1 /\ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (inter s1 s2) <-> InT y s1 /\ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y match compare_height s1 s1 s2 s2 with | Eq => linear_inter s1 s2 | Lt => filter (fun k : elt => mem k s2) s1 | Gt => filter (fun k : elt => mem k s1) s2 end <-> InT y s1 /\ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (linear_inter s1 s2) <-> InT y s1 /\ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2
InT y (filter (fun k : elt => mem k s2) s1) <-> InT y s1 /\ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2
InT y (filter (fun k : elt => mem k s1) s2) <-> InT y s1 /\ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (linear_inter s1 s2) <-> InT y s1 /\ InT y s2
now apply linear_inter_spec.
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (filter (fun k : elt => mem k s2) s1) <-> InT y s1 /\ InT y s2
rewrite filter_spec, mem_spec by auto_tc; tauto.
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (filter (fun k : elt => mem k s1) s2) <-> InT y s1 /\ InT y s2
rewrite filter_spec, mem_spec by auto_tc; tauto. Qed.

difference

l1, l2, acc:list X.t

INV l1 l2 acc -> Sorted X.lt (diff_list l1 l2 acc)
l1, l2, acc:list X.t

INV l1 l2 acc -> Sorted X.lt (diff_list l1 l2 acc)
l1:list X.t

forall l2 acc : list X.t, INV l1 l2 acc -> Sorted X.lt (diff_list l1 l2 acc)
l2, acc:list X.t
inv:INV nil l2 acc

Sorted X.lt (diff_list nil l2 acc)
x1:X.t
l1:list X.t
IH1:forall l2 acc0 : list X.t, INV l1 l2 acc0 -> Sorted X.lt (diff_list l1 l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) nil acc
Sorted X.lt (diff_list (x1 :: l1) nil acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
Sorted X.lt (diff_list (x1 :: l1) (x2 :: l2) acc)
l2, acc:list X.t
inv:INV nil l2 acc

Sorted X.lt (diff_list nil l2 acc)
eauto.
x1:X.t
l1:list X.t
IH1:forall l2 acc0 : list X.t, INV l1 l2 acc0 -> Sorted X.lt (diff_list l1 l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) nil acc

Sorted X.lt (diff_list (x1 :: l1) nil acc)
x1:X.t
l1:list X.t
IH1:forall l2 acc0 : list X.t, INV l1 l2 acc0 -> Sorted X.lt (diff_list l1 l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) nil acc

Sorted X.lt (rev_append (x1 :: l1) acc)
eapply INV_rev; eauto.
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc

Sorted X.lt (diff_list (x1 :: l1) (x2 :: l2) acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc

Sorted X.lt match X.compare x1 x2 with | Eq => diff_list l1 l2 acc | Lt => (fix diff_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => diff_list l1 l2' acc0 | Lt => diff_l1 l2' acc0 | Gt => diff_list l1 l0 (x1 :: acc0) end end) l2 acc | Gt => diff_list l1 (x2 :: l2) (x1 :: acc) end
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.eq x1 x2

Sorted X.lt (diff_list l1 l2 acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x1 x2
Sorted X.lt ((fix diff_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => diff_list l1 l2' acc0 | Lt => diff_l1 l2' acc0 | Gt => diff_list l1 l0 (x1 :: acc0) end end) l2 acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x2 x1
Sorted X.lt (diff_list l1 (x2 :: l2) (x1 :: acc))
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.eq x1 x2

Sorted X.lt (diff_list l1 l2 acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.eq x1 x2

INV l1 l2 acc
eapply INV_drop, INV_sym, INV_drop, INV_sym; eauto.
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x1 x2

Sorted X.lt ((fix diff_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => diff_list l1 l2' acc0 | Lt => diff_l1 l2' acc0 | Gt => diff_list l1 l0 (x1 :: acc0) end end) l2 acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x1 x2

INV (x1 :: l1) l2 acc
eapply INV_sym, INV_drop, INV_sym; eauto.
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x2 x1

Sorted X.lt (diff_list l1 (x2 :: l2) (x1 :: acc))
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x2 x1

INV l1 (x2 :: l2) (x1 :: acc)
x1:X.t
l1:list X.t
IH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)
x2:X.t
l2:list X.t
IH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)
acc:list X.t
inv:INV (x1 :: l1) (x2 :: l2) acc
C:X.lt x2 x1

INV (x2 :: l2) (x1 :: l1) acc
now apply INV_sym. Qed.
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (linear_diff s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (linear_diff s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (linear_diff s1 s2)
now apply treeify_ok, diff_list_ok, INV_init. Qed.
s1, s2:tree
H:Ok s2

Ok (fold remove s1 s2)
s1, s2:tree
H:Ok s2

Ok (fold remove s1 s2)
s1, s2:tree
H:Ok s2

Ok (fold_right (fun (y : elt) (x : tree) => flip remove x y) s2 (rev (elements s1)))
s1, s2:tree
H:Ok s2

Ok (fold_right (fun (y : X.t) (x : tree) => flip remove x y) s2 (rev (elements s1)))
induction (rev (elements s1)); simpl; unfold flip in *; auto_tc. Qed.
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (diff s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok (diff s1 s2)
s1, s2:tree
H:Ok s1
H0:Ok s2

Ok match compare_height s1 s1 s2 s2 with | Eq => linear_diff s1 s2 | Lt => filter (fun k : elt => negb (mem k s2)) s1 | Gt => fold remove s2 s1 end
destruct compare_height; auto_tc. Qed.
x:X.t
l1, l2:list X.t
acc:list elt

Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list l1 l2 acc) <-> InA X.eq x l1 /\ ~ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l1, l2:list X.t
acc:list elt

Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list l1 l2 acc) <-> InA X.eq x l1 /\ ~ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l1:list X.t

forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list l1 l2 acc) <-> InA X.eq x l1 /\ ~ InA X.eq x l2 \/ InA X.eq x acc
x:X.t

forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev nil) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list nil l2 acc) <-> InA X.eq x nil /\ ~ InA X.eq x l2 \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list l1 l2 acc) <-> InA X.eq x l1 /\ ~ InA X.eq x l2 \/ InA X.eq x acc
forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc
x:X.t

forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev nil) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list nil l2 acc) <-> InA X.eq x nil /\ ~ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l2:list X.t
acc:list elt

Sorted X.lt nil -> Sorted X.lt (rev l2) -> InA X.eq x acc <-> InA X.eq x nil /\ ~ InA X.eq x l2 \/ InA X.eq x acc
x:X.t
l2:list X.t
acc:list elt

Sorted X.lt nil -> Sorted X.lt (rev l2) -> InA X.eq x acc <-> False /\ ~ InA X.eq x l2 \/ InA X.eq x acc
tauto.
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list l1 l2 acc) <-> InA X.eq x l1 /\ ~ InA X.eq x l2 \/ InA X.eq x acc

forall (l2 : list X.t) (acc : list elt), Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list l1 l2 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev nil) -> InA X.eq x (diff_list (x1 :: l1) nil acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x nil \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev (x2 :: l2)) -> InA X.eq x (diff_list (x1 :: l1) (x2 :: l2) acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list l1 l2 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev nil) -> InA X.eq x (diff_list (x1 :: l1) nil acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x nil \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list l1 l2 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
H:Sorted X.lt (rev (x1 :: l1))
H0:Sorted X.lt (rev nil)

InA X.eq x (rev_append l1 (x1 :: acc)) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x nil \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list l1 l2 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
H:Sorted X.lt (rev (x1 :: l1))
H0:Sorted X.lt (rev nil)

InA X.eq x (rev l1 ++ x1 :: acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x nil \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l2 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list l1 l2 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
H:Sorted X.lt (rev (x1 :: l1))
H0:Sorted X.lt (rev nil)

InA X.eq x l1 \/ X.eq x x1 \/ InA X.eq x acc <-> (X.eq x x1 \/ InA X.eq x l1) /\ ~ False \/ InA X.eq x acc
tauto.
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev (x2 :: l2)) -> InA X.eq x (diff_list (x1 :: l1) (x2 :: l2) acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt

Sorted X.lt (rev l1 ++ x1 :: nil) -> Sorted X.lt (rev l2 ++ x2 :: nil) -> InA X.eq x match X.compare x1 x2 with | Eq => diff_list l1 l2 acc | Lt => (fix diff_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => diff_list l1 l2' acc0 | Lt => diff_l1 l2' acc0 | Gt => diff_list l1 l0 (x1 :: acc0) end end) l2 acc | Gt => diff_list l1 (x2 :: l2) (x1 :: acc) end <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)

InA X.eq x match X.compare x1 x2 with | Eq => diff_list l1 l2 acc | Lt => (fix diff_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => diff_list l1 l2' acc0 | Lt => diff_l1 l2' acc0 | Gt => diff_list l1 l0 (x1 :: acc0) end end) l2 acc | Gt => diff_list l1 (x2 :: l2) (x1 :: acc) end <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3

InA X.eq x match X.compare x1 x2 with | Eq => diff_list l1 l2 acc | Lt => (fix diff_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => diff_list l1 l2' acc0 | Lt => diff_l1 l2' acc0 | Gt => diff_list l1 l0 (x1 :: acc0) end end) l2 acc | Gt => diff_list l1 (x2 :: l2) (x1 :: acc) end <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3

InA X.eq x match X.compare x1 x2 with | Eq => diff_list l1 l2 acc | Lt => (fix diff_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => diff_list l1 l2' acc0 | Lt => diff_l1 l2' acc0 | Gt => diff_list l1 l0 (x1 :: acc0) end end) l2 acc | Gt => diff_list l1 (x2 :: l2) (x1 :: acc) end <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.eq x1 x2

InA X.eq x (diff_list l1 l2 acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2
InA X.eq x ((fix diff_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => diff_list l1 l2' acc0 | Lt => diff_l1 l2' acc0 | Gt => diff_list l1 l0 (x1 :: acc0) end end) l2 acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1
InA X.eq x (diff_list l1 (x2 :: l2) (x1 :: acc)) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.eq x1 x2

InA X.eq x (diff_list l1 l2 acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.eq x1 x2

InA X.eq x l1 /\ ~ InA X.eq x l2 \/ InA X.eq x acc <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.eq x1 x2

InA X.eq x l1 /\ ~ InA X.eq x l2 <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2)
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.eq x1 x2

InA X.eq x l1 /\ ~ InA X.eq x l2 <-> (X.eq x x1 \/ InA X.eq x l1) /\ ~ (X.eq x x2 \/ InA X.eq x l2)
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ (InA X.eq x l0 -> False) \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ (InA X.eq x l2 -> False) \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.eq x1 x2
H0:InA X.eq x l1
H1:InA X.eq x l2 -> False
H2:X.eq x x2

False
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ (InA X.eq x l0 -> False) \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ (InA X.eq x l2 -> False) \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.eq x1 x2
H0:InA X.eq x l1
H1:InA X.eq x l2 -> False
H2:X.eq x x2
H:X.lt x x1

False
order.
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2

InA X.eq x ((fix diff_l1 (l0 acc0 : list elt) {struct l0} : list elt := match l0 with | nil => rev_append l1 (x1 :: acc0) | y :: l2' => match X.compare x1 y with | Eq => diff_list l1 l2' acc0 | Lt => diff_l1 l2' acc0 | Gt => diff_list l1 l0 (x1 :: acc0) end end) l2 acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2

InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2

InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2)
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2

(X.eq x x1 \/ InA X.eq x l1) /\ ~ InA X.eq x l2 <-> (X.eq x x1 \/ InA X.eq x l1) /\ ~ (X.eq x x2 \/ InA X.eq x l2)
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ (InA X.eq x l0 -> False) \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ (InA X.eq x l2 -> False) \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2
H1:InA X.eq x l2 -> False
H:InA X.eq x l1
H2:X.eq x x2

False
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ (InA X.eq x l0 -> False) \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ (InA X.eq x l2 -> False) \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x1 x2
H1:InA X.eq x l2 -> False
H:InA X.eq x l1
H2:X.eq x x2
H0:X.lt x x1

False
order.
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1

InA X.eq x (diff_list l1 (x2 :: l2) (x1 :: acc)) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1

InA X.eq x l1 /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x (x1 :: acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ ~ InA X.eq x l0 \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x l2 \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1

InA X.eq x l1 /\ ~ (X.eq x x2 \/ InA X.eq x l2) \/ X.eq x x1 \/ InA X.eq x acc <-> (X.eq x x1 \/ InA X.eq x l1) /\ ~ (X.eq x x2 \/ InA X.eq x l2) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ (InA X.eq x l0 -> False) \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ (InA X.eq x l2 -> False) \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1
H:X.eq x x1

(X.eq x x1 \/ InA X.eq x l1) /\ (X.eq x x2 \/ InA X.eq x l2 -> False) \/ InA X.eq x acc
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ (InA X.eq x l0 -> False) \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ (InA X.eq x l2 -> False) \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1
H:X.eq x x1

X.eq x x2 \/ InA X.eq x l2 -> False
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ (InA X.eq x l0 -> False) \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ (InA X.eq x l2 -> False) \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1
H:X.eq x x1
H0:X.eq x x2

False
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ (InA X.eq x l0 -> False) \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ (InA X.eq x l2 -> False) \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1
H:X.eq x x1
H0:InA X.eq x l2
False
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ (InA X.eq x l0 -> False) \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ (InA X.eq x l2 -> False) \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1
H:X.eq x x1
H0:InA X.eq x l2

False
x, x1:X.t
l1:list X.t
IH1:forall (l0 : list X.t) (acc0 : list elt), Sorted X.lt (rev l1) -> Sorted X.lt (rev l0) -> InA X.eq x (diff_list l1 l0 acc0) <-> InA X.eq x l1 /\ (InA X.eq x l0 -> False) \/ InA X.eq x acc0
x2:X.t
l2:list X.t
IH2:forall acc0 : list elt, Sorted X.lt (rev (x1 :: l1)) -> Sorted X.lt (rev l2) -> InA X.eq x (diff_list (x1 :: l1) l2 acc0) <-> InA X.eq x (x1 :: l1) /\ (InA X.eq x l2 -> False) \/ InA X.eq x acc0
acc:list elt
U:Sorted X.lt (rev l1 ++ x1 :: nil)
V:Sorted X.lt (rev l2 ++ x2 :: nil)
U1:Sorted X.lt (rev l1)
U2:Sorted X.lt (x1 :: nil)
U3:forall x0 x3 : X.t, InA X.eq x0 (rev l1) -> InA X.eq x3 (x1 :: nil) -> X.lt x0 x3
V1:Sorted X.lt (rev l2)
V2:Sorted X.lt (x2 :: nil)
V3:forall x0 x3 : X.t, InA X.eq x0 (rev l2) -> InA X.eq x3 (x2 :: nil) -> X.lt x0 x3
C:X.lt x2 x1
H:X.eq x x1
H0:InA X.eq x l2
H1:X.lt x x2

False
order. Qed.
s1, s2:tree
x:elt
H:Ok s1
H0:Ok s2

InT x (linear_diff s1 s2) <-> InT x s1 /\ ~ InT x s2
s1, s2:tree
x:elt
H:Ok s1
H0:Ok s2

InT x (linear_diff s1 s2) <-> InT x s1 /\ ~ InT x s2
s1, s2:tree
x:elt
H:Ok s1
H0:Ok s2

InT x (treeify (diff_list (rev_elements s1) (rev_elements s2) nil)) <-> InT x s1 /\ ~ InT x s2
s1, s2:tree
x:elt
H:Ok s1
H0:Ok s2

InA X.eq x (rev (elements s1)) /\ ~ InA X.eq x (rev (elements s2)) \/ InA X.eq x nil <-> InT x s1 /\ ~ InT x s2
s1, s2:tree
x:elt
H:Ok s1
H0:Ok s2

InT x s1 /\ ~ InT x s2 \/ False <-> InT x s1 /\ ~ InT x s2
tauto. Qed.
s1, s2:tree
x:elt
H:Ok s2

InT x (fold remove s1 s2) <-> InT x s2 /\ ~ InT x s1
s1, s2:tree
x:elt
H:Ok s2

InT x (fold remove s1 s2) <-> InT x s2 /\ ~ InT x s1
s1, s2:tree
x:elt
H:Ok s2

InT x (fold_right (fun (y : elt) (x0 : tree) => flip remove x0 y) s2 (rev (elements s1))) <-> InT x s2 /\ ~ InT x s1
s1, s2:tree
x:elt
H:Ok s2

InT x (fold_right (fun (y : elt) (x0 : tree) => flip remove x0 y) s2 (rev (elements s1))) <-> InT x s2 /\ ~ InA X.eq x (rev (elements s1))
s1, s2:tree
x:X.t
H:Ok s2

InT x (fold_right (fun (y : X.t) (x0 : tree) => flip remove x0 y) s2 (rev (elements s1))) <-> InT x s2 /\ ~ InA X.eq x (rev (elements s1))
s1, s2:tree
x:X.t
H:Ok s2

InT x s2 <-> InT x s2 /\ ~ InA X.eq x nil
s1, s2:tree
x:X.t
H:Ok s2
a:X.t
l:list X.t
IHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip remove x0 y) s2 l) <-> InT x s2 /\ ~ InA X.eq x l
InT x (flip remove (fold_right (fun (y : X.t) (x0 : tree) => flip remove x0 y) s2 l) a) <-> InT x s2 /\ ~ InA X.eq x (a :: l)
s1, s2:tree
x:X.t
H:Ok s2

InT x s2 <-> InT x s2 /\ ~ InA X.eq x nil
s1, s2:tree
x:X.t
H:Ok s2

InT x s2 <-> InT x s2 /\ ~ False
intuition.
s1, s2:tree
x:X.t
H:Ok s2
a:X.t
l:list X.t
IHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip remove x0 y) s2 l) <-> InT x s2 /\ ~ InA X.eq x l

InT x (flip remove (fold_right (fun (y : X.t) (x0 : tree) => flip remove x0 y) s2 l) a) <-> InT x s2 /\ ~ InA X.eq x (a :: l)
s1, s2:tree
x:X.t
H:Ok s2
a:X.t
l:list X.t
IHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l) <-> InT x s2 /\ ~ InA X.eq x l

InT x (remove a (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l)) <-> InT x s2 /\ ~ InA X.eq x (a :: l)
s1, s2:tree
x:X.t
H:Ok s2
a:X.t
l:list X.t
IHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l) <-> InT x s2 /\ ~ InA X.eq x l

(InT x s2 /\ ~ InA X.eq x l) /\ ~ X.eq x a <-> InT x s2 /\ ~ (X.eq x a \/ InA X.eq x l)
s1, s2:tree
x:X.t
H:Ok s2
a:X.t
l:list X.t
IHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l) <-> InT x s2 /\ ~ InA X.eq x l
Ok (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l)
s1, s2:tree
x:X.t
H:Ok s2
a:X.t
l:list X.t
IHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l) <-> InT x s2 /\ ~ InA X.eq x l

Ok (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l)
s1, s2:tree
x:X.t
H:Ok s2
a:X.t
l:list X.t

Ok (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l)
induction l; simpl; auto_tc. Qed.
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (diff s1 s2) <-> InT y s1 /\ ~ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (diff s1 s2) <-> InT y s1 /\ ~ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y match compare_height s1 s1 s2 s2 with | Eq => linear_diff s1 s2 | Lt => filter (fun k : elt => negb (mem k s2)) s1 | Gt => fold remove s2 s1 end <-> InT y s1 /\ ~ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (linear_diff s1 s2) <-> InT y s1 /\ ~ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2
InT y (filter (fun k : elt => negb (mem k s2)) s1) <-> InT y s1 /\ ~ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2
InT y (fold remove s2 s1) <-> InT y s1 /\ ~ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (linear_diff s1 s2) <-> InT y s1 /\ ~ InT y s2
now apply linear_diff_spec.
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (filter (fun k : elt => negb (mem k s2)) s1) <-> InT y s1 /\ ~ InT y s2
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

Proper (X.eq ==> Logic.eq) (fun k : elt => negb (mem k s2))
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2
x1, x2:X.t
EQ:X.eq x1 x2

negb (mem x1 s2) = negb (mem x2 s2)
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2
x1, x2:X.t
EQ:X.eq x1 x2

mem x1 s2 = mem x2 s2
now apply mem_proper.
s1, s2:tree
y:elt
H:Ok s1
H0:Ok s2

InT y (fold remove s2 s1) <-> InT y s1 /\ ~ InT y s2
now apply fold_remove_spec. Qed. End MakeRaw.

Balancing properties

We now prove that all operations preserve a red-black invariant, and that trees have hence a logarithmic depth.
Module BalanceProps(X:Orders.OrderedType)(Import M : MakeRaw X).

Notation Rd := (Node Red).
Notation Bk := (Node Black).
Import M.MX.

Red-Black invariants

In a red-black tree : The black depth is here an argument of the predicate.
Inductive rbt : nat -> tree -> Prop :=
 | RB_Leaf : rbt 0 Leaf
 | RB_Rd n l k r :
   notred l -> notred r -> rbt n l -> rbt n r -> rbt n (Rd l k r)
 | RB_Bk n l k r : rbt n l -> rbt n r -> rbt (S n) (Bk l k r).
A red-red tree is almost a red-black tree, except that it has a red root node which may have red children. Note that a red-red tree is hence non-empty, and all its strict subtrees are red-black.
Inductive rrt (n:nat) : tree -> Prop :=
 | RR_Rd l k r : rbt n l -> rbt n r -> rrt n (Rd l k r).
An almost-red-black tree is almost a red-black tree, except that it's permitted to have two red nodes in a row at the very root (only). We implement this notion by saying that a quasi-red-black tree is either a red-black tree or a red-red tree.
Inductive arbt (n:nat)(t:tree) : Prop :=
 | ARB_RB : rbt n t -> arbt n t
 | ARB_RR : rrt n t -> arbt n t.
The main exported invariant : being a red-black tree for some black depth.
Class Rbt (t:tree) :=  RBT : exists d, rbt d t.

Basic tactics and results about red-black

Scheme rbt_ind := Induction for rbt Sort Prop.
Local Hint Constructors rbt rrt arbt : core.
Local Hint Extern 0 (notred _) => (exact I) : core.
Ltac invrb := intros; invtree rrt; invtree rbt; try contradiction.
Ltac desarb := match goal with H:arbt _ _ |- _ => destruct H end.
Ltac nonzero n := destruct n as [|n]; [try split; invrb|].

n:nat
t:tree

rrt n t -> notredred t -> rbt n t
n:nat
t:tree

rrt n t -> notredred t -> rbt n t
n:nat
l:tree
x:X.t
r:tree
Hl:rbt n l
Hr:rbt n r

notredred (Rd l x r) -> rbt n (Rd l x r)
destruct l, r; descolor; invrb; auto. Qed. Local Hint Resolve rr_nrr_rb : core.
n:nat
t:tree

arbt n t -> notredred t -> rbt n t
n:nat
t:tree

arbt n t -> notredred t -> rbt n t
destruct 1; auto. Qed.
n:nat
t:tree

arbt n t -> notred t -> rbt n t
n:nat
t:tree

arbt n t -> notred t -> rbt n t
destruct 1; destruct t; descolor; invrb; auto. Qed. Local Hint Resolve arb_nrr_rb arb_nr_rb : core.

A Red-Black tree has indeed a logarithmic depth

Definition redcarac s := rcase (fun _ _ _ => 1) (fun _ => 0) s.

s:tree
n:nat

rbt n s -> maxdepth s <= 2 * n + redcarac s
s:tree
n:nat

rbt n s -> maxdepth s <= 2 * n + redcarac s

maxdepth Leaf <= 2 * 0 + redcarac Leaf
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:maxdepth l <= 2 * n + redcarac l
IHrbt2:maxdepth r <= 2 * n + redcarac r
maxdepth (Rd l k r) <= 2 * n + redcarac (Rd l k r)
n:nat
l:tree
k:X.t
r:tree
H:rbt n l
H0:rbt n r
IHrbt1:maxdepth l <= 2 * n + redcarac l
IHrbt2:maxdepth r <= 2 * n + redcarac r
maxdepth (Bk l k r) <= 2 * S n + redcarac (Bk l k r)

maxdepth Leaf <= 2 * 0 + redcarac Leaf
simpl; auto.
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:maxdepth l <= 2 * n + redcarac l
IHrbt2:maxdepth r <= 2 * n + redcarac r

maxdepth (Rd l k r) <= 2 * n + redcarac (Rd l k r)
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:maxdepth l <= 2 * n + 0
IHrbt2:maxdepth r <= 2 * n + redcarac r

maxdepth (Rd l k r) <= 2 * n + redcarac (Rd l k r)
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:maxdepth l <= 2 * n + 0
IHrbt2:maxdepth r <= 2 * n + 0

maxdepth (Rd l k r) <= 2 * n + redcarac (Rd l k r)
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:maxdepth l <= 2 * n + 0
IHrbt2:maxdepth r <= 2 * n + 0

S (Init.Nat.max (maxdepth l) (maxdepth r)) <= 2 * n + redcarac (Rd l k r)
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:maxdepth l <= 2 * n + 0
IHrbt2:maxdepth r <= 2 * n + 0

S (Init.Nat.max (maxdepth l) (maxdepth r)) <= 2 * n + 1
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:maxdepth l <= 2 * n + 0
IHrbt2:maxdepth r <= 2 * n + 0

Init.Nat.max (maxdepth l) (maxdepth r) <= 2 * n + 0
now apply Nat.max_lub.
n:nat
l:tree
k:X.t
r:tree
H:rbt n l
H0:rbt n r
IHrbt1:maxdepth l <= 2 * n + redcarac l
IHrbt2:maxdepth r <= 2 * n + redcarac r

maxdepth (Bk l k r) <= 2 * S n + redcarac (Bk l k r)
n:nat
l:tree
k:X.t
r:tree
H:rbt n l
H0:rbt n r
IHrbt1:maxdepth l <= 2 * n + redcarac l
IHrbt2:maxdepth r <= 2 * n + redcarac r

S (Init.Nat.max (maxdepth l) (maxdepth r)) <= S (n + S (n + 0) + 0)
n:nat
l:tree
k:X.t
r:tree
H:rbt n l
H0:rbt n r
IHrbt1:maxdepth l <= 2 * n + redcarac l
IHrbt2:maxdepth r <= 2 * n + redcarac r

Init.Nat.max (maxdepth l) (maxdepth r) <= n + S (n + 0) + 0
apply Nat.max_lub; eapply Nat.le_trans; eauto; [destree l | destree r]; simpl; rewrite !Nat.add_0_r, ?Nat.add_1_r; auto with arith. Qed.
s:tree
n:nat

rbt n s -> n + redcarac s <= mindepth s
s:tree
n:nat

rbt n s -> n + redcarac s <= mindepth s

0 <= 0
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:n + redcarac l <= mindepth l
IHrbt2:n + redcarac r <= mindepth r
n + 1 <= S (Init.Nat.min (mindepth l) (mindepth r))
n:nat
l:tree
k:X.t
r:tree
H:rbt n l
H0:rbt n r
IHrbt1:n + redcarac l <= mindepth l
IHrbt2:n + redcarac r <= mindepth r
S (n + 0) <= S (Init.Nat.min (mindepth l) (mindepth r))

0 <= 0
trivial.
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:n + redcarac l <= mindepth l
IHrbt2:n + redcarac r <= mindepth r

n + 1 <= S (Init.Nat.min (mindepth l) (mindepth r))
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:n + redcarac l <= mindepth l
IHrbt2:n + redcarac r <= mindepth r

S (n + 0) <= S (Init.Nat.min (mindepth l) (mindepth r))
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:n + redcarac l <= mindepth l
IHrbt2:n + redcarac r <= mindepth r

n + 0 <= Init.Nat.min (mindepth l) (mindepth r)
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:n + 0 <= mindepth l
IHrbt2:n + redcarac r <= mindepth r

n + 0 <= Init.Nat.min (mindepth l) (mindepth r)
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:n + 0 <= mindepth l
IHrbt2:n + 0 <= mindepth r

n + 0 <= Init.Nat.min (mindepth l) (mindepth r)
now apply Nat.min_glb.
n:nat
l:tree
k:X.t
r:tree
H:rbt n l
H0:rbt n r
IHrbt1:n + redcarac l <= mindepth l
IHrbt2:n + redcarac r <= mindepth r

S (n + 0) <= S (Init.Nat.min (mindepth l) (mindepth r))
n:nat
l:tree
k:X.t
r:tree
H:rbt n l
H0:rbt n r
IHrbt1:n + redcarac l <= mindepth l
IHrbt2:n + redcarac r <= mindepth r

n + 0 <= Init.Nat.min (mindepth l) (mindepth r)
n:nat
l:tree
k:X.t
r:tree
H:rbt n l
H0:rbt n r
IHrbt1:n + redcarac l <= mindepth l
IHrbt2:n + redcarac r <= mindepth r

n <= Init.Nat.min (mindepth l) (mindepth r)
apply Nat.min_glb; eauto with arith. Qed.
s:tree

Rbt s -> maxdepth s <= 2 * Nat.log2 (S (cardinal s))
s:tree

Rbt s -> maxdepth s <= 2 * Nat.log2 (S (cardinal s))
s:tree
n:nat
H:rbt n s

maxdepth s <= 2 * Nat.log2 (S (cardinal s))
s:tree
n:nat
H:rbt n s

2 * n + redcarac s <= 2 * Nat.log2 (S (cardinal s))
s:tree
n:nat
H:rbt n s

2 * n + redcarac s <= 2 * (n + redcarac s)
s:tree
n:nat
H:rbt n s
2 * (n + redcarac s) <= 2 * Nat.log2 (S (cardinal s))
s:tree
n:nat
H:rbt n s

2 * n + redcarac s <= 2 * (n + redcarac s)
s:tree
n:nat
H:rbt n s

2 * n + redcarac s <= 2 * n + 2 * redcarac s
s:tree
n:nat
H:rbt n s

redcarac s <= 2 * redcarac s
s:tree
n:nat
H:rbt n s

1 * redcarac s <= 2 * redcarac s
s:tree
n:nat
H:rbt n s

1 <= 2
auto with arith.
s:tree
n:nat
H:rbt n s

2 * (n + redcarac s) <= 2 * Nat.log2 (S (cardinal s))
s:tree
n:nat
H:rbt n s

n + redcarac s <= Nat.log2 (S (cardinal s))
s:tree
n:nat
H:rbt n s

n + redcarac s <= mindepth s
s:tree
n:nat
H:rbt n s
mindepth s <= Nat.log2 (S (cardinal s))
s:tree
n:nat
H:rbt n s

n + redcarac s <= mindepth s
now apply rb_mindepth.
s:tree
n:nat
H:rbt n s

mindepth s <= Nat.log2 (S (cardinal s))
apply mindepth_log_cardinal. Qed.
s:tree

s <> Leaf -> Nat.log2 (cardinal s) < maxdepth s
s:tree

s <> Leaf -> Nat.log2 (cardinal s) < maxdepth s
apply maxdepth_log_cardinal. Qed.

Singleton

x:elt

Rbt (singleton x)
x:elt

Rbt (singleton x)
x:elt

Rbt (Bk Leaf x Leaf)
exists 1; auto. Qed.

makeBlack and makeRed

n:nat
t:tree

arbt n t -> Rbt (makeBlack t)
n:nat
t:tree

arbt n t -> Rbt (makeBlack t)
n:nat

arbt n Leaf -> Rbt (makeBlack Leaf)
n:nat
l:tree
x:X.t
r:tree
arbt n (Rd l x r) -> Rbt (makeBlack (Rd l x r))
n:nat
l:tree
x:X.t
r:tree
arbt n (Bk l x r) -> Rbt (makeBlack (Bk l x r))
n:nat

arbt n Leaf -> Rbt (makeBlack Leaf)
exists 0; auto.
n:nat
l:tree
x:X.t
r:tree

arbt n (Rd l x r) -> Rbt (makeBlack (Rd l x r))
destruct 1; invrb; exists (S n); simpl; auto.
n:nat
l:tree
x:X.t
r:tree

arbt n (Bk l x r) -> Rbt (makeBlack (Bk l x r))
exists n; auto. Qed.
t:tree
n:nat

rbt (S n) t -> notred t -> rrt n (makeRed t)
t:tree
n:nat

rbt (S n) t -> notred t -> rrt n (makeRed t)
destruct t as [|[|] l x r]; invrb; simpl; auto. Qed.

Balancing

n:nat
l:tree
k:X.t
r:tree

arbt n l -> rbt n r -> rbt (S n) (lbal l k r)
n:nat
l:tree
k:X.t
r:tree

arbt n l -> rbt n r -> rbt (S n) (lbal l k r)
case lbal_match; intros; desarb; invrb; auto. Qed.
n:nat
l:tree
k:X.t
r:tree

rbt n l -> arbt n r -> rbt (S n) (rbal l k r)
n:nat
l:tree
k:X.t
r:tree

rbt n l -> arbt n r -> rbt (S n) (rbal l k r)
case rbal_match; intros; desarb; invrb; auto. Qed.
n:nat
l:tree
k:X.t
r:tree

rbt n l -> arbt n r -> rbt (S n) (rbal' l k r)
n:nat
l:tree
k:X.t
r:tree

rbt n l -> arbt n r -> rbt (S n) (rbal' l k r)
case rbal'_match; intros; desarb; invrb; auto. Qed.
n:nat
l:tree
x:X.t
r:tree

arbt n l -> rbt (S n) r -> notred r -> rbt (S n) (lbalS l x r)
n:nat
l:tree
x:X.t
r:tree

arbt n l -> rbt (S n) r -> notred r -> rbt (S n) (lbalS l x r)
n:nat
l:tree
x:X.t
r:tree
Hl:arbt n l
Hr:rbt (S n) r
Hr':notred r

rbt (S n) (lbalS l x r)
n:nat
l:tree
x:X.t
rl:tree
rx:X.t
rr:tree
Hl:arbt n l
Hr':notred (Bk rl rx rr)
H2:rbt n rl
H4:rbt n rr

rbt (S n) (lbalS l x (Bk rl rx rr))
n:nat
l:tree
x:X.t
rl:tree
rx:X.t
rr:tree
Hl:arbt n l
H2:rbt n rl
H4:rbt n rr

rbt (S n) (lbalS l x (Bk rl rx rr))
n:nat
l:tree
x:X.t
rl:tree
rx:X.t
rr:tree
H2:rbt n rl
H4:rbt n rr

arbt n l -> rbt (S n) (lbalS l x (Bk rl rx rr))
n:nat
l:tree
x:X.t
rl:tree
rx:X.t
rr:tree
H2:rbt n rl
H4:rbt n rr

forall (a : tree) (x0 : X.t) (b : tree), arbt n (Rd a x0 b) -> rbt (S n) (Rd (Bk a x0 b) x (Bk rl rx rr))
n:nat
l:tree
x:X.t
rl:tree
rx:X.t
rr:tree
H2:rbt n rl
H4:rbt n rr
forall t0 : tree, notred t0 -> arbt n t0 -> rbt (S n) (rbal' t0 x (Rd rl rx rr))
n:nat
l:tree
x:X.t
rl:tree
rx:X.t
rr:tree
H2:rbt n rl
H4:rbt n rr

forall (a : tree) (x0 : X.t) (b : tree), arbt n (Rd a x0 b) -> rbt (S n) (Rd (Bk a x0 b) x (Bk rl rx rr))
destruct 1; invrb; auto.
n:nat
l:tree
x:X.t
rl:tree
rx:X.t
rr:tree
H2:rbt n rl
H4:rbt n rr

forall t0 : tree, notred t0 -> arbt n t0 -> rbt (S n) (rbal' t0 x (Rd rl rx rr))
n:nat
l:tree
x:X.t
rl:tree
rx:X.t
rr:tree
H2:rbt n rl
H4:rbt n rr
t0:tree
H:notred t0
Hl:arbt n t0

rbt (S n) (rbal' t0 x (Rd rl rx rr))
apply rbal'_rb; auto. Qed.
n:nat
l:tree
x:X.t
r:tree

arbt n l -> rbt (S n) r -> arbt (S n) (lbalS l x r)
n:nat
l:tree
x:X.t
r:tree

arbt n l -> rbt (S n) r -> arbt (S n) (lbalS l x r)
n:nat
l:tree
x:X.t
r:tree

forall (a : tree) (x0 : X.t) (b : tree), arbt n (Rd a x0 b) -> rbt (S n) r -> arbt (S n) (Rd (Bk a x0 b) x r)
n:nat
l:tree
x:X.t
r:tree
forall t0 : tree, notred t0 -> arbt n t0 -> rbt (S n) r -> arbt (S n) match r with | Rd (Bk a0 y b) z c => Rd (Bk t0 x a0) y (rbal' b z (makeRed c)) | Bk a z c => rbal' t0 x (Rd a z c) | _ => Rd t0 x r end
n:nat
l:tree
x:X.t
r:tree

forall (a : tree) (x0 : X.t) (b : tree), arbt n (Rd a x0 b) -> rbt (S n) r -> arbt (S n) (Rd (Bk a x0 b) x r)
destruct 1; invrb; auto.
n:nat
l:tree
x:X.t
r:tree

forall t0 : tree, notred t0 -> arbt n t0 -> rbt (S n) r -> arbt (S n) match r with | Rd (Bk a0 y b) z c => Rd (Bk t0 x a0) y (rbal' b z (makeRed c)) | Bk a z c => rbal' t0 x (Rd a z c) | _ => Rd t0 x r end
n:nat
x:X.t
r:tree

forall t0 : tree, notred t0 -> arbt n t0 -> rbt (S n) r -> arbt (S n) match r with | Rd (Bk a0 y b) z c => Rd (Bk t0 x a0) y (rbal' b z (makeRed c)) | Bk a z c => rbal' t0 x (Rd a z c) | _ => Rd t0 x r end
n:nat
x:X.t
r, l:tree
Hl:notred l
Hl':arbt n l
Hr:rbt (S n) r

arbt (S n) match r with | Rd (Bk a0 y b) z c => Rd (Bk l x a0) y (rbal' b z (makeRed c)) | Bk a z c => rbal' l x (Rd a z c) | _ => Rd l x r end
n:nat
x:X.t
rl:tree
rx:X.t
rr, l:tree
Hl:notred l
Hl':arbt n l
H2:notred rl
H4:notred rr
H5:rbt (S n) rl
H6:rbt (S n) rr

arbt (S n) match rl with | Bk a y b => Rd (Bk l x a) y (rbal' b rx (makeRed rr)) | _ => Rd l x (Rd rl rx rr) end
n:nat
x:X.t
rl:tree
rx:X.t
rr, l:tree
Hl:notred l
Hl':arbt n l
H2:rbt n rl
H4:rbt n rr
arbt (S n) (rbal' l x (Rd rl rx rr))
n:nat
x:X.t
rl:tree
rx:X.t
rr, l:tree
Hl:notred l
Hl':arbt n l
H2:notred rl
H4:notred rr
H5:rbt (S n) rl
H6:rbt (S n) rr

arbt (S n) match rl with | Bk a y b => Rd (Bk l x a) y (rbal' b rx (makeRed rr)) | _ => Rd l x (Rd rl rx rr) end
n:nat
x:X.t
rll:tree
rlx:X.t
rlr:tree
rx:X.t
rr, l:tree
Hl:notred l
Hl':arbt n l
H2:notred (Bk rll rlx rlr)
H4:notred rr
H6:rbt (S n) rr
H3:rbt n rll
H8:rbt n rlr

arbt (S n) (Rd (Bk l x rll) rlx (rbal' rlr rx (makeRed rr)))
right; auto using rbal'_rb, makeRed_rr.
n:nat
x:X.t
rl:tree
rx:X.t
rr, l:tree
Hl:notred l
Hl':arbt n l
H2:rbt n rl
H4:rbt n rr

arbt (S n) (rbal' l x (Rd rl rx rr))
left; apply rbal'_rb; auto. Qed.
n:nat
l:tree
x:X.t
r:tree

rbt (S n) l -> notred l -> arbt n r -> rbt (S n) (rbalS l x r)
n:nat
l:tree
x:X.t
r:tree

rbt (S n) l -> notred l -> arbt n r -> rbt (S n) (rbalS l x r)
n:nat
l:tree
x:X.t
r:tree
Hl:rbt (S n) l
Hl':notred l
Hr:arbt n r

rbt (S n) (rbalS l x r)
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
Hl':notred (Bk ll lx lr)
Hr:arbt n r
H2:rbt n ll
H4:rbt n lr

rbt (S n) (rbalS (Bk ll lx lr) x r)
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
Hr:arbt n r
H2:rbt n ll
H4:rbt n lr

rbt (S n) (rbalS (Bk ll lx lr) x r)
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
H2:rbt n ll
H4:rbt n lr

arbt n r -> rbt (S n) (rbalS (Bk ll lx lr) x r)
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
H2:rbt n ll
H4:rbt n lr

forall (a : tree) (x0 : X.t) (b : tree), arbt n (Rd a x0 b) -> rbt (S n) (Rd (Bk ll lx lr) x (Bk a x0 b))
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
H2:rbt n ll
H4:rbt n lr
forall t0 : tree, notred t0 -> arbt n t0 -> rbt (S n) (lbal (Rd ll lx lr) x t0)
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
H2:rbt n ll
H4:rbt n lr

forall (a : tree) (x0 : X.t) (b : tree), arbt n (Rd a x0 b) -> rbt (S n) (Rd (Bk ll lx lr) x (Bk a x0 b))
destruct 1; invrb; auto.
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
H2:rbt n ll
H4:rbt n lr

forall t0 : tree, notred t0 -> arbt n t0 -> rbt (S n) (lbal (Rd ll lx lr) x t0)
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
H2:rbt n ll
H4:rbt n lr
t0:tree
H:notred t0
Hr:arbt n t0

rbt (S n) (lbal (Rd ll lx lr) x t0)
apply lbal_rb; auto. Qed.
n:nat
l:tree
x:X.t
r:tree

rbt (S n) l -> arbt n r -> arbt (S n) (rbalS l x r)
n:nat
l:tree
x:X.t
r:tree

rbt (S n) l -> arbt n r -> arbt (S n) (rbalS l x r)
n:nat
l:tree
x:X.t
r:tree

forall (a : tree) (x0 : X.t) (b : tree), rbt (S n) l -> arbt n (Rd a x0 b) -> arbt (S n) (Rd l x (Bk a x0 b))
n:nat
l:tree
x:X.t
r:tree
forall t0 : tree, notred t0 -> rbt (S n) l -> arbt n t0 -> arbt (S n) match l with | Leaf => Rd l x t0 | Rd a y Leaf | Rd a y (Rd _ _ _) => Rd l x t0 | Rd a y (Bk b0 z c) => Rd (lbal (makeRed a) y b0) z (Bk c x t0) | Bk a y b => lbal (Rd a y b) x t0 end
n:nat
l:tree
x:X.t
r:tree

forall (a : tree) (x0 : X.t) (b : tree), rbt (S n) l -> arbt n (Rd a x0 b) -> arbt (S n) (Rd l x (Bk a x0 b))
destruct 2; invrb; auto.
n:nat
l:tree
x:X.t
r:tree

forall t0 : tree, notred t0 -> rbt (S n) l -> arbt n t0 -> arbt (S n) match l with | Leaf => Rd l x t0 | Rd a y Leaf | Rd a y (Rd _ _ _) => Rd l x t0 | Rd a y (Bk b0 z c) => Rd (lbal (makeRed a) y b0) z (Bk c x t0) | Bk a y b => lbal (Rd a y b) x t0 end
n:nat
l:tree
x:X.t

forall t0 : tree, notred t0 -> rbt (S n) l -> arbt n t0 -> arbt (S n) match l with | Leaf => Rd l x t0 | Rd a y Leaf | Rd a y (Rd _ _ _) => Rd l x t0 | Rd a y (Bk b0 z c) => Rd (lbal (makeRed a) y b0) z (Bk c x t0) | Bk a y b => lbal (Rd a y b) x t0 end
n:nat
l:tree
x:X.t
r:tree
Hr:notred r
Hr':rbt (S n) l
Hl:arbt n r

arbt (S n) match l with | Leaf => Rd l x r | Rd a y Leaf | Rd a y (Rd _ _ _) => Rd l x r | Rd a y (Bk b0 z c) => Rd (lbal (makeRed a) y b0) z (Bk c x r) | Bk a y b => lbal (Rd a y b) x r end
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
Hr:notred r
Hl:arbt n r
H2:notred ll
H4:notred lr
H5:rbt (S n) ll
H6:rbt (S n) lr

arbt (S n) match lr with | Bk b z c => Rd (lbal (makeRed ll) lx b) z (Bk c x r) | _ => Rd (Rd ll lx lr) x r end
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
Hr:notred r
Hl:arbt n r
H2:rbt n ll
H4:rbt n lr
arbt (S n) (lbal (Rd ll lx lr) x r)
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
Hr:notred r
Hl:arbt n r
H2:notred ll
H4:notred lr
H5:rbt (S n) ll
H6:rbt (S n) lr

arbt (S n) match lr with | Bk b z c => Rd (lbal (makeRed ll) lx b) z (Bk c x r) | _ => Rd (Rd ll lx lr) x r end
n:nat
ll:tree
lx:X.t
lrl:tree
lrx:X.t
lrr:tree
x:X.t
r:tree
Hr:notred r
Hl:arbt n r
H2:notred ll
H4:notred (Bk lrl lrx lrr)
H5:rbt (S n) ll
H3:rbt n lrl
H8:rbt n lrr

arbt (S n) (Rd (lbal (makeRed ll) lx lrl) lrx (Bk lrr x r))
right; auto using lbal_rb, makeRed_rr.
n:nat
ll:tree
lx:X.t
lr:tree
x:X.t
r:tree
Hr:notred r
Hl:arbt n r
H2:rbt n ll
H4:rbt n lr

arbt (S n) (lbal (Rd ll lx lr) x r)
left; apply lbal_rb; auto. Qed.

Insertion

The next lemmas combine simultaneous results about rbt and arbt. A first solution here: statement with if ... then ... else
Definition ifred s (A B:Prop) := rcase (fun _ _ _ => A) (fun _ => B) s.

s:tree
A, B:Prop

notred s -> ifred s A B <-> B
s:tree
A, B:Prop

notred s -> ifred s A B <-> B
destruct s; descolor; simpl; intuition. Qed.
s:tree
A, B:Prop

ifred s A B -> A \/ B
s:tree
A, B:Prop

ifred s A B -> A \/ B
destruct s; descolor; simpl; intuition. Qed.
x:X.t
s:tree
n:nat

rbt n s -> ifred s (rrt n (ins x s)) (rbt n (ins x s))
x:X.t
s:tree
n:nat

rbt n s -> ifred s (rrt n (ins x s)) (rbt n (ins x s))
x:X.t

ifred Leaf (rrt 0 (ins x Leaf)) (rbt 0 (ins x Leaf))
x:X.t
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHrbt2:ifred r (rrt n (ins x r)) (rbt n (ins x r))
ifred (Rd l k r) (rrt n (ins x (Rd l k r))) (rbt n (ins x (Rd l k r)))
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))
ifred (Bk l k r) (rrt (S n) (ins x (Bk l k r))) (rbt (S n) (ins x (Bk l k r)))
x:X.t

ifred Leaf (rrt 0 (ins x Leaf)) (rbt 0 (ins x Leaf))
simpl; auto.
x:X.t
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHrbt2:ifred r (rrt n (ins x r)) (rbt n (ins x r))

ifred (Rd l k r) (rrt n (ins x (Rd l k r))) (rbt n (ins x (Rd l k r)))
x:X.t
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHrbt2:ifred r (rrt n (ins x r)) (rbt n (ins x r))

rrt n match X.compare x k with | Eq => Rd l k r | Lt => Rd (ins x l) k r | Gt => Rd l k (ins x r) end
x:X.t
n:nat
l:tree
k:X.t
r:tree
n0:notred l
n1:notred r
H:rbt n l
H0:rbt n r
IHrbt1:rbt n (ins x l)
IHrbt2:rbt n (ins x r)

rrt n match X.compare x k with | Eq => Rd l k r | Lt => Rd (ins x l) k r | Gt => Rd l k (ins x r) end
elim_compare x k; auto.
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))

ifred (Bk l k r) (rrt (S n) (ins x (Bk l k r))) (rbt (S n) (ins x (Bk l k r)))
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))

rbt (S n) (ins x (Bk l k r))
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))

rbt (S n) match X.compare x k with | Eq => Bk l k r | Lt => lbal (ins x l) k r | Gt => rbal l k (ins x r) end
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))
H:X.eq x k

rbt (S n) (Bk l k r)
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))
H:X.lt x k
rbt (S n) (lbal (ins x l) k r)
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))
H:X.lt k x
rbt (S n) (rbal l k (ins x r))
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))
H:X.eq x k

rbt (S n) (Bk l k r)
auto.
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))
H:X.lt x k

rbt (S n) (lbal (ins x l) k r)
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))
H:X.lt x k

arbt n (ins x l)
apply ifred_or in IHl; intuition.
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))
H:X.lt k x

rbt (S n) (rbal l k (ins x r))
x:X.t
n:nat
l:tree
k:X.t
r:tree
Hl:rbt n l
Hr:rbt n r
IHl:ifred l (rrt n (ins x l)) (rbt n (ins x l))
IHr:ifred r (rrt n (ins x r)) (rbt n (ins x r))
H:X.lt k x

arbt n (ins x r)
apply ifred_or in IHr; intuition. Qed.
x:X.t
s:tree
n:nat

rbt n s -> arbt n (ins x s)
x:X.t
s:tree
n:nat

rbt n s -> arbt n (ins x s)
x:X.t
s:tree
n:nat
H:rbt n s

arbt n (ins x s)
x:X.t
s:tree
n:nat
H:rrt n (ins x s) \/ rbt n (ins x s)

arbt n (ins x s)
intuition. Qed.
x:X.t
s:tree

Rbt s -> Rbt (add x s)
x:X.t
s:tree

Rbt s -> Rbt (add x s)
x:X.t
s:tree
n:nat
H:rbt n s

Rbt (add x s)
x:X.t
s:tree
n:nat
H:rbt n s

Rbt (makeBlack (ins x s))
now apply (makeBlack_rb n), ins_arb. Qed.

Deletion

A second approach here: statement with ... /\ ...
n:nat
l, r:tree

rbt n l -> rbt n r -> arbt n (append l r) /\ (notred l -> notred r -> rbt n (append l r))
n:nat
l, r:tree

rbt n l -> rbt n r -> arbt n (append l r) /\ (notred l -> notred r -> rbt n (append l r))
l:tree

forall (r : tree) (n : nat), rbt n l -> rbt n r -> arbt n (append l r) /\ (notred l -> notred r -> rbt n (append l r))
r:tree

forall n : nat, rbt n Leaf -> rbt n r -> arbt n r /\ (True -> notred r -> rbt n r)
lc:Color.t
ll:tree
lx:X.t
lr:tree
IHlr:forall (r : tree) (n : nat), rbt n lr -> rbt n r -> arbt n (append lr r) /\ (notred lr -> notred r -> rbt n (append lr r))
forall n : nat, rbt n (Node lc ll lx lr) -> rbt n Leaf -> arbt n (Node lc ll lx lr) /\ (match lc with | Red => False | Black => True end -> True -> rbt n (Node lc ll lx lr))
ll:tree
lx:X.t
lr, rl:tree
IHlr:forall n : nat, rbt n lr -> rbt n rl -> arbt n (append lr rl) /\ (notred lr -> notred rl -> rbt n (append lr rl))
rx:X.t
rr:tree
forall n : nat, rbt n (Rd ll lx lr) -> rbt n (Rd rl rx rr) -> arbt n (append (Rd ll lx lr) (Rd rl rx rr)) /\ (notred (Rd ll lx lr) -> notred (Rd rl rx rr) -> rbt n (append (Rd ll lx lr) (Rd rl rx rr)))
ll:tree
lx:X.t
lr, r:tree
IHlr:forall n : nat, rbt n lr -> rbt n r -> arbt n (append lr r) /\ (notred lr -> notred r -> rbt n (append lr r))
Hr:notred r
forall n : nat, rbt n (Rd ll lx lr) -> rbt n r -> arbt n (Rd ll lx (append lr r)) /\ (False -> True -> rbt n (Rd ll lx (append lr r)))
rl:tree
rx:X.t
rr, l:tree
IHrl:forall n : nat, rbt n l -> rbt n rl -> arbt n (append l rl) /\ (notred l -> notred rl -> rbt n (append l rl))
Hl:notred l
forall n : nat, rbt n l -> rbt n (Rd rl rx rr) -> arbt n (Rd (append l rl) rx rr) /\ (notred l -> notred (Rd rl rx rr) -> rbt n (Rd (append l rl) rx rr))
ll:tree
lx:X.t
lr, rl:tree
IHlr:forall n : nat, rbt n lr -> rbt n rl -> arbt n (append lr rl) /\ (notred lr -> notred rl -> rbt n (append lr rl))
rx:X.t
rr:tree
forall n : nat, rbt n (Bk ll lx lr) -> rbt n (Bk rl rx rr) -> arbt n (append (Bk ll lx lr) (Bk rl rx rr)) /\ (notred (Bk ll lx lr) -> notred (Bk rl rx rr) -> rbt n (append (Bk ll lx lr) (Bk rl rx rr)))
r:tree

forall n : nat, rbt n Leaf -> rbt n r -> arbt n r /\ (True -> notred r -> rbt n r)
split; auto.
lc:Color.t
ll:tree
lx:X.t
lr:tree
IHlr:forall (r : tree) (n : nat), rbt n lr -> rbt n r -> arbt n (append lr r) /\ (notred lr -> notred r -> rbt n (append lr r))

forall n : nat, rbt n (Node lc ll lx lr) -> rbt n Leaf -> arbt n (Node lc ll lx lr) /\ (match lc with | Red => False | Black => True end -> True -> rbt n (Node lc ll lx lr))
split; auto.
ll:tree
lx:X.t
lr, rl:tree
IHlr:forall n : nat, rbt n lr -> rbt n rl -> arbt n (append lr rl) /\ (notred lr -> notred rl -> rbt n (append lr rl))
rx:X.t
rr:tree

forall n : nat, rbt n (Rd ll lx lr) -> rbt n (Rd rl rx rr) -> arbt n (append (Rd ll lx lr) (Rd rl rx rr)) /\ (notred (Rd ll lx lr) -> notred (Rd rl rx rr) -> rbt n (append (Rd ll lx lr) (Rd rl rx rr)))
ll:tree
lx:X.t
lr, rl:tree
IHlr:forall n0 : nat, rbt n0 lr -> rbt n0 rl -> arbt n0 (append lr rl) /\ (notred lr -> notred rl -> rbt n0 (append lr rl))
rx:X.t
rr:tree
n:nat

rbt n (Rd ll lx lr) -> rbt n (Rd rl rx rr) -> arbt n (append (Rd ll lx lr) (Rd rl rx rr)) /\ (notred (Rd ll lx lr) -> notred (Rd rl rx rr) -> rbt n (append (Rd ll lx lr) (Rd rl rx rr)))
ll:tree
lx:X.t
lr, rl:tree
IHlr:forall n0 : nat, rbt n0 lr -> rbt n0 rl -> arbt n0 (append lr rl) /\ (notred lr -> notred rl -> rbt n0 (append lr rl))
rx:X.t
rr:tree
n:nat
H4:notred rl
H6:notred rr
H7:rbt n rl
H8:rbt n rr
H3:notred ll
H9:notred lr
H10:rbt n ll
H11:rbt n lr

arbt n (append (Rd ll lx lr) (Rd rl rx rr)) /\ (notred (Rd ll lx lr) -> notred (Rd rl rx rr) -> rbt n (append (Rd ll lx lr) (Rd rl rx rr)))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:notred rl
H6:notred rr
H7:rbt n rl
H8:rbt n rr
H3:notred ll
H9:notred lr
H10:rbt n ll
H11:rbt n lr

arbt n (append lr rl) -> (notred lr -> notred rl -> rbt n (append lr rl)) -> arbt n (append (Rd ll lx lr) (Rd rl rx rr)) /\ (notred (Rd ll lx lr) -> notred (Rd rl rx rr) -> rbt n (append (Rd ll lx lr) (Rd rl rx rr)))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:notred rl
H6:notred rr
H7:rbt n rl
H8:rbt n rr
H3:notred ll
H9:notred lr
H10:rbt n ll
H11:rbt n lr

forall (a : tree) (x : X.t) (b : tree), arbt n (Rd a x b) -> (notred lr -> notred rl -> rbt n (Rd a x b)) -> arbt n (Rd (Rd ll lx a) x (Rd b rx rr)) /\ (notred (Rd ll lx lr) -> notred (Rd rl rx rr) -> rbt n (Rd (Rd ll lx a) x (Rd b rx rr)))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:notred rl
H6:notred rr
H7:rbt n rl
H8:rbt n rr
H3:notred ll
H9:notred lr
H10:rbt n ll
H11:rbt n lr
forall t0 : tree, notred t0 -> arbt n t0 -> (notred lr -> notred rl -> rbt n t0) -> arbt n (Rd ll lx (Rd t0 rx rr)) /\ (notred (Rd ll lx lr) -> notred (Rd rl rx rr) -> rbt n (Rd ll lx (Rd t0 rx rr)))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:notred rl
H6:notred rr
H7:rbt n rl
H8:rbt n rr
H3:notred ll
H9:notred lr
H10:rbt n ll
H11:rbt n lr

forall (a : tree) (x : X.t) (b : tree), arbt n (Rd a x b) -> (notred lr -> notred rl -> rbt n (Rd a x b)) -> arbt n (Rd (Rd ll lx a) x (Rd b rx rr)) /\ (notred (Rd ll lx lr) -> notred (Rd rl rx rr) -> rbt n (Rd (Rd ll lx a) x (Rd b rx rr)))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:notred rl
H6:notred rr
H7:rbt n rl
H8:rbt n rr
H3:notred ll
H9:notred lr
H10:rbt n ll
H11:rbt n lr
a:tree
x:X.t
b:tree
H:notred lr -> notred rl -> rbt n (Rd a x b)

arbt n (Rd (Rd ll lx a) x (Rd b rx rr))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:notred rl
H6:notred rr
H7:rbt n rl
H8:rbt n rr
H3:notred ll
H9:notred lr
H10:rbt n ll
H11:rbt n lr
a:tree
x:X.t
b:tree
H:notred lr -> notred rl -> rbt n (Rd a x b)
H0:rbt n (Rd a x b)

arbt n (Rd (Rd ll lx a) x (Rd b rx rr))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:notred rl
H6:notred rr
H7:rbt n rl
H8:rbt n rr
H3:notred ll
H9:notred lr
H10:rbt n ll
H11:rbt n lr
a:tree
x:X.t
b:tree
H:notred lr -> notred rl -> rbt n (Rd a x b)
H12:notred a
H14:notred b
H15:rbt n a
H16:rbt n b

arbt n (Rd (Rd ll lx a) x (Rd b rx rr))
auto.
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:notred rl
H6:notred rr
H7:rbt n rl
H8:rbt n rr
H3:notred ll
H9:notred lr
H10:rbt n ll
H11:rbt n lr

forall t0 : tree, notred t0 -> arbt n t0 -> (notred lr -> notred rl -> rbt n t0) -> arbt n (Rd ll lx (Rd t0 rx rr)) /\ (notred (Rd ll lx lr) -> notred (Rd rl rx rr) -> rbt n (Rd ll lx (Rd t0 rx rr)))
split; invrb; auto.
ll:tree
lx:X.t
lr, r:tree
IHlr:forall n : nat, rbt n lr -> rbt n r -> arbt n (append lr r) /\ (notred lr -> notred r -> rbt n (append lr r))
Hr:notred r

forall n : nat, rbt n (Rd ll lx lr) -> rbt n r -> arbt n (Rd ll lx (append lr r)) /\ (False -> True -> rbt n (Rd ll lx (append lr r)))
ll:tree
lx:X.t
lr, r:tree
IHlr:forall n0 : nat, rbt n0 lr -> rbt n0 r -> arbt n0 (append lr r) /\ (notred lr -> notred r -> rbt n0 (append lr r))
Hr:notred r
n:nat
H0:rbt n r
H4:notred ll
H6:notred lr
H7:rbt n ll
H8:rbt n lr

arbt n (Rd ll lx (append lr r))
destruct (IHlr n) as (_,IH); auto.
rl:tree
rx:X.t
rr, l:tree
IHrl:forall n : nat, rbt n l -> rbt n rl -> arbt n (append l rl) /\ (notred l -> notred rl -> rbt n (append l rl))
Hl:notred l

forall n : nat, rbt n l -> rbt n (Rd rl rx rr) -> arbt n (Rd (append l rl) rx rr) /\ (notred l -> notred (Rd rl rx rr) -> rbt n (Rd (append l rl) rx rr))
rl:tree
rx:X.t
rr, l:tree
IHrl:forall n0 : nat, rbt n0 l -> rbt n0 rl -> arbt n0 (append l rl) /\ (notred l -> notred rl -> rbt n0 (append l rl))
Hl:notred l
n:nat
H:rbt n l
H4:notred rl
H6:notred rr
H7:rbt n rl
H8:rbt n rr

arbt n (Rd (append l rl) rx rr)
destruct (IHrl n) as (_,IH); auto.
ll:tree
lx:X.t
lr, rl:tree
IHlr:forall n : nat, rbt n lr -> rbt n rl -> arbt n (append lr rl) /\ (notred lr -> notred rl -> rbt n (append lr rl))
rx:X.t
rr:tree

forall n : nat, rbt n (Bk ll lx lr) -> rbt n (Bk rl rx rr) -> arbt n (append (Bk ll lx lr) (Bk rl rx rr)) /\ (notred (Bk ll lx lr) -> notred (Bk rl rx rr) -> rbt n (append (Bk ll lx lr) (Bk rl rx rr)))
ll:tree
lx:X.t
lr, rl:tree
IHlr:forall n0 : nat, rbt n0 lr -> rbt n0 rl -> arbt n0 (append lr rl) /\ (notred lr -> notred rl -> rbt n0 (append lr rl))
rx:X.t
rr:tree
n:nat

rbt (S n) (Bk ll lx lr) -> rbt (S n) (Bk rl rx rr) -> arbt (S n) (append (Bk ll lx lr) (Bk rl rx rr)) /\ (notred (Bk ll lx lr) -> notred (Bk rl rx rr) -> rbt (S n) (append (Bk ll lx lr) (Bk rl rx rr)))
ll:tree
lx:X.t
lr, rl:tree
IHlr:forall n0 : nat, rbt n0 lr -> rbt n0 rl -> arbt n0 (append lr rl) /\ (notred lr -> notred rl -> rbt n0 (append lr rl))
rx:X.t
rr:tree
n:nat
H4:rbt n rl
H6:rbt n rr
H3:rbt n ll
H7:rbt n lr

arbt (S n) (append (Bk ll lx lr) (Bk rl rx rr)) /\ (notred (Bk ll lx lr) -> notred (Bk rl rx rr) -> rbt (S n) (append (Bk ll lx lr) (Bk rl rx rr)))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:rbt n rl
H6:rbt n rr
H3:rbt n ll
H7:rbt n lr
IH:arbt n (append lr rl)

arbt (S n) (append (Bk ll lx lr) (Bk rl rx rr)) /\ (notred (Bk ll lx lr) -> notred (Bk rl rx rr) -> rbt (S n) (append (Bk ll lx lr) (Bk rl rx rr)))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:rbt n rl
H6:rbt n rr
H3:rbt n ll
H7:rbt n lr

arbt n (append lr rl) -> arbt (S n) (append (Bk ll lx lr) (Bk rl rx rr)) /\ (notred (Bk ll lx lr) -> notred (Bk rl rx rr) -> rbt (S n) (append (Bk ll lx lr) (Bk rl rx rr)))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:rbt n rl
H6:rbt n rr
H3:rbt n ll
H7:rbt n lr

forall (a : tree) (x : X.t) (b : tree), arbt n (Rd a x b) -> arbt (S n) (Rd (Bk ll lx a) x (Bk b rx rr)) /\ (notred (Bk ll lx lr) -> notred (Bk rl rx rr) -> rbt (S n) (Rd (Bk ll lx a) x (Bk b rx rr)))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:rbt n rl
H6:rbt n rr
H3:rbt n ll
H7:rbt n lr
forall t0 : tree, notred t0 -> arbt n t0 -> arbt (S n) (lbalS ll lx (Bk t0 rx rr)) /\ (notred (Bk ll lx lr) -> notred (Bk rl rx rr) -> rbt (S n) (lbalS ll lx (Bk t0 rx rr)))
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:rbt n rl
H6:rbt n rr
H3:rbt n ll
H7:rbt n lr

forall (a : tree) (x : X.t) (b : tree), arbt n (Rd a x b) -> arbt (S n) (Rd (Bk ll lx a) x (Bk b rx rr)) /\ (notred (Bk ll lx lr) -> notred (Bk rl rx rr) -> rbt (S n) (Rd (Bk ll lx a) x (Bk b rx rr)))
intros a x b IH; split; destruct IH; invrb; auto.
ll:tree
lx:X.t
lr, rl:tree
rx:X.t
rr:tree
n:nat
H4:rbt n rl
H6:rbt n rr
H3:rbt n ll
H7:rbt n lr

forall t0 : tree, notred t0 -> arbt n t0 -> arbt (S n) (lbalS ll lx (Bk t0 rx rr)) /\ (notred (Bk ll lx lr) -> notred (Bk rl rx rr) -> rbt (S n) (lbalS ll lx (Bk t0 rx rr)))
split; [left | invrb]; auto using lbalS_rb. Qed.
A third approach : Lemma ... with ...
del_arb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s0 -> isblack s0 -> arbt n0 (del x0 s0)
del_rb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt n0 s0 -> notblack s0 -> rbt n0 (del x0 s0)
s:tree
x:X.t
n:nat

rbt (S n) s -> isblack s -> arbt n (del x s)
del_arb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s0 -> isblack s0 -> arbt n0 (del x0 s0)
del_rb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt n0 s0 -> notblack s0 -> rbt n0 (del x0 s0)
s:tree
x:X.t
n:nat
rbt n s -> notblack s -> rbt n (del x s)
del_arb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s0 -> isblack s0 -> arbt n0 (del x0 s0)
del_rb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt n0 s0 -> notblack s0 -> rbt n0 (del x0 s0)
s:tree
x:X.t
n:nat

rbt (S n) s -> isblack s -> arbt n (del x s)
del_arb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s0 -> isblack s0 -> arbt n0 (del x0 s0)
del_rb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt n0 s0 -> notblack s0 -> rbt n0 (del x0 s0)
s:tree
x:X.t
n:nat
rbt n s -> notblack s -> rbt n (del x s)
del_arb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s0 -> isblack s0 -> arbt n0 (del x0 s0)
del_rb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt n0 s0 -> notblack s0 -> rbt n0 (del x0 s0)
s:tree
x:X.t
n:nat

rbt (S n) s -> isblack s -> arbt n (del x s)
del_arb:forall (s0 : tree) (x0 : X.t) (n : nat), rbt (S n) s0 -> isblack s0 -> arbt n (del x0 s0)
del_rb:forall (s0 : tree) (x0 : X.t) (n : nat), rbt n s0 -> notblack s0 -> rbt n (del x0 s0)
s:tree
x:X.t

forall n : nat, rbt (S n) s -> isblack s -> arbt n (del x s)
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)
IHr:forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)
n:nat
H0:True
H1:X.eq x x'
H5:rbt n l
H8:rbt n r

arbt n (append l r)
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)
IHr:forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)
n:nat
H0:True
H1:X.lt x x'
H5:rbt n l
H8:rbt n r
arbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)
IHr:forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)
n:nat
H0:True
H1:X.lt x' x
H5:rbt n l
H8:rbt n r
arbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)
IHr:forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)
n:nat
H0:True
H1:X.eq x x'
H5:rbt n l
H8:rbt n r

arbt n (append l r)
apply append_arb_rb; assumption.
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)
IHr:forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)
n:nat
H0:True
H1:X.lt x x'
H5:rbt n l
H8:rbt n r

arbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)
IHr:forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)
n:nat
H0:True
H1:X.lt x x'
H5:rbt n l
H8:rbt n r
IHl':forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)

arbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)
n:nat
H0:True
H1:X.lt x x'
H5:rbt n l
H8:rbt n r
IHl':forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)

arbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
ll:tree
lx:X.t
lr:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt (S n0) (Bk ll lx lr) -> isblack (Bk ll lx lr) -> arbt n0 (del x (Bk ll lx lr))
n:nat
H0:True
H1:X.lt x x'
H5:rbt n (Bk ll lx lr)
H8:rbt n r
IHl':forall n0 : nat, rbt n0 (Bk ll lx lr) -> notblack (Bk ll lx lr) -> rbt n0 (del x (Bk ll lx lr))

arbt n (lbalS (del x (Bk ll lx lr)) x' r)
ll:tree
lx:X.t
lr:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt (S n0) (Bk ll lx lr) -> isblack (Bk ll lx lr) -> arbt n0 (del x (Bk ll lx lr))
n:nat
H0:True
H1:X.lt x x'
H5:rbt (S n) (Bk ll lx lr)
H8:rbt (S n) r
IHl':forall n0 : nat, rbt n0 (Bk ll lx lr) -> notblack (Bk ll lx lr) -> rbt n0 (del x (Bk ll lx lr))

arbt (S n) (lbalS (del x (Bk ll lx lr)) x' r)
apply lbalS_arb; auto.
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)
IHr:forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)
n:nat
H0:True
H1:X.lt x' x
H5:rbt n l
H8:rbt n r

arbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)
IHr:forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)
n:nat
H0:True
H1:X.lt x' x
H5:rbt n l
H8:rbt n r
IHr':forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)

arbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
l:tree
x':X.t
r:tree
x:X.t
IHr:forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)
n:nat
H0:True
H1:X.lt x' x
H5:rbt n l
H8:rbt n r
IHr':forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)

arbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
l:tree
x':X.t
rl:tree
rx:X.t
rr:tree
x:X.t
IHr:forall n0 : nat, rbt (S n0) (Bk rl rx rr) -> isblack (Bk rl rx rr) -> arbt n0 (del x (Bk rl rx rr))
n:nat
H0:True
H1:X.lt x' x
H5:rbt n l
H8:rbt n (Bk rl rx rr)
IHr':forall n0 : nat, rbt n0 (Bk rl rx rr) -> notblack (Bk rl rx rr) -> rbt n0 (del x (Bk rl rx rr))

arbt n (rbalS l x' (del x (Bk rl rx rr)))
l:tree
x':X.t
rl:tree
rx:X.t
rr:tree
x:X.t
IHr:forall n0 : nat, rbt (S n0) (Bk rl rx rr) -> isblack (Bk rl rx rr) -> arbt n0 (del x (Bk rl rx rr))
n:nat
H0:True
H1:X.lt x' x
H5:rbt (S n) l
H8:rbt (S n) (Bk rl rx rr)
IHr':forall n0 : nat, rbt n0 (Bk rl rx rr) -> notblack (Bk rl rx rr) -> rbt n0 (del x (Bk rl rx rr))

arbt (S n) (rbalS l x' (del x (Bk rl rx rr)))
apply rbalS_arb; auto.
del_arb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s0 -> isblack s0 -> arbt n0 (del x0 s0)
del_rb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt n0 s0 -> notblack s0 -> rbt n0 (del x0 s0)
s:tree
x:X.t
n:nat

rbt n s -> notblack s -> rbt n (del x s)
del_arb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s0 -> isblack s0 -> arbt n0 (del x0 s0)
del_rb:forall (s0 : tree) (x0 : X.t) (n0 : nat), rbt n0 s0 -> notblack s0 -> rbt n0 (del x0 s0)
s:tree
x:X.t
n:nat

rbt n s -> notblack s -> rbt n (del x s)
del_arb:forall (s0 : tree) (x0 : X.t) (n : nat), rbt (S n) s0 -> isblack s0 -> arbt n (del x0 s0)
del_rb:forall (s0 : tree) (x0 : X.t) (n : nat), rbt n s0 -> notblack s0 -> rbt n (del x0 s0)
s:tree
x:X.t

forall n : nat, rbt n s -> notblack s -> rbt n (del x s)
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)
IHr:forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)
n:nat
H0:True
H1:X.eq x x'
H7:notred l
H8:notred r
H9:rbt n l
H10:rbt n r

rbt n (append l r)
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)
IHr:forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)
n:nat
H0:True
H1:X.lt x x'
H7:notred l
H8:notred r
H9:rbt n l
H10:rbt n r
rbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)
IHr:forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)
n:nat
H0:True
H1:X.lt x' x
H7:notred l
H8:notred r
H9:rbt n l
H10:rbt n r
rbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)
IHr:forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)
n:nat
H0:True
H1:X.eq x x'
H7:notred l
H8:notred r
H9:rbt n l
H10:rbt n r

rbt n (append l r)
apply append_arb_rb; assumption.
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)
IHr:forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)
n:nat
H0:True
H1:X.lt x x'
H7:notred l
H8:notred r
H9:rbt n l
H10:rbt n r

rbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)
IHr:forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)
n:nat
H0:True
H1:X.lt x x'
H7:notred l
H8:notred r
H9:rbt n l
H10:rbt n r
IHl':forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)

rbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)
n:nat
H0:True
H1:X.lt x x'
H7:notred l
H8:notred r
H9:rbt n l
H10:rbt n r
IHl':forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)

rbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r end
ll:tree
lx:X.t
lr:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt n0 (Bk ll lx lr) -> notblack (Bk ll lx lr) -> rbt n0 (del x (Bk ll lx lr))
n:nat
H0:True
H1:X.lt x x'
H7:notred (Bk ll lx lr)
H8:notred r
H9:rbt n (Bk ll lx lr)
H10:rbt n r
IHl':forall n0 : nat, rbt (S n0) (Bk ll lx lr) -> isblack (Bk ll lx lr) -> arbt n0 (del x (Bk ll lx lr))

rbt n (lbalS (del x (Bk ll lx lr)) x' r)
ll:tree
lx:X.t
lr:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt n0 (Bk ll lx lr) -> notblack (Bk ll lx lr) -> rbt n0 (del x (Bk ll lx lr))
n:nat
H0:True
H1:X.lt x x'
H7:notred (Bk ll lx lr)
H8:notred r
H9:rbt (S n) (Bk ll lx lr)
H10:rbt (S n) r
IHl':forall n0 : nat, rbt (S n0) (Bk ll lx lr) -> isblack (Bk ll lx lr) -> arbt n0 (del x (Bk ll lx lr))

rbt (S n) (lbalS (del x (Bk ll lx lr)) x' r)
destruct n as [|n]; [invrb|]; apply lbalS_rb; auto.
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)
IHr:forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)
n:nat
H0:True
H1:X.lt x' x
H7:notred l
H8:notred r
H9:rbt n l
H10:rbt n r

rbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
del_arb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt (S n0) s -> isblack s -> arbt n0 (del x0 s)
del_rb:forall (s : tree) (x0 : X.t) (n0 : nat), rbt n0 s -> notblack s -> rbt n0 (del x0 s)
l:tree
x':X.t
r:tree
x:X.t
IHl:forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)
IHr:forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)
n:nat
H0:True
H1:X.lt x' x
H7:notred l
H8:notred r
H9:rbt n l
H10:rbt n r
IHr':forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)

rbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
l:tree
x':X.t
r:tree
x:X.t
IHr:forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)
n:nat
H0:True
H1:X.lt x' x
H7:notred l
H8:notred r
H9:rbt n l
H10:rbt n r
IHr':forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)

rbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) end
l:tree
x':X.t
rl:tree
rx:X.t
rr:tree
x:X.t
IHr:forall n0 : nat, rbt n0 (Bk rl rx rr) -> notblack (Bk rl rx rr) -> rbt n0 (del x (Bk rl rx rr))
n:nat
H0:True
H1:X.lt x' x
H7:notred l
H8:notred (Bk rl rx rr)
H9:rbt n l
H10:rbt n (Bk rl rx rr)
IHr':forall n0 : nat, rbt (S n0) (Bk rl rx rr) -> isblack (Bk rl rx rr) -> arbt n0 (del x (Bk rl rx rr))

rbt n (rbalS l x' (del x (Bk rl rx rr)))
l:tree
x':X.t
rl:tree
rx:X.t
rr:tree
x:X.t
IHr:forall n0 : nat, rbt n0 (Bk rl rx rr) -> notblack (Bk rl rx rr) -> rbt n0 (del x (Bk rl rx rr))
n:nat
H0:True
H1:X.lt x' x
H7:notred l
H8:notred (Bk rl rx rr)
H9:rbt (S n) l
H10:rbt (S n) (Bk rl rx rr)
IHr':forall n0 : nat, rbt (S n0) (Bk rl rx rr) -> isblack (Bk rl rx rr) -> arbt n0 (del x (Bk rl rx rr))

rbt (S n) (rbalS l x' (del x (Bk rl rx rr)))
apply rbalS_rb; auto. } Qed.
s:tree
x:X.t

Rbt s -> Rbt (remove x s)
s:tree
x:X.t

Rbt s -> Rbt (remove x s)
s:tree
x:X.t
n:nat
H:rbt n s

Rbt (remove x s)
s:tree
x:X.t
n:nat
H:rbt n s

Rbt (makeBlack (del x s))
x:X.t
n:nat
H:rbt n Leaf

Rbt (makeBlack (del x Leaf))
l:tree
y:X.t
r:tree
x:X.t
n:nat
H:rbt n (Rd l y r)
Rbt (makeBlack (del x (Rd l y r)))
l:tree
y:X.t
r:tree
x:X.t
n:nat
H:rbt n (Bk l y r)
Rbt (makeBlack (del x (Bk l y r)))
x:X.t
n:nat
H:rbt n Leaf

Rbt (makeBlack (del x Leaf))
x:X.t
n:nat
H:rbt n Leaf

arbt n (del x Leaf)
auto.
l:tree
y:X.t
r:tree
x:X.t
n:nat
H:rbt n (Rd l y r)

Rbt (makeBlack (del x (Rd l y r)))
l:tree
y:X.t
r:tree
x:X.t
n:nat
H:rbt n (Rd l y r)

arbt n (del x (Rd l y r))
l:tree
y:X.t
r:tree
x:X.t
n:nat
H:rbt n (Rd l y r)

rbt n (del x (Rd l y r))
apply del_rb; simpl; auto.
l:tree
y:X.t
r:tree
x:X.t
n:nat
H:rbt n (Bk l y r)

Rbt (makeBlack (del x (Bk l y r)))
l:tree
y:X.t
r:tree
x:X.t
n:nat
H:rbt (S n) (Bk l y r)

Rbt (makeBlack (del x (Bk l y r)))
l:tree
y:X.t
r:tree
x:X.t
n:nat
H:rbt (S n) (Bk l y r)

arbt n (del x (Bk l y r))
apply del_arb; simpl; auto. Qed.

Treeify

Definition treeify_rb_invariant size depth (f:treeify_t) :=
 forall acc,
 size <= length acc ->
  rbt depth (fst (f acc)) /\
  size + length (snd (f acc)) = length acc.


treeify_rb_invariant 0 0 treeify_zero

treeify_rb_invariant 0 0 treeify_zero
intros acc _; simpl; auto. Qed.

treeify_rb_invariant 1 0 treeify_one

treeify_rb_invariant 1 0 treeify_one
intros [|x acc]; simpl; auto; inversion 1. Qed.
f, g:treeify_t
size1, size2, size, d:nat

treeify_rb_invariant size1 d f -> treeify_rb_invariant size2 d g -> size = S (size1 + size2) -> treeify_rb_invariant size (S d) (treeify_cont f g)
f, g:treeify_t
size1, size2, size, d:nat

treeify_rb_invariant size1 d f -> treeify_rb_invariant size2 d g -> size = S (size1 + size2) -> treeify_rb_invariant size (S d) (treeify_cont f g)
f, g:treeify_t
size1, size2, size, d:nat
Hf:treeify_rb_invariant size1 d f
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
acc:list elt
Hacc:size <= length acc

rbt (S d) (fst (treeify_cont f g acc)) /\ size + length (snd (treeify_cont f g acc)) = length acc
f, g:treeify_t
size1, size2, size, d:nat
Hf:treeify_rb_invariant size1 d f
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
acc:list elt
Hacc:size <= length acc

rbt (S d) (fst (let (l, l0) := f acc in match l0 with | nil => bogus | x :: acc0 => let (r, acc1) := g acc0 in (Bk l x r, acc1) end)) /\ size + length (snd (let (l, l0) := f acc in match l0 with | nil => bogus | x :: acc0 => let (r, acc1) := g acc0 in (Bk l x r, acc1) end)) = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
Hf:size1 <= length acc -> rbt d (fst (f acc)) /\ size1 + length (snd (f acc)) = length acc
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc

rbt (S d) (fst (let (l, l0) := f acc in match l0 with | nil => bogus | x :: acc0 => let (r, acc1) := g acc0 in (Bk l x r, acc1) end)) /\ size + length (snd (let (l, l0) := f acc in match l0 with | nil => bogus | x :: acc0 => let (r, acc1) := g acc0 in (Bk l x r, acc1) end)) = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
acc1:list elt
Hf:size1 <= length acc -> rbt d (fst (l, acc1)) /\ size1 + length (snd (l, acc1)) = length acc
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc

rbt (S d) (fst match acc1 with | nil => bogus | x :: acc0 => let (r, acc2) := g acc0 in (Bk l x r, acc2) end) /\ size + length (snd match acc1 with | nil => bogus | x :: acc0 => let (r, acc2) := g acc0 in (Bk l x r, acc2) end) = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
acc1:list elt
Hf:size1 <= length acc -> rbt d l /\ size1 + length acc1 = length acc
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc

rbt (S d) (fst match acc1 with | nil => bogus | x :: acc0 => let (r, acc2) := g acc0 in (Bk l x r, acc2) end) /\ size + length (snd match acc1 with | nil => bogus | x :: acc0 => let (r, acc2) := g acc0 in (Bk l x r, acc2) end) = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
acc1:list elt
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc

size1 <= length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
acc1:list elt
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + length acc1 = length acc
rbt (S d) (fst match acc1 with | nil => bogus | x :: acc0 => let (r, acc2) := g acc0 in (Bk l x r, acc2) end) /\ size + length (snd match acc1 with | nil => bogus | x :: acc0 => let (r, acc2) := g acc0 in (Bk l x r, acc2) end) = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
acc1:list elt
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc

size1 <= length acc
f, g:treeify_t
size1, size2, d:nat
acc:list elt
l:tree
acc1:list elt
Hg:treeify_rb_invariant size2 d g
Hacc:S (size1 + size2) <= length acc

size1 <= length acc
eauto with arith.
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
acc1:list elt
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + length acc1 = length acc

rbt (S d) (fst match acc1 with | nil => bogus | x :: acc0 => let (r, acc2) := g acc0 in (Bk l x r, acc2) end) /\ size + length (snd match acc1 with | nil => bogus | x :: acc0 => let (r, acc2) := g acc0 in (Bk l x r, acc2) end) = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + 0 = length acc

rbt (S d) Leaf /\ size + 0 = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc
rbt (S d) (fst (let (r, acc0) := g acc2 in (Bk l x r, acc0))) /\ size + length (snd (let (r, acc0) := g acc2 in (Bk l x r, acc0))) = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + 0 = length acc

rbt (S d) Leaf /\ size + 0 = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + 0 = length acc

False
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hf1:rbt d l
Hf2:size1 + 0 = length acc

size <= length acc -> False
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hf1:rbt d l
Hf2:size1 + 0 = length acc

length acc < size
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hf1:rbt d l
Hf2:size1 + 0 = length acc

size1 + 0 < S (size1 + size2)
auto with arith.
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
Hg:treeify_rb_invariant size2 d g
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc

rbt (S d) (fst (let (r, acc0) := g acc2 in (Bk l x r, acc0))) /\ size + length (snd (let (r, acc0) := g acc2 in (Bk l x r, acc0))) = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
Hg:size2 <= length acc2 -> rbt d (fst (g acc2)) /\ size2 + length (snd (g acc2)) = length acc2
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc

rbt (S d) (fst (let (r, acc0) := g acc2 in (Bk l x r, acc0))) /\ size + length (snd (let (r, acc0) := g acc2 in (Bk l x r, acc0))) = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
r:tree
acc3:list elt
Hg:size2 <= length acc2 -> rbt d (fst (r, acc3)) /\ size2 + length (snd (r, acc3)) = length acc2
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc

rbt (S d) (fst (Bk l x r, acc3)) /\ size + length (snd (Bk l x r, acc3)) = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
r:tree
acc3:list elt
Hg:size2 <= length acc2 -> rbt d r /\ size2 + length acc3 = length acc2
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc

rbt (S d) (Bk l x r) /\ size + length acc3 = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
r:tree
acc3:list elt
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc

size2 <= length acc2
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
r:tree
acc3:list elt
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc
Hg1:rbt d r
Hg2:size2 + length acc3 = length acc2
rbt (S d) (Bk l x r) /\ size + length acc3 = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
r:tree
acc3:list elt
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc

size2 <= length acc2
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
r:tree
acc3:list elt
H:size = S (size1 + size2)
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc

size <= length acc -> size2 <= length acc2
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
r:tree
acc3:list elt
H:size = S (size1 + size2)
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc

size1 + size2 <= size1 + length acc2 -> size2 <= length acc2
apply Nat.add_le_mono_l.
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
r:tree
acc3:list elt
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc
Hg1:rbt d r
Hg2:size2 + length acc3 = length acc2

rbt (S d) (Bk l x r) /\ size + length acc3 = length acc
f, g:treeify_t
size1, size2, size, d:nat
acc:list elt
l:tree
x:elt
acc2:list elt
r:tree
acc3:list elt
H:size = S (size1 + size2)
Hacc:size <= length acc
Hf1:rbt d l
Hf2:size1 + S (length acc2) = length acc
Hg1:rbt d r
Hg2:size2 + length acc3 = length acc2

size + length acc3 = length acc
now rewrite H, <- Hf2, <- Hg2, Nat.add_succ_r, Nat.add_assoc. Qed.
n:positive

exists d : nat, forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)
n:positive

exists d : nat, forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)
n:positive
d:nat
IHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)

exists d0 : nat, forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n~1)) d0 (treeify_aux b n~1)
n:positive
d:nat
IHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)
exists d0 : nat, forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n~0)) d0 (treeify_aux b n~0)

exists d : nat, forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat 1)) d (treeify_aux b 1)
n:positive
d:nat
IHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)

exists d0 : nat, forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n~1)) d0 (treeify_aux b n~1)
n:positive
d:nat
IHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)

forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n~1)) (S d) (treeify_aux b n~1)
n:positive
d:nat
IHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)
b:bool

treeify_rb_invariant (ifpred b (Pos.to_nat n~1)) (S d) (treeify_aux b n~1)
n:positive
d:nat
IHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)
b:bool

ifpred b (Pos.to_nat n~1) = S (Pos.to_nat n + ifpred b (Pos.to_nat n))
n:positive
d:nat
IHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)
b:bool

ifpred b (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred b (Pos.to_nat n))
n:positive
d:nat
IHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)
b:bool
H:0 < Pos.to_nat n

ifpred b (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred b (Pos.to_nat n))
n:positive
d:nat
IHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)
b:bool
H:Pos.to_nat n <> 0

ifpred b (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred b (Pos.to_nat n))
n:positive
d:nat
IHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)
H:Pos.to_nat n <> 0

Pos.to_nat n + Pos.to_nat n = S (Pos.to_nat n + Init.Nat.pred (Pos.to_nat n))
now rewrite <- Nat.add_succ_r, Nat.succ_pred; trivial.
n:positive
d:nat
IHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)

exists d0 : nat, forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n~0)) d0 (treeify_aux b n~0)
n:positive
d:nat
IHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)

forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n~0)) (S d) (treeify_aux b n~0)
n:positive
d:nat
IHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)
b:bool

treeify_rb_invariant (ifpred b (Pos.to_nat n~0)) (S d) (treeify_aux b n~0)
n:positive
d:nat
IHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)
b:bool

ifpred b (Pos.to_nat n~0) = S (ifpred b (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))
n:positive
d:nat
IHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)
b:bool

ifpred b (2 * Pos.to_nat n) = S (ifpred b (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))
n:positive
d:nat
IHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)
b:bool
H:0 < Pos.to_nat n

ifpred b (2 * Pos.to_nat n) = S (ifpred b (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))
n:positive
d:nat
IHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)
b:bool
H:Pos.to_nat n <> 0

ifpred b (2 * Pos.to_nat n) = S (ifpred b (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))
n:positive
d:nat
IHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)
b:bool
H:Pos.to_nat n <> 0

ifpred b (2 * Pos.to_nat n) = ifpred b (Pos.to_nat n) + Pos.to_nat n
n:positive
d:nat
IHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)
H:Pos.to_nat n <> 0

Init.Nat.pred (Pos.to_nat n + Pos.to_nat n) = Init.Nat.pred (Pos.to_nat n) + Pos.to_nat n
n:positive
d:nat
IHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)
H:Pos.to_nat n <> 0

Init.Nat.pred (Pos.to_nat n) + Pos.to_nat n = Init.Nat.pred (Pos.to_nat n + Pos.to_nat n)
now apply Nat.add_pred_l.

exists d : nat, forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat 1)) d (treeify_aux b 1)
exists 0; destruct b; [ apply treeify_zero_rb | apply treeify_one_rb ]. Qed.
The black depth of treeify l is actually a log2, but we don't need to mention that.
l:list elt

Rbt (treeify l)
l:list elt

Rbt (treeify l)
l:list elt

Rbt (fst (treeify_aux true (plength l) l))
l:list elt
d:nat
H:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat (plength l))) d (treeify_aux b (plength l))

Rbt (fst (treeify_aux true (plength l) l))
l:list elt
d:nat
H:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat (plength l))) d (treeify_aux b (plength l))

rbt d (fst (treeify_aux true (plength l) l))
l:list elt
d:nat
H:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat (plength l))) d (treeify_aux b (plength l))

Init.Nat.pred (Pos.to_nat (plength l)) <= length l
now rewrite plength_spec. Qed.

Filtering

f:elt -> bool
s:t

Rbt (filter f s)
f:elt -> bool
s:t

Rbt (filter f s)
unfold filter; auto_tc. Qed.
f:elt -> bool
s:t

Rbt (fst (partition f s))
f:elt -> bool
s:t

Rbt (fst (partition f s))
f:elt -> bool
s:t

Rbt (fst (let (ok0, ko) := partition_aux f s nil nil in (treeify ok0, treeify ko)))
f:elt -> bool
s:t
l, l0:list X.t

Rbt (fst (treeify l, treeify l0))
f:elt -> bool
s:t
l, l0:list X.t

Rbt (treeify l)
auto_tc. Qed.
f:elt -> bool
s:t

Rbt (snd (partition f s))
f:elt -> bool
s:t

Rbt (snd (partition f s))
f:elt -> bool
s:t

Rbt (snd (let (ok0, ko) := partition_aux f s nil nil in (treeify ok0, treeify ko)))
f:elt -> bool
s:t
l, l0:list X.t

Rbt (snd (treeify l, treeify l0))
f:elt -> bool
s:t
l, l0:list X.t

Rbt (treeify l0)
auto_tc. Qed.

Union, intersection, difference

s1, s2:tree

Rbt s2 -> Rbt (fold add s1 s2)
s1, s2:tree

Rbt s2 -> Rbt (fold add s1 s2)
s1, s2:tree
H:Rbt s2

Rbt (fold add s1 s2)
s1, s2:tree
H:Rbt s2

Rbt (fold_right (fun (y : elt) (x : tree) => flip add x y) s2 (rev (elements s1)))
s1, s2:tree
H:Rbt s2

Rbt (fold_right (fun (y : X.t) (x : tree) => flip add x y) s2 (rev (elements s1)))
induction (rev (elements s1)); simpl; unfold flip in *; auto_tc. Qed.
s1, s2:tree

Rbt s2 -> Rbt (fold remove s1 s2)
s1, s2:tree

Rbt s2 -> Rbt (fold remove s1 s2)
s1, s2:tree
H:Rbt s2

Rbt (fold remove s1 s2)
s1, s2:tree
H:Rbt s2

Rbt (fold_right (fun (y : elt) (x : tree) => flip remove x y) s2 (rev (elements s1)))
s1, s2:tree
H:Rbt s2

Rbt (fold_right (fun (y : X.t) (x : tree) => flip remove x y) s2 (rev (elements s1)))
induction (rev (elements s1)); simpl; unfold flip in *; auto_tc. Qed.
s1, s2:tree

Rbt s1 -> Rbt s2 -> Rbt (union s1 s2)
s1, s2:tree

Rbt s1 -> Rbt s2 -> Rbt (union s1 s2)
s1, s2:tree
H:Rbt s1
H0:Rbt s2

Rbt (union s1 s2)
s1, s2:tree
H:Rbt s1
H0:Rbt s2

Rbt match compare_height s1 s1 s2 s2 with | Eq => treeify (union_list (rev_elements s1) (rev_elements s2) nil) | Lt => fold add s1 s2 | Gt => fold add s2 s1 end
destruct compare_height; auto_tc. Qed.
s1, s2:tree

Rbt s1 -> Rbt s2 -> Rbt (inter s1 s2)
s1, s2:tree

Rbt s1 -> Rbt s2 -> Rbt (inter s1 s2)
s1, s2:tree
H:Rbt s1
H0:Rbt s2

Rbt (inter s1 s2)
s1, s2:tree
H:Rbt s1
H0:Rbt s2

Rbt match compare_height s1 s1 s2 s2 with | Eq => treeify (inter_list (rev_elements s1) (rev_elements s2) nil) | Lt => filter (fun k : elt => mem k s2) s1 | Gt => filter (fun k : elt => mem k s1) s2 end
destruct compare_height; auto_tc. Qed.
s1, s2:tree

Rbt s1 -> Rbt s2 -> Rbt (diff s1 s2)
s1, s2:tree

Rbt s1 -> Rbt s2 -> Rbt (diff s1 s2)
s1, s2:tree
H:Rbt s1
H0:Rbt s2

Rbt (diff s1 s2)
s1, s2:tree
H:Rbt s1
H0:Rbt s2

Rbt match compare_height s1 s1 s2 s2 with | Eq => treeify (diff_list (rev_elements s1) (rev_elements s2) nil) | Lt => filter (fun k : elt => negb (mem k s2)) s1 | Gt => fold remove s2 s1 end
destruct compare_height; auto_tc. Qed. End BalanceProps.

Final Encapsulation

Now, in order to really provide a functor implementing S, we need to encapsulate everything into a type of binary search trees. They also happen to be well-balanced, but this has no influence on the correctness of operations, so we won't state this here, see BalanceProps if you need more than just the MSet interface.
Module Type MSetInterface_S_Ext := MSetInterface.S <+ MSetRemoveMin.

Module Make (X: Orders.OrderedType) <:
 MSetInterface_S_Ext with Module E := X.
 Module Raw. Include MakeRaw X. End Raw.
 Include MSetInterface.Raw2Sets X Raw.

 Definition opt_ok (x:option (elt * Raw.t)) :=
  match x with Some (_,s) => Raw.Ok s | None => True end.

 Definition mk_opt_t (x: option (elt * Raw.t))(P: opt_ok x) :
   option (elt * t) :=
 match x as o return opt_ok o -> option (elt * t) with
 | Some (k,s') => fun P : Raw.Ok s' => Some (k, Mkt s')
 | None => fun _ => None
 end P.

 Definition remove_min s : option (elt * t) :=
  mk_opt_t (Raw.remove_min (this s)) (Raw.remove_min_ok s).

 
s:t_
x:elt
s':t

remove_min s = Some (x, s') -> min_elt s = Some x /\ Equal (remove x s) s'
s:t_
x:elt
s':t

remove_min s = Some (x, s') -> min_elt s = Some x /\ Equal (remove x s) s'
s:Raw.t
Hs:Raw.Ok s
x:elt
s':t

remove_min {| this := s; is_ok := Hs |} = Some (x, s') -> min_elt {| this := s; is_ok := Hs |} = Some x /\ Equal (remove x {| this := s; is_ok := Hs |}) s'
s:Raw.t
Hs:Raw.Ok s
x:elt
s':t

match Raw.remove_min s as o return (opt_ok o -> option (elt * t)) with | Some p => let (k, s'0) as p0 return ((let (_, s0) := p0 in Raw.Ok s0) -> option (elt * t)) := p in fun P : Raw.Ok s'0 => Some (k, {| this := s'0; is_ok := P |}) | None => fun _ : True => None end (Raw.remove_min_ok s) = Some (x, s') -> Raw.min_elt s = Some x /\ (forall a : elt, Raw.In a (Raw.remove x s) <-> Raw.In a s')
s:Raw.t
Hs:Raw.Ok s
x:elt
s':t

(forall (x0 : Raw.elt) (s'0 : Raw.tree), Raw.remove_min s = Some (x0, s'0) -> Raw.min_elt s = Some x0 /\ Raw.remove x0 s = s'0) -> match Raw.remove_min s as o return (opt_ok o -> option (elt * t)) with | Some p => let (k, s'0) as p0 return ((let (_, s0) := p0 in Raw.Ok s0) -> option (elt * t)) := p in fun P : Raw.Ok s'0 => Some (k, {| this := s'0; is_ok := P |}) | None => fun _ : True => None end (Raw.remove_min_ok s) = Some (x, s') -> Raw.min_elt s = Some x /\ (forall a : elt, Raw.In a (Raw.remove x s) <-> Raw.In a s')
s:Raw.t
Hs:Raw.Ok s
x:elt
s':t
P:=Raw.remove_min_ok s:match Raw.remove_min s with | Some (_, s'0) => Raw.Ok s'0 | None => True end

(forall (x0 : Raw.elt) (s'0 : Raw.tree), Raw.remove_min s = Some (x0, s'0) -> Raw.min_elt s = Some x0 /\ Raw.remove x0 s = s'0) -> match Raw.remove_min s as o return (opt_ok o -> option (elt * t)) with | Some p => let (k, s'0) as p0 return ((let (_, s0) := p0 in Raw.Ok s0) -> option (elt * t)) := p in fun P0 : Raw.Ok s'0 => Some (k, {| this := s'0; is_ok := P0 |}) | None => fun _ : True => None end P = Some (x, s') -> Raw.min_elt s = Some x /\ (forall a : elt, Raw.In a (Raw.remove x s) <-> Raw.In a s')
s:Raw.t
Hs:Raw.Ok s
x:elt
s':t
P:match Raw.remove_min s with | Some (_, s'0) => Raw.Ok s'0 | None => True end

(forall (x0 : Raw.elt) (s'0 : Raw.tree), Raw.remove_min s = Some (x0, s'0) -> Raw.min_elt s = Some x0 /\ Raw.remove x0 s = s'0) -> match Raw.remove_min s as o return (opt_ok o -> option (elt * t)) with | Some p => let (k, s'0) as p0 return ((let (_, s0) := p0 in Raw.Ok s0) -> option (elt * t)) := p in fun P0 : Raw.Ok s'0 => Some (k, {| this := s'0; is_ok := P0 |}) | None => fun _ : True => None end P = Some (x, s') -> Raw.min_elt s = Some x /\ (forall a : elt, Raw.In a (Raw.remove x s) <-> Raw.In a s')
s:Raw.t
Hs:Raw.Ok s
x:elt
s':t
x0:Raw.elt
s0:Raw.tree
P:Raw.Ok s0

(forall (x1 : Raw.elt) (s'0 : Raw.tree), Some (x0, s0) = Some (x1, s'0) -> Raw.min_elt s = Some x1 /\ Raw.remove x1 s = s'0) -> Some (x0, {| this := s0; is_ok := P |}) = Some (x, s') -> Raw.min_elt s = Some x /\ (forall a : elt, Raw.In a (Raw.remove x s) <-> Raw.In a s')
s:Raw.t
Hs:Raw.Ok s
x:elt
s':t
x0:Raw.elt
s0:Raw.tree
P:Raw.Ok s0
H:forall (x1 : Raw.elt) (s'0 : Raw.tree), Some (x0, s0) = Some (x1, s'0) -> Raw.min_elt s = Some x1 /\ Raw.remove x1 s = s'0
U:Some (x0, {| this := s0; is_ok := P |}) = Some (x, s')

Raw.min_elt s = Some x /\ (forall a : elt, Raw.In a (Raw.remove x s) <-> Raw.In a s')
s:Raw.t
Hs:Raw.Ok s
x:elt
s0:Raw.tree
P:Raw.Ok s0
H:forall (x0 : Raw.elt) (s' : Raw.tree), Some (x, s0) = Some (x0, s') -> Raw.min_elt s = Some x0 /\ Raw.remove x0 s = s'

Raw.min_elt s = Some x /\ (forall a : elt, Raw.In a (Raw.remove x s) <-> Raw.In a {| this := s0; is_ok := P |})
s:Raw.t
Hs:Raw.Ok s
x:elt
s0:Raw.tree
P:Raw.Ok s0
H:forall (x0 : Raw.elt) (s' : Raw.tree), Some (x, s0) = Some (x0, s') -> Raw.min_elt s = Some x0 /\ Raw.remove x0 s = s'

Raw.min_elt s = Some x /\ (forall a : elt, Raw.In a (Raw.remove x s) <-> Raw.In a s0)
s:Raw.t
Hs:Raw.Ok s
x:elt
s0:Raw.tree
P:Raw.Ok s0
H:forall (x0 : Raw.elt) (s' : Raw.tree), Some (x, s0) = Some (x0, s') -> Raw.min_elt s = Some x0 /\ Raw.remove x0 s = s'
H0:Raw.min_elt s = Some x
H1:Raw.remove x s = s0

Raw.min_elt s = Some x /\ (forall a : elt, Raw.In a (Raw.remove x s) <-> Raw.In a s0)
subst; intuition. Qed.
s:t_

remove_min s = None -> Empty s
s:t_

remove_min s = None -> Empty s
s:Raw.t
Hs:Raw.Ok s

remove_min {| this := s; is_ok := Hs |} = None -> Empty {| this := s; is_ok := Hs |}
s:Raw.t
Hs:Raw.Ok s

match Raw.remove_min s as o return (opt_ok o -> option (elt * t)) with | Some p => let (k, s') as p0 return ((let (_, s0) := p0 in Raw.Ok s0) -> option (elt * t)) := p in fun P : Raw.Ok s' => Some (k, {| this := s'; is_ok := P |}) | None => fun _ : True => None end (Raw.remove_min_ok s) = None -> forall a : elt, ~ Raw.In a s
s:Raw.t
Hs:Raw.Ok s

(Raw.remove_min s = None -> Raw.Empty s) -> match Raw.remove_min s as o return (opt_ok o -> option (elt * t)) with | Some p => let (k, s') as p0 return ((let (_, s0) := p0 in Raw.Ok s0) -> option (elt * t)) := p in fun P : Raw.Ok s' => Some (k, {| this := s'; is_ok := P |}) | None => fun _ : True => None end (Raw.remove_min_ok s) = None -> forall a : elt, ~ Raw.In a s
s:Raw.t
Hs:Raw.Ok s
P:=Raw.remove_min_ok s:match Raw.remove_min s with | Some (_, s') => Raw.Ok s' | None => True end

(Raw.remove_min s = None -> Raw.Empty s) -> match Raw.remove_min s as o return (opt_ok o -> option (elt * t)) with | Some p => let (k, s') as p0 return ((let (_, s0) := p0 in Raw.Ok s0) -> option (elt * t)) := p in fun P0 : Raw.Ok s' => Some (k, {| this := s'; is_ok := P0 |}) | None => fun _ : True => None end P = None -> forall a : elt, ~ Raw.In a s
s:Raw.t
Hs:Raw.Ok s
P:match Raw.remove_min s with | Some (_, s') => Raw.Ok s' | None => True end

(Raw.remove_min s = None -> Raw.Empty s) -> match Raw.remove_min s as o return (opt_ok o -> option (elt * t)) with | Some p => let (k, s') as p0 return ((let (_, s0) := p0 in Raw.Ok s0) -> option (elt * t)) := p in fun P0 : Raw.Ok s' => Some (k, {| this := s'; is_ok := P0 |}) | None => fun _ : True => None end P = None -> forall a : elt, ~ Raw.In a s
destruct (Raw.remove_min s) as [(x0,s0)|]; now intuition. Qed. End Make.