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) *)
(************************************************************************)
Initial author: Andrew W. Appel, 2011.
Extra modifications by: Pierre Letouzey
The design decisions behind this implementation are described here:
Additional suggested reading:
- Efficient Verified Red-Black Trees, by Andrew W. Appel, September 2011. http://www.cs.princeton.edu/~appel/papers/redblack.pdf
- Red-Black Trees in a Functional Setting by Chris Okasaki.
Journal of Functional Programming, 9(4):471-477, July 1999.
http://www.eecs.usma.edu/webs/people/okasaki/jfp99redblack.pdf
- Red-black trees with types, by Stefan Kahrs.
Journal of Functional Programming, 11(4), 425-432, 2001.
- Functors for Proofs and Programs, by J.-C. Filliatre and P. Letouzey. ESOP'04: European Symposium on Programming, pp. 370-384, 2004. http://www.lri.fr/~filliatr/ftp/publis/fpp.ps.gz
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.
Module Ops (X:Orders.OrderedType) <: MSetInterface.Ops X.
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).
Definition singleton (k: elt) : tree := Bk Leaf k Leaf.
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.
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.
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).
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).
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
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).
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 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:
- Lt if height s2 is at least twice height s1;
- Gt if height s1 is at least twice height s2;
- Eq if heights are approximately equal.
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.
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.
x, y:eltInT y (singleton x) <-> X.eq y xunfold singleton; intuition_in. Qed.x, y:eltInT y (singleton x) <-> X.eq y xx:eltOk (singleton x)unfold singleton; auto. Qed.x:eltOk (singleton x)
s:treex:eltInT x (makeBlack s) <-> InT x sdestruct s; simpl; intuition_in. Qed.s:treex:eltInT x (makeBlack s) <-> InT x ss:treex:eltInT x (makeRed s) <-> InT x sdestruct s; simpl; intuition_in. Qed.s:treex:eltInT x (makeRed s) <-> InT x ss:treeH:Ok sOk (makeBlack s)destruct s; simpl; ok. Qed.s:treeH:Ok sOk (makeBlack s)s:treeH:Ok sOk (makeRed s)destruct s; simpl; ok. Qed.s:treeH:Ok sOk (makeRed s)
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:Typef:tree -> X.t -> tree -> Ag:tree -> At:treerspec 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:Typef:tree -> X.t -> tree -> Ag:tree -> At:treerspec f g t (rcase f g t)A:Typef:tree -> X.t -> tree -> X.t -> tree -> Ag:tree -> At:treerrspec f g t (rrcase f g t)A:Typef:tree -> X.t -> tree -> X.t -> tree -> Ag:tree -> At:treerrspec f g t (rrcase f g t)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:Typef:tree -> X.t -> tree -> X.t -> tree -> Ag:tree -> Al:treex:X.tr:treerrspec 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 endA:Typef:tree -> X.t -> tree -> X.t -> tree -> Ag:tree -> At:treerrspec f g t (rrcase' f g t)A:Typef:tree -> X.t -> tree -> X.t -> tree -> Ag:tree -> At:treerrspec f g t (rrcase' f g t)destruct l as [|[|] ll lx lr], r as [|[|] rl rx rr]; now constructor. Qed.A:Typef:tree -> X.t -> tree -> X.t -> tree -> Ag:tree -> Al:treex:X.tr:treerrspec 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
Balancing operations are instances of generic match
l:treek:X.tr:treerrspec (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:treek:X.tr:treerrspec (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:treek:X.tr:treerrspec (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:treek:X.tr:treerrspec (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:treek:X.tr:treerrspec (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:treek:X.tr:treerrspec (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:treex:X.tr:treerspec (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:treex:X.tr:treerspec (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:treex:X.tr:treerspec (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.l:treex:X.tr:treerspec (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:treex:X.tr:treey:eltInT y (lbal l x r) <-> X.eq y x \/ InT y l \/ InT y rcase lbal_match; intuition_in. Qed.l:treex:X.tr:treey:eltInT y (lbal l x r) <-> X.eq y x \/ InT y l \/ InT y rl:treex:X.tr:treeH:Ok lH0:Ok rlt_tree0:lt_tree x lgt_tree0:gt_tree x rOk (lbal l x r)destruct (lbal_match l x r); ok. Qed.l:treex:X.tr:treeH:Ok lH0:Ok rlt_tree0:lt_tree x lgt_tree0:gt_tree x rOk (lbal l x r)l:treex:X.tr:treey:eltInT y (rbal l x r) <-> X.eq y x \/ InT y l \/ InT y rcase rbal_match; intuition_in. Qed.l:treex:X.tr:treey:eltInT y (rbal l x r) <-> X.eq y x \/ InT y l \/ InT y rl:treex:X.tr:treeH:Ok lH0:Ok rlt_tree0:lt_tree x lgt_tree0:gt_tree x rOk (rbal l x r)destruct (rbal_match l x r); ok. Qed.l:treex:X.tr:treeH:Ok lH0:Ok rlt_tree0:lt_tree x lgt_tree0:gt_tree x rOk (rbal l x r)l:treex:X.tr:treey:eltInT y (rbal' l x r) <-> X.eq y x \/ InT y l \/ InT y rcase rbal'_match; intuition_in. Qed.l:treex:X.tr:treey:eltInT y (rbal' l x r) <-> X.eq y x \/ InT y l \/ InT y rl:treex:X.tr:treeH:Ok lH0:Ok rlt_tree0:lt_tree x lgt_tree0:gt_tree x rOk (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.l:treex:X.tr:treeH:Ok lH0:Ok rlt_tree0:lt_tree x lgt_tree0:gt_tree x rOk (rbal' l x r)
forall (s : tree) (x : X.t) (y : elt), InT y (ins x s) <-> X.eq y x \/ InT y sforall (s : tree) (x : X.t) (y : elt), InT y (ins x s) <-> X.eq y x \/ InT y sx:X.ty:eltInT y (Rd Leaf x Leaf) <-> X.eq y x \/ InT y Leafi:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 lIHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 rx:X.ty:eltH:X.eq x x'InT y (Node i l x' r) <-> X.eq y x \/ InT y (Node i l x' r)i:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 lIHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 rx:X.ty:eltH: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.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 lIHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 rx:X.ty:eltH:X.lt x' xInT 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)intuition_in.x:X.ty:eltInT y (Rd Leaf x Leaf) <-> X.eq y x \/ InT y Leafi:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 lIHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 rx:X.ty:eltH:X.eq x x'InT y (Node i l x' r) <-> X.eq y x \/ InT y (Node i l x' r)setoid_replace y with x; eauto.i:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 lIHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 rx:X.ty:eltH:X.eq x x'H1:X.eq y xInT y (Node i l x' r)descolor; autorew; rewrite IHl; intuition_in.i:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 lIHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 rx:X.ty:eltH: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 IHr; intuition_in. Qed. Hint Rewrite ins_spec : rb.i:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 l) <-> X.eq y0 x0 \/ InT y0 lIHr:forall (x0 : X.t) (y0 : elt), InT y0 (ins x0 r) <-> X.eq y0 x0 \/ InT y0 rx:X.ty:eltH:X.lt x' xInT 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)s:treex:X.tH:Ok sOk (ins x s)induct s x; auto; descolor; (apply lbal_ok || apply rbal_ok || ok); auto; intros y; autorew; intuition; order. Qed.s:treex:X.tH:Ok sOk (ins x s)s:treex:X.ty:eltInT y (add x s) <-> X.eq y x \/ InT y ss:treex:X.ty:eltInT y (add x s) <-> X.eq y x \/ InT y snow autorew. Qed. Hint Rewrite add_spec' : rb.s:treex:X.ty:eltInT y (makeBlack (ins x s)) <-> X.eq y x \/ InT y ss:treex:X.ty:eltH:Ok sInT y (add x s) <-> X.eq y x \/ InT y sapply add_spec'. Qed.s:treex:X.ty:eltH:Ok sInT y (add x s) <-> X.eq y x \/ InT y ss:treex:X.tH:Ok sOk (add x s)unfold add; auto_tc. Qed.s:treex:X.tH:Ok sOk (add x s)
l:treex:X.tr:treey:eltInT y (lbalS l x r) <-> X.eq y x \/ InT y l \/ InT y rl:treex:X.tr:treey:eltInT y (lbalS l x r) <-> X.eq y x \/ InT y l \/ InT y rl:treex:X.tr:treey:eltforall (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 rl:treex:X.tr:treey:eltforall 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 rintros; autorew; intuition_in.l:treex:X.tr:treey:eltforall (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 rl:treex:X.tr:treey:eltforall 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 rx:X.tr:treey:eltforall 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 rx:X.tr:treey:eltl:treeInT 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 rx:X.ty:eltl:treeInT y (Rd l x Leaf) <-> X.eq y x \/ InT y l \/ InT y Leafx:X.trl:treerx:X.trr:treey:eltl:treeInT 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.trl:treerx:X.trr:treey:eltl:treeInT y (rbal' l x (Rd rl rx rr)) <-> X.eq y x \/ InT y l \/ InT y (Bk rl rx rr)x:X.ty:eltl:treeInT y (Rd l x Leaf) <-> X.eq y x \/ InT y l \/ InT y Leafintuition_in.x:X.ty:eltl:treeInT y l \/ X.eq y x \/ False <-> X.eq y x \/ InT y l \/ Falsedestree rl; autorew; intuition_in.x:X.trl:treerx:X.trr:treey:eltl:treeInT 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.trl:treerx:X.trr:treey:eltl:treeInT y (rbal' l x (Rd rl rx rr)) <-> X.eq y x \/ InT y l \/ InT y (Bk rl rx rr)intuition_in. Qed.x:X.trl:treerx:X.trr:treey:eltl:treeX.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 rrl:treex:X.tr:treeOk l -> Ok r -> lt_tree x l -> gt_tree x r -> Ok (lbalS l x r)l:treex:X.tr:treeOk l -> Ok r -> lt_tree x l -> gt_tree x r -> Ok (lbalS l x r)l:treex:X.tr, a:treex0:X.tb:treeH:Ok (Rd a x0 b)H0:Ok rlt_tree0:lt_tree x (Rd a x0 b)gt_tree0:gt_tree x rOk (Rd (Bk a x0 b) x r)l:treex:X.tr, t0:treeH:notred t0H0:Ok t0H1:Ok rlt_tree0:lt_tree x t0gt_tree0:gt_tree x rOk 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 endok.l:treex:X.tr, a:treex0:X.tb:treeH:Ok (Rd a x0 b)H0:Ok rlt_tree0:lt_tree x (Rd a x0 b)gt_tree0:gt_tree x rOk (Rd (Bk a x0 b) x r)l:treex:X.tr, t0:treeH:notred t0H0:Ok t0H1:Ok rlt_tree0:lt_tree x t0gt_tree0:gt_tree x rOk 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 endl:treex:X.tt0:treeH:notred t0H0:Ok t0H1:Ok Leaflt_tree0:lt_tree x t0gt_tree0:gt_tree x LeafOk (Rd t0 x Leaf)l:treex:X.trl:treerx:X.trr, t0:treeH:notred t0H0:Ok t0H1:Ok (Rd rl rx rr)lt_tree0:lt_tree x t0gt_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) endl:treex:X.trl:treerx:X.trr, t0:treeH:notred t0H0:Ok t0H1:Ok (Bk rl rx rr)lt_tree0:lt_tree x t0gt_tree0:gt_tree x (Bk rl rx rr)Ok (rbal' t0 x (Rd rl rx rr))ok.l:treex:X.tt0:treeH:notred t0H0:Ok t0H1:Ok Leaflt_tree0:lt_tree x t0gt_tree0:gt_tree x LeafOk (Rd t0 x Leaf)l:treex:X.trl:treerx:X.trr, t0:treeH:notred t0H0:Ok t0H1:Ok (Rd rl rx rr)lt_tree0:lt_tree x t0gt_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) endl:treex:X.trll:treerlx:X.trlr:treerx:X.trr, t0:treeH:notred t0H0:Ok t0lt_tree0:lt_tree x t0gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)H7:Ok rrH8:lt_tree rx (Bk rll rlx rlr)H9:gt_tree rx rrH5:Ok rllH10:Ok rlrH11:lt_tree rlx rllH12:gt_tree rlx rlrOk (rbal' rlr rx (makeRed rr))l:treex:X.trll:treerlx:X.trlr:treerx:X.trr, t0:treeH:notred t0H0:Ok t0lt_tree0:lt_tree x t0gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)H7:Ok rrH8:lt_tree rx (Bk rll rlx rlr)H9:gt_tree rx rrH5:Ok rllH10:Ok rlrH11:lt_tree rlx rllH12:gt_tree rlx rlrgt_tree rlx (rbal' rlr rx (makeRed rr))l:treex:X.trll:treerlx:X.trlr:treerx:X.trr, t0:treeH:notred t0H0:Ok t0lt_tree0:lt_tree x t0gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)H7:Ok rrH8:lt_tree rx (Bk rll rlx rlr)H9:gt_tree rx rrH5:Ok rllH10:Ok rlrH11:lt_tree rlx rllH12:gt_tree rlx rlrOk (rbal' rlr rx (makeRed rr))intros w; autorew; auto.l:treex:X.trll:treerlx:X.trlr:treerx:X.trr, t0:treeH:notred t0H0:Ok t0lt_tree0:lt_tree x t0gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)H7:Ok rrH8:lt_tree rx (Bk rll rlx rlr)H9:gt_tree rx rrH5:Ok rllH10:Ok rlrH11:lt_tree rlx rllH12:gt_tree rlx rlrgt_tree rx (makeRed rr)l:treex:X.trll:treerlx:X.trlr:treerx:X.trr, t0:treeH:notred t0H0:Ok t0lt_tree0:lt_tree x t0gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)H7:Ok rrH8:lt_tree rx (Bk rll rlx rlr)H9:gt_tree rx rrH5:Ok rllH10:Ok rlrH11:lt_tree rlx rllH12:gt_tree rlx rlrgt_tree rlx (rbal' rlr rx (makeRed rr))destruct 1 as [Hw|[Hw|Hw]]; try rewrite Hw; eauto.l:treex:X.trll:treerlx:X.trlr:treerx:X.trr, t0:treeH:notred t0H0:Ok t0lt_tree0:lt_tree x t0gt_tree0:gt_tree x (Rd (Bk rll rlx rlr) rx rr)H7:Ok rrH8:lt_tree rx (Bk rll rlx rlr)H9:gt_tree rx rrH5:Ok rllH10:Ok rlrH11:lt_tree rlx rllH12:gt_tree rlx rlrw:eltX.eq w rx \/ InT w rlr \/ InT w rr -> X.lt rlx wl:treex:X.trl:treerx:X.trr, t0:treeH:notred t0H0:Ok t0H1:Ok (Bk rl rx rr)lt_tree0:lt_tree x t0gt_tree0:gt_tree x (Bk rl rx rr)Ok (rbal' t0 x (Rd rl rx rr))l:treex:X.trl:treerx:X.trr, t0:treeH:notred t0H0:Ok t0lt_tree0:lt_tree x t0gt_tree0:gt_tree x (Bk rl rx rr)H6:Ok rlH7:Ok rrH8:lt_tree rx rlH9:gt_tree rx rrOk (rbal' t0 x (Rd rl rx rr))apply rbal'_ok; ok. Qed.l:treex:X.trl:treerx:X.trr, t0:treeH:notred t0H0:Ok t0lt_tree0:lt_tree x t0gt_tree0:gt_tree x (Bk rl rx rr)H6:Ok rlH7:Ok rrH8:lt_tree rx rlH9:gt_tree rx rrOk (rbal' t0 x (Rd rl rx rr))l:treex:X.tr:treey:eltInT y (rbalS l x r) <-> X.eq y x \/ InT y l \/ InT y rl:treex:X.tr:treey:eltInT y (rbalS l x r) <-> X.eq y x \/ InT y l \/ InT y rl:treex:X.tr:treey:eltforall (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:treex:X.tr:treey:eltforall 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 t0intros; autorew; intuition_in.l:treex:X.tr:treey:eltforall (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:treex:X.tr:treey:eltforall 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 t0l:treex:X.tr:treey:eltt:treeInT 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 tx:X.tr:treey:eltt:treeInT y (Rd Leaf x t) <-> X.eq y x \/ InT y Leaf \/ InT y tll:treelx:X.tlr:treex:X.tr:treey:eltt:treeInT 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 tll:treelx:X.tlr:treex:X.tr:treey:eltt:treeInT y (lbal (Rd ll lx lr) x t) <-> X.eq y x \/ InT y (Bk ll lx lr) \/ InT y tx:X.tr:treey:eltt:treeInT y (Rd Leaf x t) <-> X.eq y x \/ InT y Leaf \/ InT y tintuition_in.x:X.tr:treey:eltt:treeFalse \/ X.eq y x \/ InT y t <-> X.eq y x \/ False \/ InT y tdestruct lr as [|[|] lrl lrx lrr]; autorew; intuition_in.ll:treelx:X.tlr:treex:X.tr:treey:eltt:treeInT 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 tll:treelx:X.tlr:treex:X.tr:treey:eltt:treeInT y (lbal (Rd ll lx lr) x t) <-> X.eq y x \/ InT y (Bk ll lx lr) \/ InT y tintuition_in. Qed.ll:treelx:X.tlr:treex:X.tr:treey:eltt:treeX.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 tl:treex:X.tr:treeOk l -> Ok r -> lt_tree x l -> gt_tree x r -> Ok (rbalS l x r)l:treex:X.tr:treeOk l -> Ok r -> lt_tree x l -> gt_tree x r -> Ok (rbalS l x r)l:treex:X.tr, a:treex0:X.tb:treeH:Ok lH0:Ok (Rd a x0 b)lt_tree0:lt_tree x lgt_tree0:gt_tree x (Rd a x0 b)Ok (Rd l x (Bk a x0 b))l:treex:X.tr, t0:treeH:notred t0H0:Ok lH1:Ok t0lt_tree0:lt_tree x lgt_tree0:gt_tree x t0Ok 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 endok.l:treex:X.tr, a:treex0:X.tb:treeH:Ok lH0:Ok (Rd a x0 b)lt_tree0:lt_tree x lgt_tree0:gt_tree x (Rd a x0 b)Ok (Rd l x (Bk a x0 b))l:treex:X.tr, t0:treeH:notred t0H0:Ok lH1:Ok t0lt_tree0:lt_tree x lgt_tree0:gt_tree x t0Ok 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 endx:X.tr, t0:treeH:notred t0H0:Ok LeafH1:Ok t0lt_tree0:lt_tree x Leafgt_tree0:gt_tree x t0Ok (Rd Leaf x t0)ll:treelx:X.tlr:treex:X.tr, t0:treeH:notred t0H0:Ok (Rd ll lx lr)H1:Ok t0lt_tree0:lt_tree x (Rd ll lx lr)gt_tree0:gt_tree x t0Ok 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 endll:treelx:X.tlr:treex:X.tr, t0:treeH:notred t0H0:Ok (Bk ll lx lr)H1:Ok t0lt_tree0:lt_tree x (Bk ll lx lr)gt_tree0:gt_tree x t0Ok (lbal (Rd ll lx lr) x t0)ok.x:X.tr, t0:treeH:notred t0H0:Ok LeafH1:Ok t0lt_tree0:lt_tree x Leafgt_tree0:gt_tree x t0Ok (Rd Leaf x t0)ll:treelx:X.tlr:treex:X.tr, t0:treeH:notred t0H0:Ok (Rd ll lx lr)H1:Ok t0lt_tree0:lt_tree x (Rd ll lx lr)gt_tree0:gt_tree x t0Ok 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 endll:treelx:X.tlrl:treelrx:X.tlrr:treex:X.tr, t0:treeH:notred t0H1:Ok t0lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))gt_tree0:gt_tree x t0H6:Ok llH8:lt_tree lx llH9:gt_tree lx (Bk lrl lrx lrr)H5:Ok lrlH10:Ok lrrH11:lt_tree lrx lrlH12:gt_tree lrx lrrOk (lbal (makeRed ll) lx lrl)ll:treelx:X.tlrl:treelrx:X.tlrr:treex:X.tr, t0:treeH:notred t0H1:Ok t0lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))gt_tree0:gt_tree x t0H6:Ok llH8:lt_tree lx llH9:gt_tree lx (Bk lrl lrx lrr)H5:Ok lrlH10:Ok lrrH11:lt_tree lrx lrlH12:gt_tree lrx lrrlt_tree lrx (lbal (makeRed ll) lx lrl)ll:treelx:X.tlrl:treelrx:X.tlrr:treex:X.tr, t0:treeH:notred t0H1:Ok t0lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))gt_tree0:gt_tree x t0H6:Ok llH8:lt_tree lx llH9:gt_tree lx (Bk lrl lrx lrr)H5:Ok lrlH10:Ok lrrH11:lt_tree lrx lrlH12:gt_tree lrx lrrOk (lbal (makeRed ll) lx lrl)intros w; autorew; auto.ll:treelx:X.tlrl:treelrx:X.tlrr:treex:X.tr, t0:treeH:notred t0H1:Ok t0lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))gt_tree0:gt_tree x t0H6:Ok llH8:lt_tree lx llH9:gt_tree lx (Bk lrl lrx lrr)H5:Ok lrlH10:Ok lrrH11:lt_tree lrx lrlH12:gt_tree lrx lrrlt_tree lx (makeRed ll)ll:treelx:X.tlrl:treelrx:X.tlrr:treex:X.tr, t0:treeH:notred t0H1:Ok t0lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))gt_tree0:gt_tree x t0H6:Ok llH8:lt_tree lx llH9:gt_tree lx (Bk lrl lrx lrr)H5:Ok lrlH10:Ok lrrH11:lt_tree lrx lrlH12:gt_tree lrx lrrlt_tree lrx (lbal (makeRed ll) lx lrl)destruct 1 as [Hw|[Hw|Hw]]; try rewrite Hw; eauto.ll:treelx:X.tlrl:treelrx:X.tlrr:treex:X.tr, t0:treeH:notred t0H1:Ok t0lt_tree0:lt_tree x (Rd ll lx (Bk lrl lrx lrr))gt_tree0:gt_tree x t0H6:Ok llH8:lt_tree lx llH9:gt_tree lx (Bk lrl lrx lrr)H5:Ok lrlH10:Ok lrrH11:lt_tree lrx lrlH12:gt_tree lrx lrrw:eltX.eq w lx \/ InT w ll \/ InT w lrl -> X.lt w lrxll:treelx:X.tlr:treex:X.tr, t0:treeH:notred t0H0:Ok (Bk ll lx lr)H1:Ok t0lt_tree0:lt_tree x (Bk ll lx lr)gt_tree0:gt_tree x t0Ok (lbal (Rd ll lx lr) x t0)apply lbal_ok; ok. Qed. Hint Rewrite lbalS_spec rbalS_spec : rb.ll:treelx:X.tlr:treex:X.tr, t0:treeH:notred t0H1:Ok t0lt_tree0:lt_tree x (Bk ll lx lr)gt_tree0:gt_tree x t0H6:Ok llH7:Ok lrH8:lt_tree lx llH9:gt_tree lx lrOk (lbal (Rd ll lx lr) x t0)
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:treelx:X.tlr, rl:treerx:X.trr:treerspec (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:treelx:X.tlr, rl:treerx:X.trr:treerspec (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:treelx:X.tlr, rl:treerx:X.trr:treerspec (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.ll:treelx:X.tlr, rl:treerx:X.trr:treerspec (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))l, r:treex:eltInT x (append l r) <-> InT x l \/ InT x rl, r:treex:eltInT x (append l r) <-> InT x l \/ InT x rl:treex:eltforall r : tree, InT x (append l r) <-> InT x l \/ InT x rll:treelx:X.tlr:treex:eltrl:treeIHlr:InT x (append lr rl) <-> InT x lr \/ InT x rlrx:X.trr:treeInT 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 rrll:treelx:X.tlr:treex:eltrl:treeIHlr:InT x (append lr rl) <-> InT x lr \/ InT x rlrx:X.trr:treeInT 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 rrrevert IHlr; case append_rr_match; [intros a y b | intros t Ht]; autorew; tauto.ll:treelx:X.tlr:treex:eltrl:treeIHlr:InT x (append lr rl) <-> InT x lr \/ InT x rlrx:X.trr:treeInT 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 rrrevert IHlr; case append_bb_match; [intros a y b | intros t Ht]; autorew; tauto. Qed. Hint Rewrite append_spec : rb.ll:treelx:X.tlr:treex:eltrl:treeIHlr:InT x (append lr rl) <-> InT x lr \/ InT x rlrx:X.trr:treeInT 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 rrforall (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.tr:treeOk Leaf -> Ok r -> lt_tree x Leaf -> gt_tree x r -> Ok rx:X.tlc:Color.tll:treelx:X.tlr:treeIHlr: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.tll:treelx:X.tlr, rl:treeIHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)rx:X.trr:treeOk (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.tll:treelx:X.tlr, r:treeIHlr:Ok lr -> Ok r -> lt_tree x lr -> gt_tree x r -> Ok (append lr r)Hr:notred rOk (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.trl:treerx:X.trr, l:treeIHrl:Ok l -> Ok rl -> lt_tree x l -> gt_tree x rl -> Ok (append l rl)Hl:notred lOk 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.tll:treelx:X.tlr, rl:treeIHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)rx:X.trr:treeOk (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))trivial.x:X.tr:treeOk Leaf -> Ok r -> lt_tree x Leaf -> gt_tree x r -> Ok rtrivial.x:X.tlc:Color.tll:treelx:X.tlr:treeIHlr: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.tll:treelx:X.tlr, rl:treeIHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)rx:X.trr:treeOk (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.tll:treelx:X.tlr, rl:treeIHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)rx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrOk (append (Rd ll lx lr) (Rd rl rx rr))x:X.tll:treelx:X.tlr, rl:treeIHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)rx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)Ok (append (Rd ll lx lr) (Rd rl rx rr))x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)Ok (append (Rd ll lx lr) (Rd rl rx rr))x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxOk (append (Rd ll lx lr) (Rd rl rx rr))x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxgt_tree lx (append lr rl)x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)Ok (append (Rd ll lx lr) (Rd rl rx rr))x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxgt_tree lx (append lr rl)x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxw:eltInT w (append lr rl) -> X.lt lx wdestruct 1; [|transitivity x]; eauto.x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxw:eltInT w lr \/ InT w rl -> X.lt lx wx:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)Ok (append (Rd ll lx lr) (Rd rl rx rr))x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)lt_tree rx (append lr rl)x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG: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.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)lt_tree rx (append lr rl)x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)w:eltInT w (append lr rl) -> X.lt w rxdestruct 1; [transitivity x|]; eauto.x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)w:eltInT w lr \/ InT w rl -> X.lt w rxrevert IH G L; case append_rr_match; intros; ok.x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Rd ll lx lr)H2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG: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.tll:treelx:X.tlr, r:treeIHlr:Ok lr -> Ok r -> lt_tree x lr -> gt_tree x r -> Ok (append lr r)Hr:notred rOk (Rd ll lx lr) -> Ok r -> lt_tree x (Rd ll lx lr) -> gt_tree x r -> Ok (Rd ll lx (append lr r))intros w; autorew; destruct 1; eauto.x:X.tll:treelx:X.tlr, r:treeIHlr:Ok lr -> Ok r -> lt_tree x lr -> gt_tree x r -> Ok (append lr r)Hr:notred rH0:Ok rH1:lt_tree x (Rd ll lx lr)H2:gt_tree x rH7:Ok llH8:Ok lrH9:lt_tree lx llH10:gt_tree lx lrgt_tree lx (append lr r)x:X.trl:treerx:X.trr, l:treeIHrl:Ok l -> Ok rl -> lt_tree x l -> gt_tree x rl -> Ok (append l rl)Hl:notred lOk l -> Ok (Rd rl rx rr) -> lt_tree x l -> gt_tree x (Rd rl rx rr) -> Ok (Rd (append l rl) rx rr)intros w; autorew; destruct 1; eauto.x:X.trl:treerx:X.trr, l:treeIHrl:Ok l -> Ok rl -> lt_tree x l -> gt_tree x rl -> Ok (append l rl)Hl:notred lH:Ok lH1:lt_tree x lH2:gt_tree x (Rd rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrlt_tree rx (append l rl)x:X.tll:treelx:X.tlr, rl:treeIHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)rx:X.trr:treeOk (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.tll:treelx:X.tlr, rl:treeIHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)rx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrOk (append (Bk ll lx lr) (Bk rl rx rr))x:X.tll:treelx:X.tlr, rl:treeIHlr:Ok lr -> Ok rl -> lt_tree x lr -> gt_tree x rl -> Ok (append lr rl)rx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)Ok (append (Bk ll lx lr) (Bk rl rx rr))x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)Ok (append (Bk ll lx lr) (Bk rl rx rr))x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxOk (append (Bk ll lx lr) (Bk rl rx rr))x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxgt_tree lx (append lr rl)x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)Ok (append (Bk ll lx lr) (Bk rl rx rr))x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxgt_tree lx (append lr rl)x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxw:eltInT w (append lr rl) -> X.lt lx wdestruct 1; [|transitivity x]; eauto.x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxw:eltInT w lr \/ InT w rl -> X.lt lx wx:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)Ok (append (Bk ll lx lr) (Bk rl rx rr))x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)lt_tree rx (append lr rl)x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG: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.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)lt_tree rx (append lr rl)x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)w:eltInT w (append lr rl) -> X.lt w rxdestruct 1; [transitivity x|]; eauto.x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)w:eltInT w lr \/ InT w rl -> X.lt w rxx:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrIH:Ok (append lr rl)H:X.lt lx rxG:gt_tree lx (append lr rl)L:lt_tree rx (append lr rl)Ok (append (Bk ll lx lr) (Bk rl rx rr))apply lbalS_ok; ok. Qed.x:X.tll:treelx:X.tlr, rl:treerx:X.trr:treeH1:lt_tree x (Bk ll lx lr)H2:gt_tree x (Bk rl rx rr)H7:Ok rlH8:Ok rrH9:lt_tree rx rlH10:gt_tree rx rrH6:Ok llH11:Ok lrH12:lt_tree lx llH13:gt_tree lx lrH:X.lt lx rxt0:treeH0:notred t0IH:Ok t0G:gt_tree lx t0L:lt_tree rx t0Ok (lbalS ll lx (Bk t0 rx rr))
forall (s : tree) (x : X.t) (y : elt), Ok s -> InT y (del x s) <-> InT y s /\ ~ X.eq y xforall (s : tree) (x : X.t) (y : elt), Ok s -> InT y (del x s) <-> InT y s /\ ~ X.eq y xx:X.ty:eltH:Ok LeafInT y Leaf <-> InT y Leaf /\ ~ X.eq y xi:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0x:X.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rInT y (append l r) <-> InT y (Node i l x' r) /\ ~ X.eq y xi:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0x:X.ty:eltH0:X.lt x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rInT 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 xi:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0x:X.ty:eltH0:X.lt x' xH5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rInT 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 xintuition_in.x:X.ty:eltH:Ok LeafInT y Leaf <-> InT y Leaf /\ ~ X.eq y xi:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0x:X.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rInT y (append l r) <-> InT y (Node i l x' r) /\ ~ X.eq y xi:Color.tl:treex':X.tr:treeIHl: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.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH1:InT y lH:X.eq y xFalsei:Color.tl:treex':X.tr:treeIHl: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.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH1:InT y rH:X.eq y xFalsei:Color.tl:treex':X.tr:treeIHl: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.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH2:X.eq y x -> FalseH1:X.eq y x'InT y l \/ InT y ri:Color.tl:treex':X.tr:treeIHl: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.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH1:InT y lH:X.eq y xH2:X.lt y x'Falsei:Color.tl:treex':X.tr:treeIHl: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.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH1:InT y rH:X.eq y xFalsei:Color.tl:treex':X.tr:treeIHl: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.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH2:X.eq y x -> FalseH1:X.eq y x'InT y l \/ InT y ri:Color.tl:treex':X.tr:treeIHl: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.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH1:InT y rH:X.eq y xFalsei:Color.tl:treex':X.tr:treeIHl: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.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH2:X.eq y x -> FalseH1:X.eq y x'InT y l \/ InT y ri:Color.tl:treex':X.tr:treeIHl: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.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH1:InT y rH:X.eq y xH2:X.lt x' yFalsei:Color.tl:treex':X.tr:treeIHl: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.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH2:X.eq y x -> FalseH1:X.eq y x'InT y l \/ InT y rorder.i:Color.tl:treex':X.tr:treeIHl: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.ty:eltH0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH2:X.eq y x -> FalseH1:X.eq y x'InT y l \/ InT y rdestruct l as [|[|] ll lx lr]; autorew; rewrite ?IHl by trivial; intuition_in; order.i:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0x:X.ty:eltH0:X.lt x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rInT 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 xdestruct r as [|[|] rl rx rr]; autorew; rewrite ?IHr by trivial; intuition_in; order. Qed. Hint Rewrite del_spec : rb.i:Color.tl:treex':X.tr:treeIHl:forall (x0 : X.t) (y0 : elt), Ok l -> InT y0 (del x0 l) <-> InT y0 l /\ ~ X.eq y0 x0IHr:forall (x0 : X.t) (y0 : elt), Ok r -> InT y0 (del x0 r) <-> InT y0 r /\ ~ X.eq y0 x0x:X.ty:eltH0:X.lt x' xH5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rInT 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 xs:treex:X.tH:Ok sOk (del x s)s:treex:X.tH:Ok sOk (del x s)x:X.tH:Ok LeafOk Leafi:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rOk (append l r)i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rOk match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r endi:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x' xH5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rOk match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) endtrivial.x:X.tH:Ok LeafOk Leafeapply append_ok; eauto.i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.eq x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rOk (append l r)i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rOk match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r endi:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rlt_tree x' (del x l)i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH:lt_tree x' (del x l)Ok match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r endi:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rlt_tree x' (del x l)i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rw:eltInT w (del x l) -> X.lt w x'i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rw:eltInT w l /\ ~ X.eq w x -> X.lt w x'eauto.i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rw:eltH:InT w lH1:~ X.eq w xX.lt w x'destruct l as [|[|] ll lx lr]; auto_tc.i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x x'H5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH:lt_tree x' (del x l)Ok match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r endi:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x' xH5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rOk match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) endi:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x' xH5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rgt_tree x' (del x r)i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x' xH5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH:gt_tree x' (del x r)Ok match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) endi:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x' xH5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rgt_tree x' (del x r)i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x' xH5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rw:eltInT w (del x r) -> X.lt x' wi:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x' xH5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rw:eltInT w r /\ ~ X.eq w x -> X.lt x' weauto.i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x' xH5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rw:eltH:InT w rH1:~ X.eq w xX.lt x' wdestruct r as [|[|] rl rx rr]; auto_tc. Qed.i:Color.tl:treex':X.tr:treex:X.tIHl:Ok l -> Ok (del x l)IHr:Ok r -> Ok (del x r)H0:X.lt x' xH5:Ok lH6:Ok rH7:lt_tree x' lH8:gt_tree x' rH:gt_tree x' (del x r)Ok match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) ends:treex:X.ty:eltH:Ok sInT y (remove x s) <-> InT y s /\ ~ X.eq y xs:treex:X.ty:eltH:Ok sInT y (remove x s) <-> InT y s /\ ~ X.eq y xnow autorew. Qed. Hint Rewrite remove_spec : rb.s:treex:X.ty:eltH:Ok sInT y (makeBlack (del x s)) <-> InT y s /\ ~ X.eq y xs:treex:X.tH:Ok sOk (remove x s)unfold remove; auto_tc. Qed.s:treex:X.tH:Ok sOk (remove x s)
l:treey:X.tr:treec:Color.tx:elts':treeO: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:treey:X.tr:treec:Color.tx:elts':treeO: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:treeforall (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.tll:treely:X.tlr:treeIH: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.tr:treex:elts':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'now rewrite MX.compare_refl.x:elts':treeH:(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'lc:Color.tll:treely:X.tlr:treeIH: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.tll:treely:X.tlr:treeIH: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'0y:X.tr:treec:Color.tx:elts':treeO: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.tll:treely:X.tlr:treeIH: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'0y:X.tr:treec:Color.tx:elts':treeO: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.tll:treely:X.tlr:treeIH: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'0y:X.tr:treec:Color.tx:elts':treeO: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.tll:treely:X.tlr:treex0:elts0:treeIH: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'0y:X.tr:treec:Color.tx:elts':treeO: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.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts':treeO:Ok (Node c (Node lc ll ly lr) y r)H:min_elt (Node lc ll ly lr) = Some x0H0:del x0 (Node lc ll ly lr) = s0match 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.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some x0H0:del x0 l = s0match 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.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some x0H0:del x0 l = s0match 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.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some x0H0:del x0 l = s0E: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.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some xH0:del x l = s0E: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.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some xH0:del x l = s0E: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 xlc:Color.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some xH0:del x l = s0E: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'subst l; intuition.lc:Color.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some xH0:del x l = s0E: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 xlc:Color.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some xH0:del x l = s0E: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.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some xH0:del x l = s0E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')X.lt x ylc:Color.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some xH0:del x l = s0E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')H1:X.lt x ydel x (Node c l y r) = s'lc:Color.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some xH0:del x l = s0E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')X.lt x ylc:Color.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrH:min_elt l = Some xH0:del x l = s0E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')H1:bst lH2:bst rH3:lt_tree y lH4:gt_tree y rX.lt x yauto.lc:Color.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrH:min_elt l = Some xH0:del x l = s0E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')H1:bst lH2:bst rH3:lt_tree y lH4:gt_tree y rH5:InT x lX.lt x ylc:Color.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some xH0:del x l = s0E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')H1:X.lt x ydel x (Node c l y r) = s'lc:Color.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some xH0:del x l = s0E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')H1:X.lt x ymatch 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'destruct lc; injection E; subst l s0; auto. Qed.lc:Color.tll:treely:X.tlr:treex0:elts0:treey:X.tr:treec:Color.tx:elts', l:treeHeql:l = Node lc ll ly lrO:Ok (Node c l y r)H:min_elt l = Some xH0:del x l = s0E:match lc with | Red => (x, Rd s0 y r) | Black => (x, lbalS s0 y r) end = (x, s')H1:X.lt x yX.lt x y -> match l with | Bk _ _ _ => lbalS (del x l) y r | _ => Rd (del x l) y r end = s's:treex:elts':treeH:Ok sremove_min s = Some (x, s') -> min_elt s = Some x /\ remove x s = s's:treex:elts':treeH:Ok sremove_min s = Some (x, s') -> min_elt s = Some x /\ remove x s = s's:treex:elts':treeH:Ok smatch 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.tl:treey:X.tr:treex:elts':treeH: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.tl:treey:X.tr:treex:elts':treeH: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.tl:treey:X.tr:treex:elts':treeH:Ok (Node c l y r)x0:elts0: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.tl:treey:X.tr:treex:elts':treeH:Ok (Node c l y r)x0:elts0:treeD: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'0Some (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.tl:treey:X.tr:treex:elts':treeH:Ok (Node c l y r)x0:eltD: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'0Some (x0, makeBlack (del 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.c:Color.tl:treey:X.tr:treex:elts':treeH:Ok (Node c l y r)x0:eltD: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'0Some (x0, remove x0 (Node c l y r)) = Some (x, s') -> Some x0 = Some x /\ remove x (Node c l y r) = s's:treeremove_min s = None -> Empty ss:treeremove_min s = None -> Empty ss:treematch s with | Leaf => None | Node _ l x r => let (k, t) := delmin l x r in Some (k, makeBlack t) end = None -> Empty sNone = None -> Empty Leafc:Color.tl:treey:X.tr:tree(let (k, t) := delmin l y r in Some (k, makeBlack t)) = None -> Empty (Node c l y r)easy.None = None -> Empty Leafnow destruct delmin. Qed.c:Color.tl:treey:X.tr:tree(let (k, t) := delmin l y r in Some (k, makeBlack t)) = None -> Empty (Node c l y r)s:tH:Ok smatch remove_min s with | Some (_, s') => Ok s' | None => True ends:tH:Ok smatch remove_min s with | Some (_, s') => Ok s' | None => True ends:tH: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 ends:tH:Ok sx0:elts0:tree(forall (x : elt) (s' : tree), Ok s -> Some (x0, s0) = Some (x, s') -> min_elt s = Some x /\ remove x s = s') -> Ok s0s:tH:Ok sx0:elts0:treeR:forall (x : elt) (s' : tree), Ok s -> Some (x0, s0) = Some (x, s') -> min_elt s = Some x /\ remove x s = s'Ok s0s:tH:Ok sx0:elts0:treeR: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 x0H1:remove x0 s = s0Ok s0auto_tc. Qed.s:tH:Ok sx0:eltR: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 x0Ok (remove x0 s)
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_zerotreeify_invariant 0 treeify_zeroacc:list elt0 <= length acc -> let (t0, acc') := treeify_zero acc in cardinal t0 = 0 /\ acc = elements t0 ++ acc'auto. Qed.acc:list elt0 <= length acc -> 0 = 0 /\ acc = acctreeify_invariant 1 treeify_oneintros [|x acc]; simpl; auto; inversion 1. Qed.treeify_invariant 1 treeify_onef, g:treeify_tsize1, size2, size:nattreeify_invariant size1 f -> treeify_invariant size2 g -> size = S (size1 + size2) -> treeify_invariant size (treeify_cont f g)f, g:treeify_tsize1, size2, size:nattreeify_invariant size1 f -> treeify_invariant size2 g -> size = S (size1 + size2) -> treeify_invariant size (treeify_cont f g)f, g:treeify_tsize1, size2, size:natHf:treeify_invariant size1 fHg:treeify_invariant size2 gEQ:size = S (size1 + size2)acc:list eltLE:size <= length acclet (t0, acc') := treeify_cont f g acc in cardinal t0 = size /\ acc = elements t0 ++ acc'f, g:treeify_tsize1, size2, size:natHf:treeify_invariant size1 fHg:treeify_invariant size2 gEQ:size = S (size1 + size2)acc:list eltLE:size <= length acclet (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_tsize1, size2, size:natacc:list eltHf:size1 <= length acc -> let (t0, acc') := f acc in cardinal t0 = size1 /\ acc = elements t0 ++ acc'Hg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length acclet (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_tsize1, size2, size:natacc:list eltt1:treeacc1:list eltHf:size1 <= length acc -> cardinal t1 = size1 /\ acc = elements t1 ++ acc1Hg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length acclet (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_tsize1, size2, size:natacc:list eltt1:treeacc1:list eltHg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length accsize1 <= length accf, g:treeify_tsize1, size2, size:natacc:list eltt1:treeacc1:list eltHg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ acc1let (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_tsize1, size2, size:natacc:list eltt1:treeacc1:list eltHg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length accsize1 <= length accf, g:treeify_tsize1, size2, size:natacc:list eltt1:treeacc1:list eltHg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length accsize1 <= sizeauto with arith.f, g:treeify_tsize1, size2:natacc:list eltt1:treeacc1:list eltHg:treeify_invariant size2 gLE:S (size1 + size2) <= length accsize1 <= S (size1 + size2)f, g:treeify_tsize1, size2, size:natacc:list eltt1:treeacc1:list eltHg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ acc1let (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_tsize1, size2, size:natacc:list eltt1:treeHg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ nillet (t0, acc') := bogus in cardinal t0 = size /\ acc = elements t0 ++ acc'f, g:treeify_tsize1, size2, size:natacc:list eltt1:treex:eltacc1:list eltHg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ x :: acc1let (t0, acc') := let (r, acc0) := g acc1 in (Bk t1 x r, acc0) in cardinal t0 = size /\ acc = elements t0 ++ acc'f, g:treeify_tsize1, size2, size:natacc:list eltt1:treeHg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ nillet (t0, acc') := bogus in cardinal t0 = size /\ acc = elements t0 ++ acc'f, g:treeify_tsize1, size2, size:natacc:list eltt1:treeHg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ nilFalsef, g:treeify_tsize1, size2, size:natacc:list eltt1:treeHg:treeify_invariant size2 gEQ:size = S (size1 + size2)Hf1:cardinal t1 = size1Hf2:acc = elements t1 ++ nilsize <= length acc -> Falsef, g:treeify_tsize1, size2, size:natacc:list eltt1:treeHg:treeify_invariant size2 gEQ:size = S (size1 + size2)Hf1:cardinal t1 = size1Hf2:acc = elements t1 ++ nillength acc < sizerewrite app_nil_r, <- elements_cardinal; auto with arith.f, g:treeify_tsize2:natt1:treeHg:treeify_invariant size2 glength (elements t1 ++ nil) < S (cardinal t1 + size2)f, g:treeify_tsize1, size2, size:natacc:list eltt1:treex:eltacc1:list eltHg:treeify_invariant size2 gEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ x :: acc1let (t0, acc') := let (r, acc0) := g acc1 in (Bk t1 x r, acc0) in cardinal t0 = size /\ acc = elements t0 ++ acc'f, g:treeify_tsize1, size2, size:natacc:list eltt1:treex:eltacc1:list eltHg:size2 <= length acc1 -> let (t0, acc') := g acc1 in cardinal t0 = size2 /\ acc1 = elements t0 ++ acc'EQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ x :: acc1let (t0, acc') := let (r, acc0) := g acc1 in (Bk t1 x r, acc0) in cardinal t0 = size /\ acc = elements t0 ++ acc'f, g:treeify_tsize1, size2, size:natacc:list eltt1:treex:eltacc1:list eltt2:treeacc2:list eltHg:size2 <= length acc1 -> cardinal t2 = size2 /\ acc1 = elements t2 ++ acc2EQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ x :: acc1cardinal (Bk t1 x t2) = size /\ acc = elements (Bk t1 x t2) ++ acc2f, g:treeify_tsize1, size2, size:natacc:list eltt1:treex:eltacc1:list eltt2:treeacc2:list eltEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ x :: acc1size2 <= length acc1f, g:treeify_tsize1, size2, size:natacc:list eltt1:treex:eltacc1:list eltt2:treeacc2:list eltEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ x :: acc1Hg1:cardinal t2 = size2Hg2:acc1 = elements t2 ++ acc2cardinal (Bk t1 x t2) = size /\ acc = elements (Bk t1 x t2) ++ acc2f, g:treeify_tsize1, size2, size:natacc:list eltt1:treex:eltacc1:list eltt2:treeacc2:list eltEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ x :: acc1size2 <= length acc1f, g:treeify_tsize1, size2, size:natacc:list eltt1:treex:eltacc1:list eltt2:treeacc2:list eltEQ:size = S (size1 + size2)Hf1:cardinal t1 = size1Hf2:acc = elements t1 ++ x :: acc1size <= length acc -> size2 <= length acc1f, g:treeify_tsize2:natt1:treex:eltacc1:list eltt2:treeacc2:list eltS (cardinal t1 + size2) <= length (elements t1 ++ x :: acc1) -> size2 <= length acc1f, g:treeify_tsize2:natt1:treex:eltacc1:list eltt2:treeacc2:list eltS (cardinal t1 + size2) <= cardinal t1 + length (x :: acc1) -> size2 <= length acc1f, g:treeify_tsize2:natt1:treex:eltacc1:list eltt2:treeacc2:list eltS (cardinal t1 + size2) <= cardinal t1 + S (length acc1) -> size2 <= length acc1apply Nat.add_le_mono_l.f, g:treeify_tsize2:natt1:treex:eltacc1:list eltt2:treeacc2:list eltcardinal t1 + size2 <= cardinal t1 + length acc1 -> size2 <= length acc1f, g:treeify_tsize1, size2, size:natacc:list eltt1:treex:eltacc1:list eltt2:treeacc2:list eltEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ x :: acc1Hg1:cardinal t2 = size2Hg2:acc1 = elements t2 ++ acc2cardinal (Bk t1 x t2) = size /\ acc = elements (Bk t1 x t2) ++ acc2now subst. Qed.f, g:treeify_tsize1, size2, size:natacc:list eltt1:treex:eltacc1:list eltt2:treeacc2:list eltEQ:size = S (size1 + size2)LE:size <= length accHf1:cardinal t1 = size1Hf2:acc = elements t1 ++ x :: acc1Hg1:cardinal t2 = size2Hg2:acc1 = elements t2 ++ acc2cardinal (Bk t1 x t2) = size /\ acc = elements t1 ++ (x :: elements t2) ++ acc2n:positivep:booltreeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)n:positivep:booltreeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)n:positiveforall p : bool, treeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:booltreeify_invariant (ifpred p (Pos.to_nat n~1)) (treeify_cont (treeify_aux false n) (treeify_aux p n))n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:booltreeify_invariant (ifpred p (Pos.to_nat n~0)) (treeify_cont (treeify_aux p n) (treeify_aux true n))p:booltreeify_invariant (ifpred p (Pos.to_nat 1)) (if p then treeify_zero else treeify_one)n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:booltreeify_invariant (ifpred p (Pos.to_nat n~1)) (treeify_cont (treeify_aux false n) (treeify_aux p n))n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:boolifpred p (Pos.to_nat n~1) = S (Pos.to_nat n + ifpred p (Pos.to_nat n))n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:boolifpred p (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred p (Pos.to_nat n))n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:boolH:0 < Pos.to_nat nifpred p (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred p (Pos.to_nat n))n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:boolH:Pos.to_nat n <> 0ifpred p (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred p (Pos.to_nat n))now rewrite <- Nat.add_succ_r, Nat.succ_pred; trivial.n:positiveIHn:forall p : bool, treeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)H:Pos.to_nat n <> 0Pos.to_nat n + Pos.to_nat n = S (Pos.to_nat n + Init.Nat.pred (Pos.to_nat n))n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:booltreeify_invariant (ifpred p (Pos.to_nat n~0)) (treeify_cont (treeify_aux p n) (treeify_aux true n))n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:boolifpred p (Pos.to_nat n~0) = S (ifpred p (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:boolifpred p (2 * Pos.to_nat n) = S (ifpred p (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:boolH:0 < Pos.to_nat nifpred p (2 * Pos.to_nat n) = S (ifpred p (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:boolH:Pos.to_nat n <> 0ifpred p (2 * Pos.to_nat n) = S (ifpred p (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))n:positiveIHn:forall p0 : bool, treeify_invariant (ifpred p0 (Pos.to_nat n)) (treeify_aux p0 n)p:boolH:Pos.to_nat n <> 0ifpred p (2 * Pos.to_nat n) = ifpred p (Pos.to_nat n) + Pos.to_nat nn:positiveIHn:forall p : bool, treeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)H:Pos.to_nat n <> 0Init.Nat.pred (Pos.to_nat n + Pos.to_nat n) = Init.Nat.pred (Pos.to_nat n) + Pos.to_nat nnow apply Nat.add_pred_l.n:positiveIHn:forall p : bool, treeify_invariant (ifpred p (Pos.to_nat n)) (treeify_aux p n)H:Pos.to_nat n <> 0Init.Nat.pred (Pos.to_nat n) + Pos.to_nat n = Init.Nat.pred (Pos.to_nat n + Pos.to_nat n)destruct p; [ apply treeify_zero_spec | apply treeify_one_spec ]. Qed.p:booltreeify_invariant (ifpred p (Pos.to_nat 1)) (if p then treeify_zero else treeify_one)l:list eltp:positivePos.to_nat (plength_aux l p) = length l + Pos.to_nat pl:list eltp:positivePos.to_nat (plength_aux l p) = length l + Pos.to_nat pl:list eltforall p : positive, Pos.to_nat (plength_aux l p) = length l + Pos.to_nat pa:eltl:list eltIHl:forall p : positive, Pos.to_nat (plength_aux l p) = length l + Pos.to_nat pforall p : positive, Pos.to_nat (plength_aux (a :: l) p) = length (a :: l) + Pos.to_nat pa:eltl:list eltIHl:forall p : positive, Pos.to_nat (plength_aux l p) = length l + Pos.to_nat pforall p : positive, Pos.to_nat (plength_aux l (Pos.succ p)) = length (a :: l) + Pos.to_nat pnow rewrite IHl, Pos2Nat.inj_succ, Nat.add_succ_r. Qed.a:eltl:list eltIHl:forall p0 : positive, Pos.to_nat (plength_aux l p0) = length l + Pos.to_nat p0p:positivePos.to_nat (plength_aux l (Pos.succ p)) = length (a :: l) + Pos.to_nat pl:list eltPos.to_nat (plength l) = S (length l)l:list eltPos.to_nat (plength l) = S (length l)l:list eltPos.to_nat (plength_aux l 1) = S (length l)apply Nat.add_1_r. Qed.l:list eltlength l + Pos.to_nat 1 = S (length l)l:list eltelements (treeify l) = ll:list eltelements (treeify l) = ll:list eltH: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) = ll:list eltH: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)) = ll:list eltt:treeacc:list eltH:Init.Nat.pred (Pos.to_nat (plength l)) <= length l -> cardinal t = Init.Nat.pred (Pos.to_nat (plength l)) /\ l = elements t ++ accelements t = ll:list eltt:treeacc:list eltInit.Nat.pred (Pos.to_nat (plength l)) <= length ll:list eltt:treeacc:list eltH:cardinal t = Init.Nat.pred (Pos.to_nat (plength l))H':l = elements t ++ accelements t = lnow rewrite plength_spec.l:list eltt:treeacc:list eltInit.Nat.pred (Pos.to_nat (plength l)) <= length ll:list eltt:treeacc:list eltH:cardinal t = Init.Nat.pred (Pos.to_nat (plength l))H':l = elements t ++ accelements t = lt:treeacc:list eltH:cardinal t = Init.Nat.pred (Pos.to_nat (plength (elements t ++ acc)))elements t = elements t ++ acct:treeacc:list eltH:cardinal t = Init.Nat.pred (S (cardinal t + length acc))elements t = elements t ++ acct:treeH:cardinal t = Init.Nat.pred (S (cardinal t + length nil))elements t = elements t ++ nilt:treee:eltacc:list eltH:cardinal t = Init.Nat.pred (S (cardinal t + length (e :: acc)))elements t = elements t ++ e :: accnow rewrite app_nil_r.t:treeH:cardinal t = Init.Nat.pred (S (cardinal t + length nil))elements t = elements t ++ nilt:treee:eltacc:list eltH:cardinal t = Init.Nat.pred (S (cardinal t + length (e :: acc)))elements t = elements t ++ e :: acct:treee:eltacc:list eltH:cardinal t = Init.Nat.pred (S (cardinal t + length (e :: acc)))Falset:treee:eltacc:list eltcardinal t = Init.Nat.pred (S (cardinal t + length (e :: acc))) -> Falset:treee:eltacc:list eltcardinal t = cardinal t + S (length acc) -> Falseapply Nat.succ_add_discr. Qed.t:treee:eltacc:list eltcardinal t = S (length acc + cardinal t) -> Falsex:eltl:list eltInT x (treeify l) <-> InA X.eq x lx:eltl:list eltInT x (treeify l) <-> InA X.eq x lnow rewrite <- elements_spec1, treeify_elements. Qed.x:eltl:list eltInT x (treeify l) <-> InA X.eq x ll:list X.tSorted X.lt l -> Ok (treeify l)l:list X.tSorted X.lt l -> Ok (treeify l)l:list X.tH:Sorted X.lt lOk (treeify l)rewrite treeify_elements; auto. Qed.l:list X.tH:Sorted X.lt lSorted X.lt (elements (treeify l))
A:Typef:A -> booll, l':list AList.filter f (l ++ l') = List.filter f l ++ List.filter f l'A:Typef:A -> booll, l':list AList.filter f (l ++ l') = List.filter f l ++ List.filter f l'destruct (f x); simpl; now rewrite IH. Qed.A:Typef:A -> boolx:Al, l':list AIH: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's:treef:elt -> boolacc:list X.tfilter_aux f s acc = List.filter f (elements s) ++ accs:treef:elt -> boolacc:list X.tfilter_aux f s acc = List.filter f (elements s) ++ accs:treef:elt -> boolforall acc : list X.t, filter_aux f s acc = List.filter f (elements s) ++ accc:Color.tl:treex:X.tr:treef:elt -> boolIHl:forall acc : list X.t, filter_aux f l acc = List.filter f (elements l) ++ accIHr:forall acc : list X.t, filter_aux f r acc = List.filter f (elements r) ++ accforall acc : list X.t, filter_aux f (Node c l x r) acc = List.filter f (elements (Node c l x r)) ++ accc:Color.tl:treex:X.tr:treef:elt -> boolIHl:forall acc0 : list X.t, filter_aux f l acc0 = List.filter f (elements l) ++ acc0IHr:forall acc0 : list X.t, filter_aux f r acc0 = List.filter f (elements r) ++ acc0acc:list X.tfilter_aux f (Node c l x r) acc = List.filter f (elements (Node c l x r)) ++ accc:Color.tl:treex:X.tr:treef:elt -> boolIHl:forall acc0 : list X.t, filter_aux f l acc0 = List.filter f (elements l) ++ acc0IHr:forall acc0 : list X.t, filter_aux f r acc0 = List.filter f (elements r) ++ acc0acc:list X.tfilter_aux f (Node c l x r) acc = (List.filter f (elements l) ++ List.filter f (x :: elements r)) ++ accdestruct (f x); now rewrite IHl, IHr, app_ass. Qed.c:Color.tl:treex:X.tr:treef:elt -> boolIHl:forall acc0 : list X.t, filter_aux f l acc0 = List.filter f (elements l) ++ acc0IHr:forall acc0 : list X.t, filter_aux f r acc0 = List.filter f (elements r) ++ acc0acc: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))) ++ accs:tf:elt -> boolelements (filter f s) = List.filter f (elements s)s:tf:elt -> boolelements (filter f s) = List.filter f (elements s)now rewrite treeify_elements, filter_aux_elements, app_nil_r. Qed.s:tf:elt -> boolelements (treeify (filter_aux f s nil)) = List.filter f (elements s)s:tx:eltf:X.t -> boolProper (X.eq ==> Logic.eq) f -> InT x (filter f s) <-> InT x s /\ f x = trues:tx:eltf:X.t -> boolProper (X.eq ==> Logic.eq) f -> InT x (filter f s) <-> InT x s /\ f x = truerewrite <- elements_spec1, filter_elements, filter_InA, elements_spec1; now auto_tc. Qed.s:tx:eltf:X.t -> boolHf:Proper (X.eq ==> Logic.eq) fInT x (filter f s) <-> InT x s /\ f x = trues:treef:elt -> boolH:Ok sOk (filter f s)s:treef:elt -> boolH:Ok sOk (filter f s)s:treef:elt -> boolH:Ok sSorted X.lt (elements (filter f s))apply filter_sort with X.eq; auto_tc. Qed.s:treef:elt -> boolH:Ok sSorted X.lt (List.filter f (elements s))
s:treef:elt -> boolacc1, acc2:list X.tpartition_aux f s acc1 acc2 = (filter_aux f s acc1, filter_aux (fun x : elt => negb (f x)) s acc2)s:treef:elt -> boolacc1, acc2:list X.tpartition_aux f s acc1 acc2 = (filter_aux f s acc1, filter_aux (fun x : elt => negb (f x)) s acc2)s:treef:elt -> boolforall 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 -> boolforall acc1 acc2 : list X.t, (acc1, acc2) = (acc1, acc2)c:Color.tl:treex:X.tr:treef:elt -> boolHl: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))trivial.f:elt -> boolforall acc1 acc2 : list X.t, (acc1, acc2) = (acc1, acc2)c:Color.tl:treex:X.tr:treef:elt -> boolHl: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))destruct (f x); simpl; now rewrite Hr, Hl. Qed.c:Color.tl:treex:X.tr:treef:elt -> boolHl: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))s:tf:elt -> boolpartition f s = (filter f s, filter (fun x : elt => negb (f x)) s)s:tf:elt -> boolpartition f s = (filter f s, filter (fun x : elt => negb (f x)) s)now rewrite partition_aux_spec. Qed.s:tf: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))s:tf:X.t -> boolProper (X.eq ==> Logic.eq) f -> Equal (fst (partition f s)) (filter f s)now rewrite partition_spec. Qed.s:tf:X.t -> boolProper (X.eq ==> Logic.eq) f -> Equal (fst (partition f s)) (filter f s)s:tf:X.t -> boolProper (X.eq ==> Logic.eq) f -> Equal (snd (partition f s)) (filter (fun x : elt => negb (f x)) s)now rewrite partition_spec. Qed.s:tf:X.t -> boolProper (X.eq ==> Logic.eq) f -> Equal (snd (partition f s)) (filter (fun x : elt => negb (f x)) s)s:treef:elt -> boolH:Ok sOk (fst (partition f s))rewrite partition_spec; now apply filter_ok. Qed.s:treef:elt -> boolH:Ok sOk (fst (partition f s))s:treef:elt -> boolH:Ok sOk (snd (partition f s))rewrite partition_spec; now apply filter_ok. Qed.s:treef:elt -> boolH:Ok sOk (snd (partition f s))
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:treeH:Ok s1H0:Ok s2INV (rev_elements s1) (rev_elements s2) nils1, s2:treeH:Ok s1H0:Ok s2INV (rev_elements s1) (rev_elements s2) nilsplit; rewrite ?rev_involutive; auto; intros; now inA. Qed.s1, s2:treeH:Ok s1H0:Ok s2INV (rev (elements s1)) (rev (elements s2)) nill1, l2, acc:list X.tINV l1 l2 acc -> INV l2 l1 accdestruct 1; now split. Qed.l1, l2, acc:list X.tINV l1 l2 acc -> INV l2 l1 accx1:X.tl1, l2, acc:list X.tINV (x1 :: l1) l2 acc -> INV l1 l2 accx1:X.tl1, l2, acc:list X.tINV (x1 :: l1) l2 acc -> INV l1 l2 accx1:X.tl1, l2, acc:list X.tl1s:Sorted X.lt (rev (x1 :: l1))l2s:Sorted X.lt (rev l2)accs:Sorted X.lt accl1a:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yl2a:forall x y : X.t, InA X.eq x l2 -> InA X.eq y acc -> X.lt x yINV l1 l2 accx1:X.tl1, l2, acc:list X.tl1s:Sorted X.lt (rev l1 ++ x1 :: nil)l2s:Sorted X.lt (rev l2)accs:Sorted X.lt accl1a:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yl2a:forall x y : X.t, InA X.eq x l2 -> InA X.eq y acc -> X.lt x yINV l1 l2 accsplit; auto. Qed.x1:X.tl1, l2, acc:list X.tl1s:Sorted X.lt (rev l1 ++ x1 :: nil)l2s:Sorted X.lt (rev l2)accs:Sorted X.lt accl1a:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yl2a:forall x y : X.t, InA X.eq x l2 -> InA X.eq y acc -> X.lt x yU: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 x2INV l1 l2 accx1, x2:X.tl1, l2, acc:list X.tINV (x1 :: l1) (x2 :: l2) acc -> X.eq x1 x2 -> INV l1 l2 (x1 :: acc)x1, x2:X.tl1, l2, acc:list X.tINV (x1 :: l1) (x2 :: l2) acc -> X.eq x1 x2 -> INV l1 l2 (x1 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev (x1 :: l1))V:Sorted X.lt (rev (x2 :: l2))W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.eq x1 x2INV l1 l2 (x1 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.eq x1 x2INV l1 l2 (x1 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.eq x1 x2U1: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 x3INV l1 l2 (x1 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.eq x1 x2U1: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 x3V1: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 x3INV l1 l2 (x1 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.eq x1 x2U1: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 x3V1: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 x3Sorted X.lt (x1 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.eq x1 x2U1: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 x3V1: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 x3forall x y : X.t, InA X.eq x l1 -> InA X.eq y (x1 :: acc) -> X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.eq x1 x2U1: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 x3V1: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 x3forall x y : X.t, InA X.eq x l2 -> InA X.eq y (x1 :: acc) -> X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.eq x1 x2U1: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 x3V1: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 x3Sorted X.lt (x1 :: acc)apply InA_InfA with X.eq; auto_tc.x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.eq x1 x2U1: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 x3V1: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 x3HdRel X.lt x1 accx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.eq x1 x2U1: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 x3V1: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 x3forall x y : X.t, InA X.eq x l1 -> InA X.eq y (x1 :: acc) -> X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.eq x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l1Hy:X.eq y x1X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.eq x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l1Hy:InA X.eq y accX.lt x yapply U3; inA.x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.eq x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l1Hy:X.eq y x1X.lt x yapply X; inA.x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.eq x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l1Hy:InA X.eq y accX.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.eq x1 x2U1: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 x3V1: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 x3forall x y : X.t, InA X.eq x l2 -> InA X.eq y (x1 :: acc) -> X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.eq x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l2Hy:X.eq y x1X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.eq x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l2Hy:InA X.eq y accX.lt x yrewrite Hy, EQ; apply V3; inA.x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.eq x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l2Hy:X.eq y x1X.lt x yapply Y; inA. Qed.x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.eq x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l2Hy:InA X.eq y accX.lt x yx1, x2:X.tl1, l2, acc:list X.tINV (x1 :: l1) (x2 :: l2) acc -> X.lt x1 x2 -> INV (x1 :: l1) l2 (x2 :: acc)x1, x2:X.tl1, l2, acc:list X.tINV (x1 :: l1) (x2 :: l2) acc -> X.lt x1 x2 -> INV (x1 :: l1) l2 (x2 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev (x1 :: l1))V:Sorted X.lt (rev (x2 :: l2))W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.lt x1 x2INV (x1 :: l1) l2 (x2 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.lt x1 x2INV (x1 :: l1) l2 (x2 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.lt x1 x2U1: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 x3INV (x1 :: l1) l2 (x2 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.lt x1 x2U1: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 x3V1: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 x3INV (x1 :: l1) l2 (x2 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.lt x1 x2U1: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 x3V1: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 x3Sorted X.lt (x2 :: acc)x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.lt x1 x2U1: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 x3V1: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 x3forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y (x2 :: acc) -> X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.lt x1 x2U1: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 x3V1: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 x3forall x y : X.t, InA X.eq x l2 -> InA X.eq y (x2 :: acc) -> X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.lt x1 x2U1: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 x3V1: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 x3Sorted X.lt (x2 :: acc)apply InA_InfA with X.eq; auto_tc.x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.lt x1 x2U1: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 x3V1: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 x3HdRel X.lt x2 accx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.lt x1 x2U1: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 x3V1: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 x3forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y (x2 :: acc) -> X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tHx:X.eq x x1 \/ InA X.eq x l1Hy:X.eq y x2X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tHx:X.eq x x1 \/ InA X.eq x l1Hy:InA X.eq y accX.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tHx:X.eq x x1 \/ InA X.eq x l1Hy:X.eq y x2X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tHx:X.eq x x1 \/ InA X.eq x l1X.lt x x2x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tH:InA X.eq x l1X.lt x x2apply U3; inA.x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tH:InA X.eq x l1X.lt x x1apply X; inA.x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tHx:X.eq x x1 \/ InA X.eq x l1Hy:InA X.eq y accX.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x y : X.t, InA X.eq x (x1 :: l1) -> InA X.eq y acc -> X.lt x yY:forall x y : X.t, InA X.eq x (x2 :: l2) -> InA X.eq y acc -> X.lt x yEQ:X.lt x1 x2U1: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 x3V1: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 x3forall x y : X.t, InA X.eq x l2 -> InA X.eq y (x2 :: acc) -> X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l2Hy:X.eq y x2X.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l2Hy:InA X.eq y accX.lt x yx1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l2Hy:X.eq y x2X.lt x yapply V3; inA.x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l2Hy:X.eq y x2X.lt x x2apply Y; inA. Qed.x1, x2:X.tl1, l2, acc:list X.tU:Sorted X.lt (rev l1 ++ x1 :: nil)V:Sorted X.lt (rev l2 ++ x2 :: nil)W:Sorted X.lt accX:forall x0 y0 : X.t, InA X.eq x0 (x1 :: l1) -> InA X.eq y0 acc -> X.lt x0 y0Y:forall x0 y0 : X.t, InA X.eq x0 (x2 :: l2) -> InA X.eq y0 acc -> X.lt x0 y0EQ:X.lt x1 x2U1: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 x3V1: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 x3x, y:X.tHx:InA X.eq x l2Hy:InA X.eq y accX.lt x yl1, l2, acc:list X.tINV l1 l2 acc -> Sorted X.lt (rev_append l1 acc)l1, l2, acc:list X.tINV l1 l2 acc -> Sorted X.lt (rev_append l1 acc)l1, l2, acc:list X.tH:INV l1 l2 accSorted X.lt (rev_append l1 acc)l1, l2, acc:list X.tH:INV l1 l2 accSorted X.lt (rev l1 ++ acc)l1, l2, acc:list X.tH:INV l1 l2 accforall x y : X.t, InA X.eq x (rev l1) -> InA X.eq y acc -> X.lt x yl1, l2, acc:list X.tH:INV l1 l2 accx, y:X.tInA X.eq x (rev l1) -> InA X.eq y acc -> X.lt x yeapply @l1_lt_acc; eauto. Qed.l1, l2, acc:list X.tH:INV l1 l2 accx, y:X.tInA X.eq x l1 -> InA X.eq y acc -> X.lt x y
l1, l2, acc:list X.tINV l1 l2 acc -> Sorted X.lt (union_list l1 l2 acc)l1, l2, acc:list X.tINV l1 l2 acc -> Sorted X.lt (union_list l1 l2 acc)l1:list X.tforall l2 acc : list X.t, INV l1 l2 acc -> Sorted X.lt (union_list l1 l2 acc)l2, acc:list X.tinv:INV nil l2 accSorted X.lt (union_list nil l2 acc)x1:X.tl1:list X.tIH1:forall l2 acc0 : list X.t, INV l1 l2 acc0 -> Sorted X.lt (union_list l1 l2 acc0)acc:list X.tinv:INV (x1 :: l1) nil accSorted X.lt (union_list (x1 :: l1) nil acc)x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accSorted X.lt (union_list (x1 :: l1) (x2 :: l2) acc)eapply INV_rev, INV_sym; eauto.l2, acc:list X.tinv:INV nil l2 accSorted X.lt (union_list nil l2 acc)eapply INV_rev; eauto.x1:X.tl1:list X.tIH1:forall l2 acc0 : list X.t, INV l1 l2 acc0 -> Sorted X.lt (union_list l1 l2 acc0)acc:list X.tinv:INV (x1 :: l1) nil accSorted X.lt (union_list (x1 :: l1) nil acc)x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accSorted X.lt (union_list (x1 :: l1) (x2 :: l2) acc)x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accSorted 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) endx1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.eq x1 x2Sorted X.lt (union_list l1 l2 (x1 :: acc))x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x1 x2Sorted 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.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x2 x1Sorted X.lt (union_list l1 (x2 :: l2) (x1 :: acc))x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.eq x1 x2Sorted X.lt (union_list l1 l2 (x1 :: acc))eapply INV_eq; eauto.x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.eq x1 x2INV l1 l2 (x1 :: acc)x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x1 x2Sorted 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))eapply INV_lt; eauto.x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x1 x2INV (x1 :: l1) l2 (x2 :: acc)x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x2 x1Sorted X.lt (union_list l1 (x2 :: l2) (x1 :: acc))x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x2 x1INV l1 (x2 :: l2) (x1 :: acc)now apply INV_sym. Qed.x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (union_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (union_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x2 x1INV (x2 :: l2) (x1 :: l1) accs1, s2:treeH:Ok s1H0:Ok s2Ok (linear_union s1 s2)s1, s2:treeH:Ok s1H0:Ok s2Ok (linear_union s1 s2)now apply treeify_ok, union_list_ok, INV_init. Qed.s1, s2:treeH:Ok s1H0:Ok s2Ok (treeify (union_list (rev_elements s1) (rev_elements s2) nil))s1, s2:treeH:Ok s1H0:Ok s2Ok (fold add s1 s2)s1, s2:treeH:Ok s1H0:Ok s2Ok (fold add s1 s2)s1, s2:treeH:Ok s1H0:Ok s2Ok (fold_right (fun (y : elt) (x : tree) => flip add x y) s2 (rev (elements s1)))induction (rev (elements s1)); simpl; unfold flip in *; auto_tc. Qed.s1, s2:treeH:Ok s1H0:Ok s2Ok (fold_right (fun (y : X.t) (x : tree) => flip add x y) s2 (rev (elements s1)))s1, s2:treeH:Ok s1H0:Ok s2Ok (union s1 s2)s1, s2:treeH:Ok s1H0:Ok s2Ok (union s1 s2)destruct compare_height; auto_tc. Qed.s1, s2:treeH:Ok s1H0:Ok s2Ok match compare_height s1 s1 s2 s2 with | Eq => linear_union s1 s2 | Lt => fold add s1 s2 | Gt => fold add s2 s1 endx:X.tl1, l2, acc:list eltInA X.eq x (union_list l1 l2 acc) <-> InA X.eq x l1 \/ InA X.eq x l2 \/ InA X.eq x accx:X.tl1, l2, acc:list eltInA X.eq x (union_list l1 l2 acc) <-> InA X.eq x l1 \/ InA X.eq x l2 \/ InA X.eq x accx:X.tl1:list eltforall 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 accx:X.tforall 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 accx:X.tx1:eltl1:list eltIH1: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 accforall 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 accx:X.tforall 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 accx:X.tl2, acc:list eltInA X.eq x (rev_append l2 acc) <-> InA X.eq x nil \/ InA X.eq x l2 \/ InA X.eq x accx:X.tl2, acc:list eltInA X.eq x (rev l2 ++ acc) <-> InA X.eq x nil \/ InA X.eq x l2 \/ InA X.eq x acctauto.x:X.tl2, acc:list eltInA X.eq x l2 \/ InA X.eq x acc <-> False \/ InA X.eq x l2 \/ InA X.eq x accx:X.tx1:eltl1:list eltIH1: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 accforall 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 accx:X.tx1:eltl1:list eltIH1: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 acc0acc:list eltInA X.eq x (rev_append l1 (x1 :: acc)) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x nil \/ InA X.eq x accx:X.tx1:eltl1:list eltIH1: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 acc0x2:eltl2:list eltIH2: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 acc0acc:list eltInA 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 accx:X.tx1:eltl1:list eltIH1: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 acc0acc:list eltInA X.eq x (rev_append l1 (x1 :: acc)) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x nil \/ InA X.eq x accx:X.tx1:eltl1:list eltIH1: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 acc0acc:list eltInA X.eq x (rev l1 ++ x1 :: acc) <-> InA X.eq x (x1 :: l1) \/ InA X.eq x nil \/ InA X.eq x acctauto.x:X.tx1:eltl1:list eltIH1: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 acc0acc:list eltInA 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 accx:X.tx1:eltl1:list eltIH1: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 acc0x2:eltl2:list eltIH2: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 acc0acc:list eltInA 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 accx:X.tx1:eltl1:list eltIH1: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 acc0x2:eltl2:list eltIH2: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 acc0acc:list eltC:X.eq x1 x2InA 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 accx:X.tx1:eltl1:list eltIH1: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 acc0x2:eltl2:list eltIH2: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 acc0acc:list eltC:X.lt x1 x2InA 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 accx:X.tx1:eltl1:list eltIH1: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 acc0x2:eltl2:list eltIH2: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 acc0acc:list eltC:X.lt x2 x1InA 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 accrewrite IH1, !InA_cons, C; tauto.x:X.tx1:eltl1:list eltIH1: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 acc0x2:eltl2:list eltIH2: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 acc0acc:list eltC:X.eq x1 x2InA 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 accx:X.tx1:eltl1:list eltIH1: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 acc0x2:eltl2:list eltIH2: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 acc0acc:list eltC:X.lt x1 x2InA 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 acctauto.x:X.tx1:eltl1:list eltIH1: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 acc0x2:eltl2:list eltIH2: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 acc0acc:list eltC: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 accrewrite IH1, !InA_cons; tauto. Qed.x:X.tx1:eltl1:list eltIH1: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 acc0x2:eltl2:list eltIH2: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 acc0acc:list eltC:X.lt x2 x1InA 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 accs1, s2:treex:eltInT x (linear_union s1 s2) <-> InT x s1 \/ InT x s2s1, s2:treex:eltInT x (linear_union s1 s2) <-> InT x s1 \/ InT x s2s1, s2:treex:eltInT x (treeify (union_list (rev_elements s1) (rev_elements s2) nil)) <-> InT x s1 \/ InT x s2s1, s2:treex:eltInA X.eq x (rev (elements s1)) \/ InA X.eq x (rev (elements s2)) \/ InA X.eq x nil <-> InT x s1 \/ InT x s2tauto. Qed.s1, s2:treex:eltInT x s1 \/ InT x s2 \/ False <-> InT x s1 \/ InT x s2s1, s2:treex:eltInT x (fold add s1 s2) <-> InT x s1 \/ InT x s2s1, s2:treex:eltInT x (fold add s1 s2) <-> InT x s1 \/ InT x s2s1, s2:treex:eltInT x (fold_right (fun (y : elt) (x0 : tree) => flip add x0 y) s2 (rev (elements s1))) <-> InT x s1 \/ InT x s2s1, s2:treex:eltInT 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 s2s1, s2:treex:X.tInT 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 s2s1, s2:treex:X.tInT x s2 <-> InA X.eq x nil \/ InT x s2s1, s2:treex, a:X.tl:list X.tIHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip add x0 y) s2 l) <-> InA X.eq x l \/ InT x s2InT 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 s2s1, s2:treex:X.tInT x s2 <-> InA X.eq x nil \/ InT x s2tauto.s1, s2:treex:X.tInT x s2 <-> False \/ InT x s2s1, s2:treex, a:X.tl:list X.tIHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip add x0 y) s2 l) <-> InA X.eq x l \/ InT x s2InT 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 s2s1, s2:treex, a:X.tl:list X.tIHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip add x0 y) s2 l) <-> InA X.eq x l \/ InT x s2InT x (add a (fold_right (fun (y : X.t) (x0 : tree) => add y x0) s2 l)) <-> InA X.eq x (a :: l) \/ InT x s2tauto. Qed.s1, s2:treex, a:X.tl:list X.tIHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip add x0 y) s2 l) <-> InA X.eq x l \/ InT x s2X.eq x a \/ InA X.eq x l \/ InT x s2 <-> (X.eq x a \/ InA X.eq x l) \/ InT x s2s1, s2:tx:eltInT x (union s1 s2) <-> InT x s1 \/ InT x s2s1, s2:tx:eltInT x (union s1 s2) <-> InT x s1 \/ InT x s2s1, s2:tx:eltInT 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 s2s1, s2:tx:eltInT x (linear_union s1 s2) <-> InT x s1 \/ InT x s2s1, s2:tx:eltInT x (fold add s1 s2) <-> InT x s1 \/ InT x s2s1, s2:tx:eltInT x (fold add s2 s1) <-> InT x s1 \/ InT x s2apply linear_union_spec.s1, s2:tx:eltInT x (linear_union s1 s2) <-> InT x s1 \/ InT x s2apply fold_add_spec.s1, s2:tx:eltInT x (fold add s1 s2) <-> InT x s1 \/ InT x s2s1, s2:tx:eltInT x (fold add s2 s1) <-> InT x s1 \/ InT x s2tauto. Qed.s1, s2:tx:eltInT x s2 \/ InT x s1 <-> InT x s1 \/ InT x s2forall (s1 s2 : tree) (y : elt), Ok s1 -> Ok s2 -> InT y (union s1 s2) <-> InT y s1 \/ InT y s2intros; apply union_spec'. Qed.forall (s1 s2 : tree) (y : elt), Ok s1 -> Ok s2 -> InT y (union s1 s2) <-> InT y s1 \/ InT y s2
l1, l2, acc:list X.tINV l1 l2 acc -> Sorted X.lt (inter_list l1 l2 acc)l1, l2, acc:list X.tINV l1 l2 acc -> Sorted X.lt (inter_list l1 l2 acc)l1:list X.tforall 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 accx1:X.tl1:list X.tIH1: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 accx1:X.tl1:list X.tIH1:forall l0 acc : list X.t, INV l1 l0 acc -> Sorted X.lt (inter_list l1 l0 acc)x2:X.tl2:list X.tIH2: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 endeauto.forall l2 acc : list X.t, INV nil l2 acc -> Sorted X.lt acceauto.x1:X.tl1:list X.tIH1: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 accx1:X.tl1:list X.tIH1:forall l0 acc : list X.t, INV l1 l0 acc -> Sorted X.lt (inter_list l1 l0 acc)x2:X.tl2:list X.tIH2: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 endx1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accSorted 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 endx1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.eq x1 x2Sorted X.lt (inter_list l1 l2 (x1 :: acc))x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x1 x2Sorted 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.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x2 x1Sorted X.lt (inter_list l1 (x2 :: l2) acc)x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.eq x1 x2Sorted X.lt (inter_list l1 l2 (x1 :: acc))eapply INV_eq; eauto.x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.eq x1 x2INV l1 l2 (x1 :: acc)x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x1 x2Sorted 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)eapply INV_sym, INV_drop, INV_sym; eauto.x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x1 x2INV (x1 :: l1) l2 accx1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x2 x1Sorted X.lt (inter_list l1 (x2 :: l2) acc)eapply INV_drop; eauto. Qed.x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (inter_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (inter_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x2 x1INV l1 (x2 :: l2) accs1, s2:treeH:Ok s1H0:Ok s2Ok (linear_inter s1 s2)s1, s2:treeH:Ok s1H0:Ok s2Ok (linear_inter s1 s2)now apply treeify_ok, inter_list_ok, INV_init. Qed.s1, s2:treeH:Ok s1H0:Ok s2Ok (treeify (inter_list (rev_elements s1) (rev_elements s2) nil))s1, s2:treeH:Ok s1H0:Ok s2Ok (inter s1 s2)s1, s2:treeH:Ok s1H0:Ok s2Ok (inter s1 s2)destruct compare_height; auto_tc. Qed.s1, s2:treeH:Ok s1H0:Ok s2Ok 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 endx:X.tl1, l2:list X.tacc:list eltSorted 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 accx:X.tl1, l2:list X.tacc:list eltSorted 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 accx:X.tl1:list X.tforall (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 accx:X.tforall (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 accx, x1:X.tl1:list X.tIH1: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 accforall (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 accx:X.tforall (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 accx:X.tl2:list X.tacc:list eltSorted 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 acctauto.x:X.tl2:list X.tacc:list eltSorted X.lt nil -> Sorted X.lt (rev l2) -> InA X.eq x acc <-> False /\ InA X.eq x l2 \/ InA X.eq x accx, x1:X.tl1:list X.tIH1: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 accforall (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 accx, x1:X.tl1:list X.tIH1: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 acc0acc:list eltSorted 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltSorted 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 accx, x1:X.tl1:list X.tIH1: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 acc0acc:list eltSorted 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 accx, x1:X.tl1:list X.tIH1: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 acc0acc:list eltSorted 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 acctauto.x, x1:X.tl1:list X.tIH1: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 acc0acc:list eltSorted 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltSorted 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltSorted 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.eq x1 x2InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x1 x2InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1InA 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 accrewrite IH1, !InA_cons, C; tauto.x, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.eq x1 x2InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x1 x2InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x1 x2InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C: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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x1 x2H0:InA X.eq x l1H:X.eq x x2(X.eq x x1 \/ InA X.eq x l1) /\ InA X.eq x l2 \/ InA X.eq x accorder.x, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x1 x2H0:InA X.eq x l1H:X.eq x x2H1:X.lt x x1(X.eq x x1 \/ InA X.eq x l1) /\ InA X.eq x l2 \/ InA X.eq x accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1H0:X.eq x x1H:InA X.eq x l2InA X.eq x l1 /\ (X.eq x x2 \/ InA X.eq x l2) \/ InA X.eq x accorder. Qed.x, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1H0:X.eq x x1H:InA X.eq x l2H1:X.lt x x2InA X.eq x l1 /\ (X.eq x x2 \/ InA X.eq x l2) \/ InA X.eq x accs1, s2:treex:eltH:Ok s1H0:Ok s2InT x (linear_inter s1 s2) <-> InT x s1 /\ InT x s2s1, s2:treex:eltH:Ok s1H0:Ok s2InT x (linear_inter s1 s2) <-> InT x s1 /\ InT x s2s1, s2:treex:eltH:Ok s1H0:Ok s2InT x (treeify (inter_list (rev_elements s1) (rev_elements s2) nil)) <-> InT x s1 /\ InT x s2s1, s2:treex:eltH:Ok s1H0:Ok s2InA X.eq x (rev (elements s1)) /\ InA X.eq x (rev (elements s2)) \/ InA X.eq x nil <-> InT x s1 /\ InT x s2tauto. Qed.s1, s2:treex:eltH:Ok s1H0:Ok s2InT x s1 /\ InT x s2 \/ False <-> InT x s1 /\ InT x s2s:treeH:Ok sProper (X.eq ==> Logic.eq) (fun k : X.t => mem k s)s:treeH:Ok sProper (X.eq ==> Logic.eq) (fun k : X.t => mem k s)s:treeH:Ok sx, y:X.tEQ:X.eq x ymem x s = mem y snow rewrite EQ. Qed.s:treeH:Ok sx, y:X.tEQ:X.eq x yInT x s <-> InT y ss1, s2:treey:eltH:Ok s1H0:Ok s2InT y (inter s1 s2) <-> InT y s1 /\ InT y s2s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (inter s1 s2) <-> InT y s1 /\ InT y s2s1, s2:treey:eltH:Ok s1H0:Ok s2InT 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 s2s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (linear_inter s1 s2) <-> InT y s1 /\ InT y s2s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (filter (fun k : elt => mem k s2) s1) <-> InT y s1 /\ InT y s2s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (filter (fun k : elt => mem k s1) s2) <-> InT y s1 /\ InT y s2now apply linear_inter_spec.s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (linear_inter s1 s2) <-> InT y s1 /\ InT y s2rewrite filter_spec, mem_spec by auto_tc; tauto.s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (filter (fun k : elt => mem k s2) s1) <-> InT y s1 /\ InT y s2rewrite filter_spec, mem_spec by auto_tc; tauto. Qed.s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (filter (fun k : elt => mem k s1) s2) <-> InT y s1 /\ InT y s2
l1, l2, acc:list X.tINV l1 l2 acc -> Sorted X.lt (diff_list l1 l2 acc)l1, l2, acc:list X.tINV l1 l2 acc -> Sorted X.lt (diff_list l1 l2 acc)l1:list X.tforall l2 acc : list X.t, INV l1 l2 acc -> Sorted X.lt (diff_list l1 l2 acc)l2, acc:list X.tinv:INV nil l2 accSorted X.lt (diff_list nil l2 acc)x1:X.tl1:list X.tIH1:forall l2 acc0 : list X.t, INV l1 l2 acc0 -> Sorted X.lt (diff_list l1 l2 acc0)acc:list X.tinv:INV (x1 :: l1) nil accSorted X.lt (diff_list (x1 :: l1) nil acc)x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accSorted X.lt (diff_list (x1 :: l1) (x2 :: l2) acc)eauto.l2, acc:list X.tinv:INV nil l2 accSorted X.lt (diff_list nil l2 acc)x1:X.tl1:list X.tIH1:forall l2 acc0 : list X.t, INV l1 l2 acc0 -> Sorted X.lt (diff_list l1 l2 acc0)acc:list X.tinv:INV (x1 :: l1) nil accSorted X.lt (diff_list (x1 :: l1) nil acc)eapply INV_rev; eauto.x1:X.tl1:list X.tIH1:forall l2 acc0 : list X.t, INV l1 l2 acc0 -> Sorted X.lt (diff_list l1 l2 acc0)acc:list X.tinv:INV (x1 :: l1) nil accSorted X.lt (rev_append (x1 :: l1) acc)x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accSorted X.lt (diff_list (x1 :: l1) (x2 :: l2) acc)x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accSorted 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) endx1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.eq x1 x2Sorted X.lt (diff_list l1 l2 acc)x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x1 x2Sorted 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.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x2 x1Sorted X.lt (diff_list l1 (x2 :: l2) (x1 :: acc))x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.eq x1 x2Sorted X.lt (diff_list l1 l2 acc)eapply INV_drop, INV_sym, INV_drop, INV_sym; eauto.x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.eq x1 x2INV l1 l2 accx1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x1 x2Sorted 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)eapply INV_sym, INV_drop, INV_sym; eauto.x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x1 x2INV (x1 :: l1) l2 accx1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x2 x1Sorted X.lt (diff_list l1 (x2 :: l2) (x1 :: acc))x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x2 x1INV l1 (x2 :: l2) (x1 :: acc)now apply INV_sym. Qed.x1:X.tl1:list X.tIH1:forall l0 acc0 : list X.t, INV l1 l0 acc0 -> Sorted X.lt (diff_list l1 l0 acc0)x2:X.tl2:list X.tIH2:forall acc0 : list X.t, INV (x1 :: l1) l2 acc0 -> Sorted X.lt (diff_list (x1 :: l1) l2 acc0)acc:list X.tinv:INV (x1 :: l1) (x2 :: l2) accC:X.lt x2 x1INV (x2 :: l2) (x1 :: l1) accs1, s2:treeH:Ok s1H0:Ok s2Ok (linear_diff s1 s2)s1, s2:treeH:Ok s1H0:Ok s2Ok (linear_diff s1 s2)now apply treeify_ok, diff_list_ok, INV_init. Qed.s1, s2:treeH:Ok s1H0:Ok s2Ok (linear_diff s1 s2)s1, s2:treeH:Ok s2Ok (fold remove s1 s2)s1, s2:treeH:Ok s2Ok (fold remove s1 s2)s1, s2:treeH:Ok s2Ok (fold_right (fun (y : elt) (x : tree) => flip remove x y) s2 (rev (elements s1)))induction (rev (elements s1)); simpl; unfold flip in *; auto_tc. Qed.s1, s2:treeH:Ok s2Ok (fold_right (fun (y : X.t) (x : tree) => flip remove x y) s2 (rev (elements s1)))s1, s2:treeH:Ok s1H0:Ok s2Ok (diff s1 s2)s1, s2:treeH:Ok s1H0:Ok s2Ok (diff s1 s2)destruct compare_height; auto_tc. Qed.s1, s2:treeH:Ok s1H0:Ok s2Ok 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 endx:X.tl1, l2:list X.tacc:list eltSorted 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 accx:X.tl1, l2:list X.tacc:list eltSorted 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 accx:X.tl1:list X.tforall (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 accx:X.tforall (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 accx, x1:X.tl1:list X.tIH1: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 accforall (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 accx:X.tforall (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 accx:X.tl2:list X.tacc:list eltSorted 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 acctauto.x:X.tl2:list X.tacc:list eltSorted X.lt nil -> Sorted X.lt (rev l2) -> InA X.eq x acc <-> False /\ ~ InA X.eq x l2 \/ InA X.eq x accx, x1:X.tl1:list X.tIH1: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 accforall (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 accx, x1:X.tl1:list X.tIH1: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 acc0acc:list eltSorted 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltSorted 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 accx, x1:X.tl1:list X.tIH1: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 acc0acc:list eltSorted 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 accx, x1:X.tl1:list X.tIH1: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 acc0acc:list eltH: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 accx, x1:X.tl1:list X.tIH1: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 acc0acc:list eltH: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 acctauto.x, x1:X.tl1:list X.tIH1: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 acc0acc:list eltH: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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltSorted 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltSorted 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.eq x1 x2InA X.eq x (diff_list l1 l2 acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x1 x2InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.eq x1 x2InA X.eq x (diff_list l1 l2 acc) <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2) \/ InA X.eq x accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.eq x1 x2InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.eq x1 x2InA X.eq x l1 /\ ~ InA X.eq x l2 <-> InA X.eq x (x1 :: l1) /\ ~ InA X.eq x (x2 :: l2)x, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.eq x1 x2InA 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.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.eq x1 x2H0:InA X.eq x l1H1:InA X.eq x l2 -> FalseH2:X.eq x x2Falseorder.x, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.eq x1 x2H0:InA X.eq x l1H1:InA X.eq x l2 -> FalseH2:X.eq x x2H:X.lt x x1Falsex, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x1 x2InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x1 x2InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x1 x2InA 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.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C: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.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x1 x2H1:InA X.eq x l2 -> FalseH:InA X.eq x l1H2:X.eq x x2Falseorder.x, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x1 x2H1:InA X.eq x l2 -> FalseH:InA X.eq x l1H2:X.eq x x2H0:X.lt x x1Falsex, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1InA 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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1H: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 accx, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1H:X.eq x x1X.eq x x2 \/ InA X.eq x l2 -> Falsex, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1H:X.eq x x1H0:X.eq x x2Falsex, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1H:X.eq x x1H0:InA X.eq x l2Falsex, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1H:X.eq x x1H0:InA X.eq x l2Falseorder. Qed.x, x1:X.tl1:list X.tIH1: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 acc0x2:X.tl2:list X.tIH2: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 acc0acc:list eltU: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 x3V1: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 x3C:X.lt x2 x1H:X.eq x x1H0:InA X.eq x l2H1:X.lt x x2Falses1, s2:treex:eltH:Ok s1H0:Ok s2InT x (linear_diff s1 s2) <-> InT x s1 /\ ~ InT x s2s1, s2:treex:eltH:Ok s1H0:Ok s2InT x (linear_diff s1 s2) <-> InT x s1 /\ ~ InT x s2s1, s2:treex:eltH:Ok s1H0:Ok s2InT x (treeify (diff_list (rev_elements s1) (rev_elements s2) nil)) <-> InT x s1 /\ ~ InT x s2s1, s2:treex:eltH:Ok s1H0:Ok s2InA X.eq x (rev (elements s1)) /\ ~ InA X.eq x (rev (elements s2)) \/ InA X.eq x nil <-> InT x s1 /\ ~ InT x s2tauto. Qed.s1, s2:treex:eltH:Ok s1H0:Ok s2InT x s1 /\ ~ InT x s2 \/ False <-> InT x s1 /\ ~ InT x s2s1, s2:treex:eltH:Ok s2InT x (fold remove s1 s2) <-> InT x s2 /\ ~ InT x s1s1, s2:treex:eltH:Ok s2InT x (fold remove s1 s2) <-> InT x s2 /\ ~ InT x s1s1, s2:treex:eltH:Ok s2InT x (fold_right (fun (y : elt) (x0 : tree) => flip remove x0 y) s2 (rev (elements s1))) <-> InT x s2 /\ ~ InT x s1s1, s2:treex:eltH:Ok s2InT 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:treex:X.tH:Ok s2InT 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:treex:X.tH:Ok s2InT x s2 <-> InT x s2 /\ ~ InA X.eq x nils1, s2:treex:X.tH:Ok s2a:X.tl:list X.tIHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip remove x0 y) s2 l) <-> InT x s2 /\ ~ InA X.eq x lInT 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:treex:X.tH:Ok s2InT x s2 <-> InT x s2 /\ ~ InA X.eq x nilintuition.s1, s2:treex:X.tH:Ok s2InT x s2 <-> InT x s2 /\ ~ Falses1, s2:treex:X.tH:Ok s2a:X.tl:list X.tIHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => flip remove x0 y) s2 l) <-> InT x s2 /\ ~ InA X.eq x lInT 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:treex:X.tH:Ok s2a:X.tl:list X.tIHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l) <-> InT x s2 /\ ~ InA X.eq x lInT 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:treex:X.tH:Ok s2a:X.tl:list X.tIHl: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:treex:X.tH:Ok s2a:X.tl:list X.tIHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l) <-> InT x s2 /\ ~ InA X.eq x lOk (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l)s1, s2:treex:X.tH:Ok s2a:X.tl:list X.tIHl:InT x (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l) <-> InT x s2 /\ ~ InA X.eq x lOk (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l)induction l; simpl; auto_tc. Qed.s1, s2:treex:X.tH:Ok s2a:X.tl:list X.tOk (fold_right (fun (y : X.t) (x0 : tree) => remove y x0) s2 l)s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (diff s1 s2) <-> InT y s1 /\ ~ InT y s2s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (diff s1 s2) <-> InT y s1 /\ ~ InT y s2s1, s2:treey:eltH:Ok s1H0:Ok s2InT 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 s2s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (linear_diff s1 s2) <-> InT y s1 /\ ~ InT y s2s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (filter (fun k : elt => negb (mem k s2)) s1) <-> InT y s1 /\ ~ InT y s2s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (fold remove s2 s1) <-> InT y s1 /\ ~ InT y s2now apply linear_diff_spec.s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (linear_diff s1 s2) <-> InT y s1 /\ ~ InT y s2s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (filter (fun k : elt => negb (mem k s2)) s1) <-> InT y s1 /\ ~ InT y s2s1, s2:treey:eltH:Ok s1H0:Ok s2Proper (X.eq ==> Logic.eq) (fun k : elt => negb (mem k s2))s1, s2:treey:eltH:Ok s1H0:Ok s2x1, x2:X.tEQ:X.eq x1 x2negb (mem x1 s2) = negb (mem x2 s2)now apply mem_proper.s1, s2:treey:eltH:Ok s1H0:Ok s2x1, x2:X.tEQ:X.eq x1 x2mem x1 s2 = mem x2 s2now apply fold_remove_spec. Qed. End MakeRaw.s1, s2:treey:eltH:Ok s1H0:Ok s2InT y (fold remove s2 s1) <-> InT y s1 /\ ~ InT y s2
Balancing properties
Module BalanceProps(X:Orders.OrderedType)(Import M : MakeRaw X). Notation Rd := (Node Red). Notation Bk := (Node Black). Import M.MX.
In a red-black tree :
- a red node has no red children
- the black depth at each node is the same along all paths.
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.
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:natt:treerrt n t -> notredred t -> rbt n tn:natt:treerrt n t -> notredred t -> rbt n tdestruct l, r; descolor; invrb; auto. Qed. Local Hint Resolve rr_nrr_rb : core.n:natl:treex:X.tr:treeHl:rbt n lHr:rbt n rnotredred (Rd l x r) -> rbt n (Rd l x r)n:natt:treearbt n t -> notredred t -> rbt n tdestruct 1; auto. Qed.n:natt:treearbt n t -> notredred t -> rbt n tn:natt:treearbt n t -> notred t -> rbt n tdestruct 1; destruct t; descolor; invrb; auto. Qed. Local Hint Resolve arb_nrr_rb arb_nr_rb : core.n:natt:treearbt n t -> notred t -> rbt n t
Definition redcarac s := rcase (fun _ _ _ => 1) (fun _ => 0) s.s:treen:natrbt n s -> maxdepth s <= 2 * n + redcarac ss:treen:natrbt n s -> maxdepth s <= 2 * n + redcarac smaxdepth Leaf <= 2 * 0 + redcarac Leafn:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:maxdepth l <= 2 * n + redcarac lIHrbt2:maxdepth r <= 2 * n + redcarac rmaxdepth (Rd l k r) <= 2 * n + redcarac (Rd l k r)n:natl:treek:X.tr:treeH:rbt n lH0:rbt n rIHrbt1:maxdepth l <= 2 * n + redcarac lIHrbt2:maxdepth r <= 2 * n + redcarac rmaxdepth (Bk l k r) <= 2 * S n + redcarac (Bk l k r)simpl; auto.maxdepth Leaf <= 2 * 0 + redcarac Leafn:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:maxdepth l <= 2 * n + redcarac lIHrbt2:maxdepth r <= 2 * n + redcarac rmaxdepth (Rd l k r) <= 2 * n + redcarac (Rd l k r)n:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:maxdepth l <= 2 * n + 0IHrbt2:maxdepth r <= 2 * n + redcarac rmaxdepth (Rd l k r) <= 2 * n + redcarac (Rd l k r)n:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:maxdepth l <= 2 * n + 0IHrbt2:maxdepth r <= 2 * n + 0maxdepth (Rd l k r) <= 2 * n + redcarac (Rd l k r)n:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:maxdepth l <= 2 * n + 0IHrbt2:maxdepth r <= 2 * n + 0S (Init.Nat.max (maxdepth l) (maxdepth r)) <= 2 * n + redcarac (Rd l k r)n:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:maxdepth l <= 2 * n + 0IHrbt2:maxdepth r <= 2 * n + 0S (Init.Nat.max (maxdepth l) (maxdepth r)) <= 2 * n + 1now apply Nat.max_lub.n:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:maxdepth l <= 2 * n + 0IHrbt2:maxdepth r <= 2 * n + 0Init.Nat.max (maxdepth l) (maxdepth r) <= 2 * n + 0n:natl:treek:X.tr:treeH:rbt n lH0:rbt n rIHrbt1:maxdepth l <= 2 * n + redcarac lIHrbt2:maxdepth r <= 2 * n + redcarac rmaxdepth (Bk l k r) <= 2 * S n + redcarac (Bk l k r)n:natl:treek:X.tr:treeH:rbt n lH0:rbt n rIHrbt1:maxdepth l <= 2 * n + redcarac lIHrbt2:maxdepth r <= 2 * n + redcarac rS (Init.Nat.max (maxdepth l) (maxdepth r)) <= S (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.n:natl:treek:X.tr:treeH:rbt n lH0:rbt n rIHrbt1:maxdepth l <= 2 * n + redcarac lIHrbt2:maxdepth r <= 2 * n + redcarac rInit.Nat.max (maxdepth l) (maxdepth r) <= n + S (n + 0) + 0s:treen:natrbt n s -> n + redcarac s <= mindepth ss:treen:natrbt n s -> n + redcarac s <= mindepth s0 <= 0n:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:n + redcarac l <= mindepth lIHrbt2:n + redcarac r <= mindepth rn + 1 <= S (Init.Nat.min (mindepth l) (mindepth r))n:natl:treek:X.tr:treeH:rbt n lH0:rbt n rIHrbt1:n + redcarac l <= mindepth lIHrbt2:n + redcarac r <= mindepth rS (n + 0) <= S (Init.Nat.min (mindepth l) (mindepth r))trivial.0 <= 0n:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:n + redcarac l <= mindepth lIHrbt2:n + redcarac r <= mindepth rn + 1 <= S (Init.Nat.min (mindepth l) (mindepth r))n:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:n + redcarac l <= mindepth lIHrbt2:n + redcarac r <= mindepth rS (n + 0) <= S (Init.Nat.min (mindepth l) (mindepth r))n:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:n + redcarac l <= mindepth lIHrbt2:n + redcarac r <= mindepth rn + 0 <= Init.Nat.min (mindepth l) (mindepth r)n:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:n + 0 <= mindepth lIHrbt2:n + redcarac r <= mindepth rn + 0 <= Init.Nat.min (mindepth l) (mindepth r)now apply Nat.min_glb.n:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1:n + 0 <= mindepth lIHrbt2:n + 0 <= mindepth rn + 0 <= Init.Nat.min (mindepth l) (mindepth r)n:natl:treek:X.tr:treeH:rbt n lH0:rbt n rIHrbt1:n + redcarac l <= mindepth lIHrbt2:n + redcarac r <= mindepth rS (n + 0) <= S (Init.Nat.min (mindepth l) (mindepth r))n:natl:treek:X.tr:treeH:rbt n lH0:rbt n rIHrbt1:n + redcarac l <= mindepth lIHrbt2:n + redcarac r <= mindepth rn + 0 <= Init.Nat.min (mindepth l) (mindepth r)apply Nat.min_glb; eauto with arith. Qed.n:natl:treek:X.tr:treeH:rbt n lH0:rbt n rIHrbt1:n + redcarac l <= mindepth lIHrbt2:n + redcarac r <= mindepth rn <= Init.Nat.min (mindepth l) (mindepth r)s:treeRbt s -> maxdepth s <= 2 * Nat.log2 (S (cardinal s))s:treeRbt s -> maxdepth s <= 2 * Nat.log2 (S (cardinal s))s:treen:natH:rbt n smaxdepth s <= 2 * Nat.log2 (S (cardinal s))s:treen:natH:rbt n s2 * n + redcarac s <= 2 * Nat.log2 (S (cardinal s))s:treen:natH:rbt n s2 * n + redcarac s <= 2 * (n + redcarac s)s:treen:natH:rbt n s2 * (n + redcarac s) <= 2 * Nat.log2 (S (cardinal s))s:treen:natH:rbt n s2 * n + redcarac s <= 2 * (n + redcarac s)s:treen:natH:rbt n s2 * n + redcarac s <= 2 * n + 2 * redcarac ss:treen:natH:rbt n sredcarac s <= 2 * redcarac ss:treen:natH:rbt n s1 * redcarac s <= 2 * redcarac sauto with arith.s:treen:natH:rbt n s1 <= 2s:treen:natH:rbt n s2 * (n + redcarac s) <= 2 * Nat.log2 (S (cardinal s))s:treen:natH:rbt n sn + redcarac s <= Nat.log2 (S (cardinal s))s:treen:natH:rbt n sn + redcarac s <= mindepth ss:treen:natH:rbt n smindepth s <= Nat.log2 (S (cardinal s))now apply rb_mindepth.s:treen:natH:rbt n sn + redcarac s <= mindepth sapply mindepth_log_cardinal. Qed.s:treen:natH:rbt n smindepth s <= Nat.log2 (S (cardinal s))s:trees <> Leaf -> Nat.log2 (cardinal s) < maxdepth sapply maxdepth_log_cardinal. Qed.s:trees <> Leaf -> Nat.log2 (cardinal s) < maxdepth s
x:eltRbt (singleton x)x:eltRbt (singleton x)exists 1; auto. Qed.x:eltRbt (Bk Leaf x Leaf)
n:natt:treearbt n t -> Rbt (makeBlack t)n:natt:treearbt n t -> Rbt (makeBlack t)n:natarbt n Leaf -> Rbt (makeBlack Leaf)n:natl:treex:X.tr:treearbt n (Rd l x r) -> Rbt (makeBlack (Rd l x r))n:natl:treex:X.tr:treearbt n (Bk l x r) -> Rbt (makeBlack (Bk l x r))exists 0; auto.n:natarbt n Leaf -> Rbt (makeBlack Leaf)destruct 1; invrb; exists (S n); simpl; auto.n:natl:treex:X.tr:treearbt n (Rd l x r) -> Rbt (makeBlack (Rd l x r))exists n; auto. Qed.n:natl:treex:X.tr:treearbt n (Bk l x r) -> Rbt (makeBlack (Bk l x r))t:treen:natrbt (S n) t -> notred t -> rrt n (makeRed t)destruct t as [|[|] l x r]; invrb; simpl; auto. Qed.t:treen:natrbt (S n) t -> notred t -> rrt n (makeRed t)
n:natl:treek:X.tr:treearbt n l -> rbt n r -> rbt (S n) (lbal l k r)case lbal_match; intros; desarb; invrb; auto. Qed.n:natl:treek:X.tr:treearbt n l -> rbt n r -> rbt (S n) (lbal l k r)n:natl:treek:X.tr:treerbt n l -> arbt n r -> rbt (S n) (rbal l k r)case rbal_match; intros; desarb; invrb; auto. Qed.n:natl:treek:X.tr:treerbt n l -> arbt n r -> rbt (S n) (rbal l k r)n:natl:treek:X.tr:treerbt n l -> arbt n r -> rbt (S n) (rbal' l k r)case rbal'_match; intros; desarb; invrb; auto. Qed.n:natl:treek:X.tr:treerbt n l -> arbt n r -> rbt (S n) (rbal' l k r)n:natl:treex:X.tr:treearbt n l -> rbt (S n) r -> notred r -> rbt (S n) (lbalS l x r)n:natl:treex:X.tr:treearbt n l -> rbt (S n) r -> notred r -> rbt (S n) (lbalS l x r)n:natl:treex:X.tr:treeHl:arbt n lHr:rbt (S n) rHr':notred rrbt (S n) (lbalS l x r)n:natl:treex:X.trl:treerx:X.trr:treeHl:arbt n lHr':notred (Bk rl rx rr)H2:rbt n rlH4:rbt n rrrbt (S n) (lbalS l x (Bk rl rx rr))n:natl:treex:X.trl:treerx:X.trr:treeHl:arbt n lH2:rbt n rlH4:rbt n rrrbt (S n) (lbalS l x (Bk rl rx rr))n:natl:treex:X.trl:treerx:X.trr:treeH2:rbt n rlH4:rbt n rrarbt n l -> rbt (S n) (lbalS l x (Bk rl rx rr))n:natl:treex:X.trl:treerx:X.trr:treeH2:rbt n rlH4:rbt n rrforall (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:natl:treex:X.trl:treerx:X.trr:treeH2:rbt n rlH4:rbt n rrforall t0 : tree, notred t0 -> arbt n t0 -> rbt (S n) (rbal' t0 x (Rd rl rx rr))destruct 1; invrb; auto.n:natl:treex:X.trl:treerx:X.trr:treeH2:rbt n rlH4:rbt n rrforall (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:natl:treex:X.trl:treerx:X.trr:treeH2:rbt n rlH4:rbt n rrforall t0 : tree, notred t0 -> arbt n t0 -> rbt (S n) (rbal' t0 x (Rd rl rx rr))apply rbal'_rb; auto. Qed.n:natl:treex:X.trl:treerx:X.trr:treeH2:rbt n rlH4:rbt n rrt0:treeH:notred t0Hl:arbt n t0rbt (S n) (rbal' t0 x (Rd rl rx rr))n:natl:treex:X.tr:treearbt n l -> rbt (S n) r -> arbt (S n) (lbalS l x r)n:natl:treex:X.tr:treearbt n l -> rbt (S n) r -> arbt (S n) (lbalS l x r)n:natl:treex:X.tr:treeforall (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:natl:treex:X.tr:treeforall 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 enddestruct 1; invrb; auto.n:natl:treex:X.tr:treeforall (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:natl:treex:X.tr:treeforall 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 endn:natx:X.tr:treeforall 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 endn:natx:X.tr, l:treeHl:notred lHl':arbt n lHr:rbt (S n) rarbt (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 endn:natx:X.trl:treerx:X.trr, l:treeHl:notred lHl':arbt n lH2:notred rlH4:notred rrH5:rbt (S n) rlH6:rbt (S n) rrarbt (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) endn:natx:X.trl:treerx:X.trr, l:treeHl:notred lHl':arbt n lH2:rbt n rlH4:rbt n rrarbt (S n) (rbal' l x (Rd rl rx rr))n:natx:X.trl:treerx:X.trr, l:treeHl:notred lHl':arbt n lH2:notred rlH4:notred rrH5:rbt (S n) rlH6:rbt (S n) rrarbt (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) endright; auto using rbal'_rb, makeRed_rr.n:natx:X.trll:treerlx:X.trlr:treerx:X.trr, l:treeHl:notred lHl':arbt n lH2:notred (Bk rll rlx rlr)H4:notred rrH6:rbt (S n) rrH3:rbt n rllH8:rbt n rlrarbt (S n) (Rd (Bk l x rll) rlx (rbal' rlr rx (makeRed rr)))left; apply rbal'_rb; auto. Qed.n:natx:X.trl:treerx:X.trr, l:treeHl:notred lHl':arbt n lH2:rbt n rlH4:rbt n rrarbt (S n) (rbal' l x (Rd rl rx rr))n:natl:treex:X.tr:treerbt (S n) l -> notred l -> arbt n r -> rbt (S n) (rbalS l x r)n:natl:treex:X.tr:treerbt (S n) l -> notred l -> arbt n r -> rbt (S n) (rbalS l x r)n:natl:treex:X.tr:treeHl:rbt (S n) lHl':notred lHr:arbt n rrbt (S n) (rbalS l x r)n:natll:treelx:X.tlr:treex:X.tr:treeHl':notred (Bk ll lx lr)Hr:arbt n rH2:rbt n llH4:rbt n lrrbt (S n) (rbalS (Bk ll lx lr) x r)n:natll:treelx:X.tlr:treex:X.tr:treeHr:arbt n rH2:rbt n llH4:rbt n lrrbt (S n) (rbalS (Bk ll lx lr) x r)n:natll:treelx:X.tlr:treex:X.tr:treeH2:rbt n llH4:rbt n lrarbt n r -> rbt (S n) (rbalS (Bk ll lx lr) x r)n:natll:treelx:X.tlr:treex:X.tr:treeH2:rbt n llH4:rbt n lrforall (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:natll:treelx:X.tlr:treex:X.tr:treeH2:rbt n llH4:rbt n lrforall t0 : tree, notred t0 -> arbt n t0 -> rbt (S n) (lbal (Rd ll lx lr) x t0)destruct 1; invrb; auto.n:natll:treelx:X.tlr:treex:X.tr:treeH2:rbt n llH4:rbt n lrforall (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:natll:treelx:X.tlr:treex:X.tr:treeH2:rbt n llH4:rbt n lrforall t0 : tree, notred t0 -> arbt n t0 -> rbt (S n) (lbal (Rd ll lx lr) x t0)apply lbal_rb; auto. Qed.n:natll:treelx:X.tlr:treex:X.tr:treeH2:rbt n llH4:rbt n lrt0:treeH:notred t0Hr:arbt n t0rbt (S n) (lbal (Rd ll lx lr) x t0)n:natl:treex:X.tr:treerbt (S n) l -> arbt n r -> arbt (S n) (rbalS l x r)n:natl:treex:X.tr:treerbt (S n) l -> arbt n r -> arbt (S n) (rbalS l x r)n:natl:treex:X.tr:treeforall (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:natl:treex:X.tr:treeforall 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 enddestruct 2; invrb; auto.n:natl:treex:X.tr:treeforall (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:natl:treex:X.tr:treeforall 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 endn:natl:treex:X.tforall 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 endn:natl:treex:X.tr:treeHr:notred rHr':rbt (S n) lHl:arbt n rarbt (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 endn:natll:treelx:X.tlr:treex:X.tr:treeHr:notred rHl:arbt n rH2:notred llH4:notred lrH5:rbt (S n) llH6:rbt (S n) lrarbt (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 endn:natll:treelx:X.tlr:treex:X.tr:treeHr:notred rHl:arbt n rH2:rbt n llH4:rbt n lrarbt (S n) (lbal (Rd ll lx lr) x r)n:natll:treelx:X.tlr:treex:X.tr:treeHr:notred rHl:arbt n rH2:notred llH4:notred lrH5:rbt (S n) llH6:rbt (S n) lrarbt (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 endright; auto using lbal_rb, makeRed_rr.n:natll:treelx:X.tlrl:treelrx:X.tlrr:treex:X.tr:treeHr:notred rHl:arbt n rH2:notred llH4:notred (Bk lrl lrx lrr)H5:rbt (S n) llH3:rbt n lrlH8:rbt n lrrarbt (S n) (Rd (lbal (makeRed ll) lx lrl) lrx (Bk lrr x r))left; apply lbal_rb; auto. Qed.n:natll:treelx:X.tlr:treex:X.tr:treeHr:notred rHl:arbt n rH2:rbt n llH4:rbt n lrarbt (S n) (lbal (Rd ll lx lr) x r)
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:treeA, B:Propnotred s -> ifred s A B <-> Bdestruct s; descolor; simpl; intuition. Qed.s:treeA, B:Propnotred s -> ifred s A B <-> Bs:treeA, B:Propifred s A B -> A \/ Bdestruct s; descolor; simpl; intuition. Qed.s:treeA, B:Propifred s A B -> A \/ Bx:X.ts:treen:natrbt n s -> ifred s (rrt n (ins x s)) (rbt n (ins x s))x:X.ts:treen:natrbt n s -> ifred s (rrt n (ins x s)) (rbt n (ins x s))x:X.tifred Leaf (rrt 0 (ins x Leaf)) (rbt 0 (ins x Leaf))x:X.tn:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1: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.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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)))simpl; auto.x:X.tifred Leaf (rrt 0 (ins x Leaf)) (rbt 0 (ins x Leaf))x:X.tn:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1: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.tn:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1: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) endelim_compare x k; auto.x:X.tn:natl:treek:X.tr:treen0:notred ln1:notred rH:rbt n lH0:rbt n rIHrbt1: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) endx:X.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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) endx:X.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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 krbt (S n) (Bk l k r)x:X.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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 krbt (S n) (lbal (ins x l) k r)x:X.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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 xrbt (S n) (rbal l k (ins x r))auto.x:X.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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 krbt (S n) (Bk l k r)x:X.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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 krbt (S n) (lbal (ins x l) k r)apply ifred_or in IHl; intuition.x:X.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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 karbt n (ins x l)x:X.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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 xrbt (S n) (rbal l k (ins x r))apply ifred_or in IHr; intuition. Qed.x:X.tn:natl:treek:X.tr:treeHl:rbt n lHr:rbt n rIHl: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 xarbt n (ins x r)x:X.ts:treen:natrbt n s -> arbt n (ins x s)x:X.ts:treen:natrbt n s -> arbt n (ins x s)x:X.ts:treen:natH:rbt n sarbt n (ins x s)intuition. Qed.x:X.ts:treen:natH:rrt n (ins x s) \/ rbt n (ins x s)arbt n (ins x s)x:X.ts:treeRbt s -> Rbt (add x s)x:X.ts:treeRbt s -> Rbt (add x s)x:X.ts:treen:natH:rbt n sRbt (add x s)now apply (makeBlack_rb n), ins_arb. Qed.x:X.ts:treen:natH:rbt n sRbt (makeBlack (ins x s))
A second approach here: statement with ... /\ ...
n:natl, r:treerbt n l -> rbt n r -> arbt n (append l r) /\ (notred l -> notred r -> rbt n (append l r))n:natl, r:treerbt n l -> rbt n r -> arbt n (append l r) /\ (notred l -> notred r -> rbt n (append l r))l:treeforall (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:treeforall n : nat, rbt n Leaf -> rbt n r -> arbt n r /\ (True -> notred r -> rbt n r)lc:Color.tll:treelx:X.tlr:treeIHlr: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:treelx:X.tlr, rl:treeIHlr: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.trr:treeforall 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:treelx:X.tlr, r:treeIHlr: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 rforall 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:treerx:X.trr, l:treeIHrl: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 lforall 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:treelx:X.tlr, rl:treeIHlr: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.trr:treeforall 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)))split; auto.r:treeforall n : nat, rbt n Leaf -> rbt n r -> arbt n r /\ (True -> notred r -> rbt n r)split; auto.lc:Color.tll:treelx:X.tlr:treeIHlr: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:treelx:X.tlr, rl:treeIHlr: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.trr:treeforall 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:treelx:X.tlr, rl:treeIHlr: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.trr:treen:natrbt 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:treelx:X.tlr, rl:treeIHlr: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.trr:treen:natH4:notred rlH6:notred rrH7:rbt n rlH8:rbt n rrH3:notred llH9:notred lrH10:rbt n llH11:rbt n lrarbt 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:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:notred rlH6:notred rrH7:rbt n rlH8:rbt n rrH3:notred llH9:notred lrH10:rbt n llH11:rbt n lrarbt 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:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:notred rlH6:notred rrH7:rbt n rlH8:rbt n rrH3:notred llH9:notred lrH10:rbt n llH11:rbt n lrforall (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:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:notred rlH6:notred rrH7:rbt n rlH8:rbt n rrH3:notred llH9:notred lrH10:rbt n llH11:rbt n lrforall 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:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:notred rlH6:notred rrH7:rbt n rlH8:rbt n rrH3:notred llH9:notred lrH10:rbt n llH11:rbt n lrforall (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:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:notred rlH6:notred rrH7:rbt n rlH8:rbt n rrH3:notred llH9:notred lrH10:rbt n llH11:rbt n lra:treex:X.tb:treeH:notred lr -> notred rl -> rbt n (Rd a x b)arbt n (Rd (Rd ll lx a) x (Rd b rx rr))ll:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:notred rlH6:notred rrH7:rbt n rlH8:rbt n rrH3:notred llH9:notred lrH10:rbt n llH11:rbt n lra:treex:X.tb:treeH: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))auto.ll:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:notred rlH6:notred rrH7:rbt n rlH8:rbt n rrH3:notred llH9:notred lrH10:rbt n llH11:rbt n lra:treex:X.tb:treeH:notred lr -> notred rl -> rbt n (Rd a x b)H12:notred aH14:notred bH15:rbt n aH16:rbt n barbt n (Rd (Rd ll lx a) x (Rd b rx rr))split; invrb; auto.ll:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:notred rlH6:notred rrH7:rbt n rlH8:rbt n rrH3:notred llH9:notred lrH10:rbt n llH11:rbt n lrforall 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:treelx:X.tlr, r:treeIHlr: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 rforall 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)))destruct (IHlr n) as (_,IH); auto.ll:treelx:X.tlr, r:treeIHlr: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 rn:natH0:rbt n rH4:notred llH6:notred lrH7:rbt n llH8:rbt n lrarbt n (Rd ll lx (append lr r))rl:treerx:X.trr, l:treeIHrl: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 lforall 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))destruct (IHrl n) as (_,IH); auto.rl:treerx:X.trr, l:treeIHrl: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 ln:natH:rbt n lH4:notred rlH6:notred rrH7:rbt n rlH8:rbt n rrarbt n (Rd (append l rl) rx rr)ll:treelx:X.tlr, rl:treeIHlr: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.trr:treeforall 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:treelx:X.tlr, rl:treeIHlr: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.trr:treen:natrbt (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:treelx:X.tlr, rl:treeIHlr: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.trr:treen:natH4:rbt n rlH6:rbt n rrH3:rbt n llH7:rbt n lrarbt (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:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:rbt n rlH6:rbt n rrH3:rbt n llH7:rbt n lrIH: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:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:rbt n rlH6:rbt n rrH3:rbt n llH7:rbt n lrarbt 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:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:rbt n rlH6:rbt n rrH3:rbt n llH7:rbt n lrforall (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:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:rbt n rlH6:rbt n rrH3:rbt n llH7:rbt n lrforall 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)))intros a x b IH; split; destruct IH; invrb; auto.ll:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:rbt n rlH6:rbt n rrH3:rbt n llH7:rbt n lrforall (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)))split; [left | invrb]; auto using lbalS_rb. Qed.ll:treelx:X.tlr, rl:treerx:X.trr:treen:natH4:rbt n rlH6:rbt n rrH3:rbt n llH7:rbt n lrforall 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)))
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:treex:X.tn:natrbt (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:treex:X.tn:natrbt 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:treex:X.tn:natrbt (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:treex:X.tn:natrbt 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:treex:X.tn:natrbt (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:treex:X.tforall 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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.eq x x'H5:rbt n lH8:rbt n rarbt 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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x x'H5:rbt n lH8:rbt n rarbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r enddel_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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x' xH5:rbt n lH8:rbt n rarbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) endapply 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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.eq x x'H5:rbt n lH8:rbt n rarbt 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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x x'H5:rbt n lH8:rbt n rarbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r enddel_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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x x'H5:rbt n lH8:rbt n rIHl':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 endl:treex':X.tr:treex:X.tIHl:forall n0 : nat, rbt (S n0) l -> isblack l -> arbt n0 (del x l)n:natH0:TrueH1:X.lt x x'H5:rbt n lH8:rbt n rIHl':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 endll:treelx:X.tlr:treex':X.tr:treex:X.tIHl:forall n0 : nat, rbt (S n0) (Bk ll lx lr) -> isblack (Bk ll lx lr) -> arbt n0 (del x (Bk ll lx lr))n:natH0:TrueH1:X.lt x x'H5:rbt n (Bk ll lx lr)H8:rbt n rIHl':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)apply lbalS_arb; auto.ll:treelx:X.tlr:treex':X.tr:treex:X.tIHl:forall n0 : nat, rbt (S n0) (Bk ll lx lr) -> isblack (Bk ll lx lr) -> arbt n0 (del x (Bk ll lx lr))n:natH0:TrueH1:X.lt x x'H5:rbt (S n) (Bk ll lx lr)H8:rbt (S n) rIHl':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)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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x' xH5:rbt n lH8:rbt n rarbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) enddel_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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x' xH5:rbt n lH8:rbt n rIHr':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) endl:treex':X.tr:treex:X.tIHr:forall n0 : nat, rbt (S n0) r -> isblack r -> arbt n0 (del x r)n:natH0:TrueH1:X.lt x' xH5:rbt n lH8:rbt n rIHr':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) endl:treex':X.trl:treerx:X.trr:treex:X.tIHr:forall n0 : nat, rbt (S n0) (Bk rl rx rr) -> isblack (Bk rl rx rr) -> arbt n0 (del x (Bk rl rx rr))n:natH0:TrueH1:X.lt x' xH5:rbt n lH8: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)))apply rbalS_arb; auto.l:treex':X.trl:treerx:X.trr:treex:X.tIHr:forall n0 : nat, rbt (S n0) (Bk rl rx rr) -> isblack (Bk rl rx rr) -> arbt n0 (del x (Bk rl rx rr))n:natH0:TrueH1:X.lt x' xH5:rbt (S n) lH8: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)))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:treex:X.tn:natrbt 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:treex:X.tn:natrbt 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:treex:X.tforall 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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.eq x x'H7:notred lH8:notred rH9:rbt n lH10:rbt n rrbt 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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x x'H7:notred lH8:notred rH9:rbt n lH10:rbt n rrbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r enddel_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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x' xH7:notred lH8:notred rH9:rbt n lH10:rbt n rrbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) endapply 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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.eq x x'H7:notred lH8:notred rH9:rbt n lH10:rbt n rrbt 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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x x'H7:notred lH8:notred rH9:rbt n lH10:rbt n rrbt n match l with | Bk _ _ _ => lbalS (del x l) x' r | _ => Rd (del x l) x' r enddel_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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x x'H7:notred lH8:notred rH9:rbt n lH10:rbt n rIHl':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 endl:treex':X.tr:treex:X.tIHl:forall n0 : nat, rbt n0 l -> notblack l -> rbt n0 (del x l)n:natH0:TrueH1:X.lt x x'H7:notred lH8:notred rH9:rbt n lH10:rbt n rIHl':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 endll:treelx:X.tlr:treex':X.tr:treex:X.tIHl:forall n0 : nat, rbt n0 (Bk ll lx lr) -> notblack (Bk ll lx lr) -> rbt n0 (del x (Bk ll lx lr))n:natH0:TrueH1:X.lt x x'H7:notred (Bk ll lx lr)H8:notred rH9:rbt n (Bk ll lx lr)H10:rbt n rIHl':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)destruct n as [|n]; [invrb|]; apply lbalS_rb; auto.ll:treelx:X.tlr:treex':X.tr:treex:X.tIHl:forall n0 : nat, rbt n0 (Bk ll lx lr) -> notblack (Bk ll lx lr) -> rbt n0 (del x (Bk ll lx lr))n:natH0:TrueH1:X.lt x x'H7:notred (Bk ll lx lr)H8:notred rH9:rbt (S n) (Bk ll lx lr)H10:rbt (S n) rIHl':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)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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x' xH7:notred lH8:notred rH9:rbt n lH10:rbt n rrbt n match r with | Bk _ _ _ => rbalS l x' (del x r) | _ => Rd l x' (del x r) enddel_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:treex':X.tr:treex:X.tIHl: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:natH0:TrueH1:X.lt x' xH7:notred lH8:notred rH9:rbt n lH10:rbt n rIHr':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) endl:treex':X.tr:treex:X.tIHr:forall n0 : nat, rbt n0 r -> notblack r -> rbt n0 (del x r)n:natH0:TrueH1:X.lt x' xH7:notred lH8:notred rH9:rbt n lH10:rbt n rIHr':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) endl:treex':X.trl:treerx:X.trr:treex:X.tIHr:forall n0 : nat, rbt n0 (Bk rl rx rr) -> notblack (Bk rl rx rr) -> rbt n0 (del x (Bk rl rx rr))n:natH0:TrueH1:X.lt x' xH7:notred lH8:notred (Bk rl rx rr)H9:rbt n lH10: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)))apply rbalS_rb; auto. } Qed.l:treex':X.trl:treerx:X.trr:treex:X.tIHr:forall n0 : nat, rbt n0 (Bk rl rx rr) -> notblack (Bk rl rx rr) -> rbt n0 (del x (Bk rl rx rr))n:natH0:TrueH1:X.lt x' xH7:notred lH8:notred (Bk rl rx rr)H9:rbt (S n) lH10: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)))s:treex:X.tRbt s -> Rbt (remove x s)s:treex:X.tRbt s -> Rbt (remove x s)s:treex:X.tn:natH:rbt n sRbt (remove x s)s:treex:X.tn:natH:rbt n sRbt (makeBlack (del x s))x:X.tn:natH:rbt n LeafRbt (makeBlack (del x Leaf))l:treey:X.tr:treex:X.tn:natH:rbt n (Rd l y r)Rbt (makeBlack (del x (Rd l y r)))l:treey:X.tr:treex:X.tn:natH:rbt n (Bk l y r)Rbt (makeBlack (del x (Bk l y r)))x:X.tn:natH:rbt n LeafRbt (makeBlack (del x Leaf))auto.x:X.tn:natH:rbt n Leafarbt n (del x Leaf)l:treey:X.tr:treex:X.tn:natH:rbt n (Rd l y r)Rbt (makeBlack (del x (Rd l y r)))l:treey:X.tr:treex:X.tn:natH:rbt n (Rd l y r)arbt n (del x (Rd l y r))apply del_rb; simpl; auto.l:treey:X.tr:treex:X.tn:natH:rbt n (Rd l y r)rbt n (del x (Rd l y r))l:treey:X.tr:treex:X.tn:natH:rbt n (Bk l y r)Rbt (makeBlack (del x (Bk l y r)))l:treey:X.tr:treex:X.tn:natH:rbt (S n) (Bk l y r)Rbt (makeBlack (del x (Bk l y r)))apply del_arb; simpl; auto. Qed.l:treey:X.tr:treex:X.tn:natH:rbt (S n) (Bk l y r)arbt n (del x (Bk l y r))
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_zerointros acc _; simpl; auto. Qed.treeify_rb_invariant 0 0 treeify_zerotreeify_rb_invariant 1 0 treeify_oneintros [|x acc]; simpl; auto; inversion 1. Qed.treeify_rb_invariant 1 0 treeify_onef, g:treeify_tsize1, size2, size, d:nattreeify_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_tsize1, size2, size, d:nattreeify_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_tsize1, size2, size, d:natHf:treeify_rb_invariant size1 d fHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)acc:list eltHacc:size <= length accrbt (S d) (fst (treeify_cont f g acc)) /\ size + length (snd (treeify_cont f g acc)) = length accf, g:treeify_tsize1, size2, size, d:natHf:treeify_rb_invariant size1 d fHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)acc:list eltHacc:size <= length accrbt (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 accf, g:treeify_tsize1, size2, size, d:natacc:list eltHf:size1 <= length acc -> rbt d (fst (f acc)) /\ size1 + length (snd (f acc)) = length accHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accrbt (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 accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeacc1:list eltHf:size1 <= length acc -> rbt d (fst (l, acc1)) /\ size1 + length (snd (l, acc1)) = length accHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accrbt (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 accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeacc1:list eltHf:size1 <= length acc -> rbt d l /\ size1 + length acc1 = length accHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accrbt (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 accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeacc1:list eltHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accsize1 <= length accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeacc1:list eltHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + length acc1 = length accrbt (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 accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeacc1:list eltHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accsize1 <= length acceauto with arith.f, g:treeify_tsize1, size2, d:natacc:list eltl:treeacc1:list eltHg:treeify_rb_invariant size2 d gHacc:S (size1 + size2) <= length accsize1 <= length accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeacc1:list eltHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + length acc1 = length accrbt (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 accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + 0 = length accrbt (S d) Leaf /\ size + 0 = length accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + S (length acc2) = length accrbt (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 accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + 0 = length accrbt (S d) Leaf /\ size + 0 = length accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + 0 = length accFalsef, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hf1:rbt d lHf2:size1 + 0 = length accsize <= length acc -> Falsef, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hf1:rbt d lHf2:size1 + 0 = length acclength acc < sizeauto with arith.f, g:treeify_tsize1, size2, size, d:natacc:list eltl:treeHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hf1:rbt d lHf2:size1 + 0 = length accsize1 + 0 < S (size1 + size2)f, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltHg:treeify_rb_invariant size2 d gH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + S (length acc2) = length accrbt (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 accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltHg:size2 <= length acc2 -> rbt d (fst (g acc2)) /\ size2 + length (snd (g acc2)) = length acc2H:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + S (length acc2) = length accrbt (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 accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltr:treeacc3:list eltHg:size2 <= length acc2 -> rbt d (fst (r, acc3)) /\ size2 + length (snd (r, acc3)) = length acc2H:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + S (length acc2) = length accrbt (S d) (fst (Bk l x r, acc3)) /\ size + length (snd (Bk l x r, acc3)) = length accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltr:treeacc3:list eltHg:size2 <= length acc2 -> rbt d r /\ size2 + length acc3 = length acc2H:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + S (length acc2) = length accrbt (S d) (Bk l x r) /\ size + length acc3 = length accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltr:treeacc3:list eltH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + S (length acc2) = length accsize2 <= length acc2f, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltr:treeacc3:list eltH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + S (length acc2) = length accHg1:rbt d rHg2:size2 + length acc3 = length acc2rbt (S d) (Bk l x r) /\ size + length acc3 = length accf, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltr:treeacc3:list eltH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + S (length acc2) = length accsize2 <= length acc2f, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltr:treeacc3:list eltH:size = S (size1 + size2)Hf1:rbt d lHf2:size1 + S (length acc2) = length accsize <= length acc -> size2 <= length acc2apply Nat.add_le_mono_l.f, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltr:treeacc3:list eltH:size = S (size1 + size2)Hf1:rbt d lHf2:size1 + S (length acc2) = length accsize1 + size2 <= size1 + length acc2 -> size2 <= length acc2f, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltr:treeacc3:list eltH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + S (length acc2) = length accHg1:rbt d rHg2:size2 + length acc3 = length acc2rbt (S d) (Bk l x r) /\ size + length acc3 = length accnow rewrite H, <- Hf2, <- Hg2, Nat.add_succ_r, Nat.add_assoc. Qed.f, g:treeify_tsize1, size2, size, d:natacc:list eltl:treex:eltacc2:list eltr:treeacc3:list eltH:size = S (size1 + size2)Hacc:size <= length accHf1:rbt d lHf2:size1 + S (length acc2) = length accHg1:rbt d rHg2:size2 + length acc3 = length acc2size + length acc3 = length accn:positiveexists d : nat, forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)n:positiveexists d : nat, forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)n:positived:natIHn: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:positived:natIHn: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:positived:natIHn: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:positived:natIHn: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:positived:natIHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)b:booltreeify_rb_invariant (ifpred b (Pos.to_nat n~1)) (S d) (treeify_aux b n~1)n:positived:natIHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)b:boolifpred b (Pos.to_nat n~1) = S (Pos.to_nat n + ifpred b (Pos.to_nat n))n:positived:natIHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)b:boolifpred b (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred b (Pos.to_nat n))n:positived:natIHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)b:boolH:0 < Pos.to_nat nifpred b (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred b (Pos.to_nat n))n:positived:natIHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)b:boolH:Pos.to_nat n <> 0ifpred b (S (2 * Pos.to_nat n)) = S (Pos.to_nat n + ifpred b (Pos.to_nat n))now rewrite <- Nat.add_succ_r, Nat.succ_pred; trivial.n:positived:natIHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)H:Pos.to_nat n <> 0Pos.to_nat n + Pos.to_nat n = S (Pos.to_nat n + Init.Nat.pred (Pos.to_nat n))n:positived:natIHn: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:positived:natIHn: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:positived:natIHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)b:booltreeify_rb_invariant (ifpred b (Pos.to_nat n~0)) (S d) (treeify_aux b n~0)n:positived:natIHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)b:boolifpred b (Pos.to_nat n~0) = S (ifpred b (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))n:positived:natIHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)b:boolifpred b (2 * Pos.to_nat n) = S (ifpred b (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))n:positived:natIHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)b:boolH:0 < Pos.to_nat nifpred b (2 * Pos.to_nat n) = S (ifpred b (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))n:positived:natIHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)b:boolH:Pos.to_nat n <> 0ifpred b (2 * Pos.to_nat n) = S (ifpred b (Pos.to_nat n) + Init.Nat.pred (Pos.to_nat n))n:positived:natIHn:forall b0 : bool, treeify_rb_invariant (ifpred b0 (Pos.to_nat n)) d (treeify_aux b0 n)b:boolH:Pos.to_nat n <> 0ifpred b (2 * Pos.to_nat n) = ifpred b (Pos.to_nat n) + Pos.to_nat nn:positived:natIHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)H:Pos.to_nat n <> 0Init.Nat.pred (Pos.to_nat n + Pos.to_nat n) = Init.Nat.pred (Pos.to_nat n) + Pos.to_nat nnow apply Nat.add_pred_l.n:positived:natIHn:forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat n)) d (treeify_aux b n)H:Pos.to_nat n <> 0Init.Nat.pred (Pos.to_nat n) + Pos.to_nat n = Init.Nat.pred (Pos.to_nat n + Pos.to_nat n)exists 0; destruct b; [ apply treeify_zero_rb | apply treeify_one_rb ]. Qed.exists d : nat, forall b : bool, treeify_rb_invariant (ifpred b (Pos.to_nat 1)) d (treeify_aux b 1)
The black depth of treeify l is actually a log2, but
we don't need to mention that.
l:list eltRbt (treeify l)l:list eltRbt (treeify l)l:list eltRbt (fst (treeify_aux true (plength l) l))l:list eltd:natH: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 eltd:natH: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))now rewrite plength_spec. Qed.l:list eltd:natH: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
f:elt -> bools:tRbt (filter f s)unfold filter; auto_tc. Qed.f:elt -> bools:tRbt (filter f s)f:elt -> bools:tRbt (fst (partition f s))f:elt -> bools:tRbt (fst (partition f s))f:elt -> bools:tRbt (fst (let (ok0, ko) := partition_aux f s nil nil in (treeify ok0, treeify ko)))f:elt -> bools:tl, l0:list X.tRbt (fst (treeify l, treeify l0))auto_tc. Qed.f:elt -> bools:tl, l0:list X.tRbt (treeify l)f:elt -> bools:tRbt (snd (partition f s))f:elt -> bools:tRbt (snd (partition f s))f:elt -> bools:tRbt (snd (let (ok0, ko) := partition_aux f s nil nil in (treeify ok0, treeify ko)))f:elt -> bools:tl, l0:list X.tRbt (snd (treeify l, treeify l0))auto_tc. Qed.f:elt -> bools:tl, l0:list X.tRbt (treeify l0)
s1, s2:treeRbt s2 -> Rbt (fold add s1 s2)s1, s2:treeRbt s2 -> Rbt (fold add s1 s2)s1, s2:treeH:Rbt s2Rbt (fold add s1 s2)s1, s2:treeH:Rbt s2Rbt (fold_right (fun (y : elt) (x : tree) => flip add x y) s2 (rev (elements s1)))induction (rev (elements s1)); simpl; unfold flip in *; auto_tc. Qed.s1, s2:treeH:Rbt s2Rbt (fold_right (fun (y : X.t) (x : tree) => flip add x y) s2 (rev (elements s1)))s1, s2:treeRbt s2 -> Rbt (fold remove s1 s2)s1, s2:treeRbt s2 -> Rbt (fold remove s1 s2)s1, s2:treeH:Rbt s2Rbt (fold remove s1 s2)s1, s2:treeH:Rbt s2Rbt (fold_right (fun (y : elt) (x : tree) => flip remove x y) s2 (rev (elements s1)))induction (rev (elements s1)); simpl; unfold flip in *; auto_tc. Qed.s1, s2:treeH:Rbt s2Rbt (fold_right (fun (y : X.t) (x : tree) => flip remove x y) s2 (rev (elements s1)))s1, s2:treeRbt s1 -> Rbt s2 -> Rbt (union s1 s2)s1, s2:treeRbt s1 -> Rbt s2 -> Rbt (union s1 s2)s1, s2:treeH:Rbt s1H0:Rbt s2Rbt (union s1 s2)destruct compare_height; auto_tc. Qed.s1, s2:treeH:Rbt s1H0:Rbt s2Rbt 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 ends1, s2:treeRbt s1 -> Rbt s2 -> Rbt (inter s1 s2)s1, s2:treeRbt s1 -> Rbt s2 -> Rbt (inter s1 s2)s1, s2:treeH:Rbt s1H0:Rbt s2Rbt (inter s1 s2)destruct compare_height; auto_tc. Qed.s1, s2:treeH:Rbt s1H0:Rbt s2Rbt 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 ends1, s2:treeRbt s1 -> Rbt s2 -> Rbt (diff s1 s2)s1, s2:treeRbt s1 -> Rbt s2 -> Rbt (diff s1 s2)s1, s2:treeH:Rbt s1H0:Rbt s2Rbt (diff s1 s2)destruct compare_height; auto_tc. Qed. End BalanceProps.s1, s2:treeH:Rbt s1H0:Rbt s2Rbt 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
Final Encapsulation
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:elts':tremove_min s = Some (x, s') -> min_elt s = Some x /\ Equal (remove x s) s's:t_x:elts':tremove_min s = Some (x, s') -> min_elt s = Some x /\ Equal (remove x s) s's:Raw.tHs:Raw.Ok sx:elts':tremove_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.tHs:Raw.Ok sx:elts':tmatch 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.tHs:Raw.Ok sx:elts':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.tHs:Raw.Ok sx:elts':tP:=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.tHs:Raw.Ok sx:elts':tP: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.tHs:Raw.Ok sx:elts':tx0:Raw.elts0:Raw.treeP: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.tHs:Raw.Ok sx:elts':tx0:Raw.elts0:Raw.treeP:Raw.Ok s0H: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'0U: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.tHs:Raw.Ok sx:elts0:Raw.treeP:Raw.Ok s0H: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.tHs:Raw.Ok sx:elts0:Raw.treeP:Raw.Ok s0H: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)subst; intuition. Qed.s:Raw.tHs:Raw.Ok sx:elts0:Raw.treeP:Raw.Ok s0H: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 xH1:Raw.remove x s = s0Raw.min_elt s = Some x /\ (forall a : elt, Raw.In a (Raw.remove x s) <-> Raw.In a s0)s:t_remove_min s = None -> Empty ss:t_remove_min s = None -> Empty ss:Raw.tHs:Raw.Ok sremove_min {| this := s; is_ok := Hs |} = None -> Empty {| this := s; is_ok := Hs |}s:Raw.tHs:Raw.Ok smatch 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 ss:Raw.tHs: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 ss:Raw.tHs:Raw.Ok sP:=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 sdestruct (Raw.remove_min s) as [(x0,s0)|]; now intuition. Qed. End Make.s:Raw.tHs:Raw.Ok sP: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