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) *)
(************************************************************************)
Efficient implementation of FSetInterface.S for positive keys,
inspired from the FMapPositive module.
This module was adapted by Alexandre Ren, Damien Pous, and Thomas
Braibant (2010, LIG, CNRS, UMR 5217), from the FMapPositive
module of Pierre Letouzey and Jean-Christophe Filliâtre, which in
turn comes from the FMap framework of a work by Xavier Leroy and
Sandrine Blazy (used for building certified compilers).
Require Import Bool BinPos OrderedType OrderedTypeEx FSetInterface. Set Implicit Arguments. Local Open Scope lazy_bool_scope. Local Open Scope positive_scope. Local Unset Elimination Schemes. Module PositiveSet <: S with Module E:=PositiveOrderedTypeBits. Module E:=PositiveOrderedTypeBits. Definition elt := positive : Type. Inductive tree := | Leaf : tree | Node : tree -> bool -> tree -> tree. Scheme tree_ind := Induction for tree Sort Prop. Definition t := tree : Type. Definition empty : t := Leaf. Fixpoint is_empty (m : t) : bool := match m with | Leaf => true | Node l b r => negb b &&& is_empty l &&& is_empty r end. Fixpoint mem (i : elt) (m : t) {struct m} : bool := match m with | Leaf => false | Node l o r => match i with | 1 => o | i~0 => mem i l | i~1 => mem i r end end. Fixpoint add (i : elt) (m : t) : t := match m with | Leaf => match i with | 1 => Node Leaf true Leaf | i~0 => Node (add i Leaf) false Leaf | i~1 => Node Leaf false (add i Leaf) end | Node l o r => match i with | 1 => Node l true r | i~0 => Node (add i l) o r | i~1 => Node l o (add i r) end end. Definition singleton i := add i empty.
helper function to avoid creating empty trees that are not leaves
Definition node (l : t) (b: bool) (r : t) : t := if b then Node l b r else match l,r with | Leaf,Leaf => Leaf | _,_ => Node l false r end. Fixpoint remove (i : elt) (m : t) {struct m} : t := match m with | Leaf => Leaf | Node l o r => match i with | 1 => node l false r | i~0 => node (remove i l) o r | i~1 => node l o (remove i r) end end. Fixpoint union (m m': t) : t := match m with | Leaf => m' | Node l o r => match m' with | Leaf => m | Node l' o' r' => Node (union l l') (o||o') (union r r') end end. Fixpoint inter (m m': t) : t := match m with | Leaf => Leaf | Node l o r => match m' with | Leaf => Leaf | Node l' o' r' => node (inter l l') (o&&o') (inter r r') end end. Fixpoint diff (m m': t) : t := match m with | Leaf => Leaf | Node l o r => match m' with | Leaf => m | Node l' o' r' => node (diff l l') (o&&negb o') (diff r r') end end. Fixpoint equal (m m': t): bool := match m with | Leaf => is_empty m' | Node l o r => match m' with | Leaf => is_empty m | Node l' o' r' => eqb o o' &&& equal l l' &&& equal r r' end end. Fixpoint subset (m m': t): bool := match m with | Leaf => true | Node l o r => match m' with | Leaf => is_empty m | Node l' o' r' => (negb o ||| o') &&& subset l l' &&& subset r r' end end.
reverses y and concatenate it with x
Fixpoint rev_append (y x : elt) : elt := match y with | 1 => x | y~1 => rev_append y x~1 | y~0 => rev_append y x~0 end. Infix "@" := rev_append (at level 60). Definition rev x := x@1. Section Fold. Variable B : Type. Variable f : elt -> B -> B.
the additional argument, i, records the current path, in
reverse order (this should be more efficient: we reverse this argument
only at present nodes only, rather than at each node of the tree).
we also use this convention in all functions below
Fixpoint xfold (m : t) (v : B) (i : elt) := match m with | Leaf => v | Node l true r => xfold r (f (rev i) (xfold l v i~0)) i~1 | Node l false r => xfold r (xfold l v i~0) i~1 end. Definition fold m i := xfold m i 1. End Fold. Section Quantifiers. Variable f : elt -> bool. Fixpoint xforall (m : t) (i : elt) := match m with | Leaf => true | Node l o r => (negb o ||| f (rev i)) &&& xforall r i~1 &&& xforall l i~0 end. Definition for_all m := xforall m 1. Fixpoint xexists (m : t) (i : elt) := match m with | Leaf => false | Node l o r => (o &&& f (rev i)) ||| xexists r i~1 ||| xexists l i~0 end. Definition exists_ m := xexists m 1. Fixpoint xfilter (m : t) (i : elt) : t := match m with | Leaf => Leaf | Node l o r => node (xfilter l i~0) (o &&& f (rev i)) (xfilter r i~1) end. Definition filter m := xfilter m 1. Fixpoint xpartition (m : t) (i : elt) : t * t := match m with | Leaf => (Leaf,Leaf) | Node l o r => let (lt,lf) := xpartition l i~0 in let (rt,rf) := xpartition r i~1 in if o then let fi := f (rev i) in (node lt fi rt, node lf (negb fi) rf) else (node lt false rt, node lf false rf) end. Definition partition m := xpartition m 1. End Quantifiers.
uses a to accumulate values rather than doing a lot of concatenations
Fixpoint xelements (m : t) (i : elt) (a: list elt) := match m with | Leaf => a | Node l false r => xelements l i~0 (xelements r i~1 a) | Node l true r => xelements l i~0 (rev i :: xelements r i~1 a) end. Definition elements (m : t) := xelements m 1 nil. Fixpoint cardinal (m : t) : nat := match m with | Leaf => O | Node l false r => (cardinal l + cardinal r)%nat | Node l true r => S (cardinal l + cardinal r) end. Definition omap (f: elt -> elt) x := match x with | None => None | Some i => Some (f i) end.
would it be more efficient to use a path like in the above functions ?
Fixpoint choose (m: t) : option elt := match m with | Leaf => None | Node l o r => if o then Some 1 else match choose l with | None => omap xI (choose r) | Some i => Some i~0 end end. Fixpoint min_elt (m: t) : option elt := match m with | Leaf => None | Node l o r => match min_elt l with | None => if o then Some 1 else omap xI (min_elt r) | Some i => Some i~0 end end. Fixpoint max_elt (m: t) : option elt := match m with | Leaf => None | Node l o r => match max_elt r with | None => if o then Some 1 else omap xO (max_elt l) | Some i => Some i~1 end end.
lexicographic product, defined using a notation to keep things lazy
Notation lex u v := match u with Eq => v | Lt => Lt | Gt => Gt end. Definition compare_bool a b := match a,b with | false, true => Lt | true, false => Gt | _,_ => Eq end. Fixpoint compare_fun (m m': t): comparison := match m,m' with | Leaf,_ => if is_empty m' then Eq else Lt | _,Leaf => if is_empty m then Eq else Gt | Node l o r,Node l' o' r' => lex (compare_bool o o') (lex (compare_fun l l') (compare_fun r r')) end. Definition In i t := mem i t = true. Definition Equal s s' := forall a : elt, In a s <-> In a s'. Definition Subset s s' := forall a : elt, In a s -> In a s'. Definition Empty s := forall a : elt, ~ In a s. Definition For_all (P : elt -> Prop) s := forall x, In x s -> P x. Definition Exists (P : elt -> Prop) s := exists x, In x s /\ P x. Notation "s [=] t" := (Equal s t) (at level 70, no associativity). Notation "s [<=] t" := (Subset s t) (at level 70, no associativity). Definition eq := Equal. Declare Equivalent Keys Equal eq. Definition lt m m' := compare_fun m m' = Lt.
Specification of In
forall (s : t) (x y : positive), E.eq x y -> In x s -> In y sforall (s : t) (x y : positive), E.eq x y -> In x s -> In y strivial. Qed.s:ty:positiveIn y s -> In y s
Specification of eq
forall s : t, eq s sforall s : t, eq s sreflexivity. Qed.forall (s : t) (a : elt), In a s <-> In a sforall s s' : t, eq s s' -> eq s' sforall s s' : t, eq s s' -> eq s' sforall s s' : t, (forall a : elt, In a s <-> In a s') -> forall a : elt, In a s' <-> In a ss, s':tH:forall a0 : elt, In a0 s <-> In a0 s'a:eltIn a s' <-> In a strivial. Qed.s, s':tH:forall a0 : elt, In a0 s <-> In a0 s'a:eltIn a s <-> In a s'forall s s' s'' : t, eq s s' -> eq s' s'' -> eq s s''forall s s' s'' : t, eq s s' -> eq s' s'' -> eq s s''forall s s' s'' : t, (forall a : elt, In a s <-> In a s') -> (forall a : elt, In a s' <-> In a s'') -> forall a : elt, In a s <-> In a s''s, s', s'':tH:forall a0 : elt, In a0 s <-> In a0 s'H0:forall a0 : elt, In a0 s' <-> In a0 s''a:eltIn a s <-> In a s''trivial. Qed.s, s', s'':tH:forall a0 : elt, In a0 s <-> In a0 s'H0:forall a0 : elt, In a0 s' <-> In a0 s''a:eltIn a s' <-> In a s''
Specification of mem
forall (s : t) (x : elt), In x s -> mem x s = trueforall (s : t) (x : elt), In x s -> mem x s = truetrivial. Qed.forall (s : t) (x : elt), mem x s = true -> mem x s = trueforall (s : t) (x : elt), mem x s = true -> In x sforall (s : t) (x : elt), mem x s = true -> In x strivial. Qed.forall (s : t) (x : elt), mem x s = true -> mem x s = true
Additional lemmas for mem
forall x : elt, mem x Leaf = falsedestruct x; trivial. Qed.forall x : elt, mem x Leaf = false
Specification of empty
Empty emptyEmpty emptyforall a : elt, mem a empty <> truea:eltmem a empty <> truediscriminate. Qed.a:eltfalse <> true
Specification of node
forall (x : elt) (l : t) (o : bool) (r : t), mem x (node l o r) = mem x (Node l o r)forall (x : elt) (l : t) (o : bool) (r : t), mem x (node l o r) = mem x (Node l o r)x:eltl:to:boolr:tmem x (node l o r) = mem x (Node l o r)x:eltl:to:boolr:tmem x (node l false r) = mem x (Node l false r)x:elto:boolr:tmem x (node Leaf false r) = mem x (Node Leaf false r)now destruct x. Qed. Local Opaque node.x:elto:boolmem x (node Leaf false Leaf) = mem x (Node Leaf false Leaf)
Specification of is_empty
forall s : t, Empty s <-> is_empty s = trueforall s : t, Empty s <-> is_empty s = trueforall s : t, (forall a : elt, mem a s <> true) <-> is_empty s = true(elt -> false <> true) <-> true = truel:treeo:boolr:treeIHl:(forall a : elt, mem a l <> true) <-> is_empty l = trueIHr:(forall a : elt, mem a r <> true) <-> is_empty r = true(forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => o end <> true) <-> negb o &&& is_empty l &&& is_empty r = truel:treeo:boolr:treeIHl:(forall a : elt, mem a l <> true) <-> is_empty l = trueIHr:(forall a : elt, mem a r <> true) <-> is_empty r = true(forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => o end <> true) <-> negb o &&& is_empty l &&& is_empty r = truel:treeo:boolr:treeIHl:(forall a : elt, mem a l <> true) <-> is_empty l = trueIHr:(forall a : elt, mem a r <> true) <-> is_empty r = true(forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => o end <> true) <-> (negb o = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true)l:treeo:boolr:tree(forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => o end <> true) <-> (negb o = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true)l, r:tree(forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => true end <> true) -> (false = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true)l, r:tree(false = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => true end <> truel, r:tree(forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> true) -> (true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true)l, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truel, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => true end <> true(false = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true)l, r:tree(false = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => true end <> truel, r:tree(forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> true) -> (true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true)l, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truel, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => true end <> truetrue = truel, r:tree(false = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => true end <> truel, r:tree(forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> true) -> (true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true)l, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truel, r:tree(false = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => true end <> truel, r:tree(forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> true) -> (true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true)l, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truel, r:tree(forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> true) -> (true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true)l, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truel, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> true(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true)l, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truel, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truetrue = true /\ (forall a : elt, mem a l <> true)l, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> trueforall a : elt, mem a r <> truel, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truel, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truetrue = truel, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> trueforall a : elt, mem a l <> truel, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> trueforall a : elt, mem a r <> truel, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truel, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> trueforall a : elt, mem a l <> truel, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> trueforall a : elt, mem a r <> truel, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truel, r:treeH:forall a0 : elt, match a0 with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truea:eltmem a l <> truel, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> trueforall a : elt, mem a r <> truel, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truel, r:treeH:forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> trueforall a : elt, mem a r <> truel, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truel, r:treeH:forall a0 : elt, match a0 with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> truea:eltmem a r <> truel, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> trueintros H [a|a|]; apply H || intro; discriminate. Qed.l, r:tree(true = true /\ (forall a : elt, mem a l <> true)) /\ (forall a : elt, mem a r <> true) -> forall a : elt, match a with | i~1 => mem i r | i~0 => mem i l | 1 => false end <> trueforall s : t, Empty s -> is_empty s = trueforall s : t, Empty s -> is_empty s = trues:tEmpty s -> is_empty s = truetrivial. Qed.s:tis_empty s = true -> is_empty s = trueforall s : t, is_empty s = true -> Empty sforall s : t, is_empty s = true -> Empty ss:tis_empty s = true -> Empty strivial. Qed.s:tis_empty s = true -> is_empty s = true
Specification of subset
forall s : t, Leaf [<=] sforall s : t, Leaf [<=] selim (empty_1 Hi). Qed.s:ti:eltHi:In i LeafIn i sforall s s' : t, s [<=] s' <-> subset s s' = trueforall s s' : t, s [<=] s' <-> subset s s' = trueLeaf [<=] Leaf <-> true = truel':treeo':boolr':treeLeaf [<=] Node l' o' r' <-> true = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l o r [<=] Leaf <-> negb o &&& is_empty l &&& is_empty r = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = trueH:Leaf [<=] Leaftrue = trueH:true = trueLeaf [<=] Leafl':treeo':boolr':treeLeaf [<=] Node l' o' r' <-> true = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l o r [<=] Leaf <-> negb o &&& is_empty l &&& is_empty r = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = trueH:true = trueLeaf [<=] Leafl':treeo':boolr':treeLeaf [<=] Node l' o' r' <-> true = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l o r [<=] Leaf <-> negb o &&& is_empty l &&& is_empty r = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel':treeo':boolr':treeLeaf [<=] Node l' o' r' <-> true = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l o r [<=] Leaf <-> negb o &&& is_empty l &&& is_empty r = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel':treeo':boolr':treeH:Leaf [<=] Node l' o' r'true = truel':treeo':boolr':treeH:true = trueLeaf [<=] Node l' o' r'l:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l o r [<=] Leaf <-> negb o &&& is_empty l &&& is_empty r = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel':treeo':boolr':treeH:true = trueLeaf [<=] Node l' o' r'l:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l o r [<=] Leaf <-> negb o &&& is_empty l &&& is_empty r = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l o r [<=] Leaf <-> negb o &&& is_empty l &&& is_empty r = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l o r [<=] Leaf <-> (negb o = true /\ Empty l) /\ Empty rl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l true r [<=] Leaf <-> (false = true /\ Empty l) /\ Empty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l false r [<=] Leaf <-> (true = true /\ Empty l) /\ Empty rl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l true r [<=] Leaf -> (false = true /\ Empty l) /\ Empty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = true(false = true /\ Empty l) /\ Empty r -> Node l true r [<=] Leafl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l false r [<=] Leaf <-> (true = true /\ Empty l) /\ Empty rl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l true r [<=] Leaf(false = true /\ Empty l) /\ Empty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = true(false = true /\ Empty l) /\ Empty r -> Node l true r [<=] Leafl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l false r [<=] Leaf <-> (true = true /\ Empty l) /\ Empty rl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l true r [<=] LeafIn 1 emptyl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = true(false = true /\ Empty l) /\ Empty r -> Node l true r [<=] Leafl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l false r [<=] Leaf <-> (true = true /\ Empty l) /\ Empty rl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l true r [<=] LeafIn 1 (Node l true r)l, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = true(false = true /\ Empty l) /\ Empty r -> Node l true r [<=] Leafl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l false r [<=] Leaf <-> (true = true /\ Empty l) /\ Empty rl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = true(false = true /\ Empty l) /\ Empty r -> Node l true r [<=] Leafl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l false r [<=] Leaf <-> (true = true /\ Empty l) /\ Empty rl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueNode l false r [<=] Leaf <-> (true = true /\ Empty l) /\ Empty rl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] Leaf(true = true /\ Empty l) /\ Empty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] Leaftrue = true /\ Empty ll, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] LeafEmpty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] Leaftrue = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] LeafEmpty ll, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] LeafEmpty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] LeafEmpty ll, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] LeafEmpty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] Leafforall a : elt, ~ In a ll, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] LeafEmpty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] Leafa:eltH1:In a lFalsel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] LeafEmpty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] Leafa:eltH1:In a lIn a~0 emptyl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] LeafEmpty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] Leafa:eltH1:In a lIn a~0 (Node l false r)l, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] LeafEmpty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] LeafEmpty rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] Leafforall a : elt, ~ In a rl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] Leafa:eltH1:In a rFalsel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] Leafa:eltH1:In a rIn a~1 emptyl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:Node l false r [<=] Leafa:eltH1:In a rIn a~1 (Node l false r)l, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueH:(true = true /\ Empty l) /\ Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueHl:Empty lHr:Empty rNode l false r [<=] Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueHl:Empty lHr:Empty ri:positiveHi:In i~1 (Node l false r)In i~1 Leafl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueHl:Empty lHr:Empty ri:positiveHi:In i~0 (Node l false r)In i~0 Leafl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueHl:Empty lHr:Empty rHi:In 1 (Node l false r)In 1 Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueHl:Empty lHr:Empty ri:positiveHi:In i~0 (Node l false r)In i~0 Leafl, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueHl:Empty lHr:Empty rHi:In 1 (Node l false r)In 1 Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel, r:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = trueHl:Empty lHr:Empty rHi:In 1 (Node l false r)In 1 Leafl:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o') &&& subset l l' &&& subset r r' = truel:treeo:boolr:treeIHl:forall s' : t, l [<=] s' <-> subset l s' = trueIHr:forall s' : t, r [<=] s' <-> subset r s' = truel':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o' = true /\ l [<=] l') /\ r [<=] r'l:treeo:boolr, l':treeo':boolr':treeNode l o r [<=] Node l' o' r' <-> (negb o ||| o' = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeNode l true r [<=] Node l' o' r' <-> (o' = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'(o' = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'o' = true /\ l [<=] l'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'o' = truel, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'l [<=] l'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l', r':treeH:Node l true r [<=] Node l' false r'false = truel, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'l [<=] l'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l', r':treeH:In 1 (Node l true r) -> In 1 (Node l' false r')false = truel, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'l [<=] l'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l', r':treeH:mem 1 (Node l true r) = true -> mem 1 (Node l' false r') = truefalse = truel, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'l [<=] l'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l', r':treeH:true = true -> false = truefalse = truel, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'l [<=] l'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l', r':treeH:true = true -> false = truetrue = truel, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'l [<=] l'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'l [<=] l'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'i:eltHi:In i lIn i l'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'i:eltHi:In i lIn i~0 (Node l true r)l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'i:eltHi:In i rIn i r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:Node l true r [<=] Node l' o' r'i:eltHi:In i rIn i~1 (Node l true r)l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:(o' = true /\ l [<=] l') /\ r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'Node l true r [<=] Node l' o' r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'Node l true r [<=] Node l' true r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'i:eltHi:In i (Node l true r)In i (Node l' true r')l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~1 (Node l true r)In i~1 (Node l' true r')l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~0 (Node l true r)In i~0 (Node l' true r')l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'Hi:In 1 (Node l true r)In 1 (Node l' true r')l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~1 (Node l true r)In i rl, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~0 (Node l true r)In i~0 (Node l' true r')l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'Hi:In 1 (Node l true r)In 1 (Node l' true r')l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~0 (Node l true r)In i~0 (Node l' true r')l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'Hi:In 1 (Node l true r)In 1 (Node l' true r')l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~0 (Node l true r)In i ll, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'Hi:In 1 (Node l true r)In 1 (Node l' true r')l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeHo':o' = trueHl:l [<=] l'Hr:r [<=] r'Hi:In 1 (Node l true r)In 1 (Node l' true r')l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeNode l false r [<=] Node l' o' r' <-> (true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'(true = true /\ l [<=] l') /\ r [<=] r'l, r, l':treeo':boolr':treeH:(true = true /\ l [<=] l') /\ r [<=] r'Node l false r [<=] Node l' o' r'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'true = true /\ l [<=] l'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(true = true /\ l [<=] l') /\ r [<=] r'Node l false r [<=] Node l' o' r'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'true = truel, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'l [<=] l'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(true = true /\ l [<=] l') /\ r [<=] r'Node l false r [<=] Node l' o' r'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'l [<=] l'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(true = true /\ l [<=] l') /\ r [<=] r'Node l false r [<=] Node l' o' r'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'i:eltHi:In i lIn i l'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(true = true /\ l [<=] l') /\ r [<=] r'Node l false r [<=] Node l' o' r'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'i:eltHi:In i lIn i~0 (Node l false r)l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(true = true /\ l [<=] l') /\ r [<=] r'Node l false r [<=] Node l' o' r'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'r [<=] r'l, r, l':treeo':boolr':treeH:(true = true /\ l [<=] l') /\ r [<=] r'Node l false r [<=] Node l' o' r'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'i:eltHi:In i rIn i r'l, r, l':treeo':boolr':treeH:(true = true /\ l [<=] l') /\ r [<=] r'Node l false r [<=] Node l' o' r'l, r, l':treeo':boolr':treeH:Node l false r [<=] Node l' o' r'i:eltHi:In i rIn i~1 (Node l false r)l, r, l':treeo':boolr':treeH:(true = true /\ l [<=] l') /\ r [<=] r'Node l false r [<=] Node l' o' r'l, r, l':treeo':boolr':treeH:(true = true /\ l [<=] l') /\ r [<=] r'Node l false r [<=] Node l' o' r'l, r, l':treeo':boolr':treeH:(true = true /\ l [<=] l') /\ r [<=] r'i:eltHi:In i (Node l false r)In i (Node l' o' r')l, r, l':treeo':boolr':treeH:true = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~1 (Node l false r)In i~1 (Node l' o' r')l, r, l':treeo':boolr':treeH:true = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~0 (Node l false r)In i~0 (Node l' o' r')l, r, l':treeo':boolr':treeH:true = trueHl:l [<=] l'Hr:r [<=] r'Hi:In 1 (Node l false r)In 1 (Node l' o' r')l, r, l':treeo':boolr':treeH:true = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~1 (Node l false r)In i rl, r, l':treeo':boolr':treeH:true = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~0 (Node l false r)In i~0 (Node l' o' r')l, r, l':treeo':boolr':treeH:true = trueHl:l [<=] l'Hr:r [<=] r'Hi:In 1 (Node l false r)In 1 (Node l' o' r')l, r, l':treeo':boolr':treeH:true = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~0 (Node l false r)In i~0 (Node l' o' r')l, r, l':treeo':boolr':treeH:true = trueHl:l [<=] l'Hr:r [<=] r'Hi:In 1 (Node l false r)In 1 (Node l' o' r')l, r, l':treeo':boolr':treeH:true = trueHl:l [<=] l'Hr:r [<=] r'i:positiveHi:In i~0 (Node l false r)In i ll, r, l':treeo':boolr':treeH:true = trueHl:l [<=] l'Hr:r [<=] r'Hi:In 1 (Node l false r)In 1 (Node l' o' r')discriminate Hi. Qed.l, r, l':treeo':boolr':treeH:true = trueHl:l [<=] l'Hr:r [<=] r'Hi:In 1 (Node l false r)In 1 (Node l' o' r')forall s s' : t, s [<=] s' -> subset s s' = trueforall s s' : t, s [<=] s' -> subset s s' = trueapply -> subset_spec; trivial. Qed.s, s':ts [<=] s' -> subset s s' = trueforall s s' : t, subset s s' = true -> s [<=] s'forall s s' : t, subset s s' = true -> s [<=] s'apply <- subset_spec; trivial. Qed.s, s':tsubset s s' = true -> s [<=] s'
Specification of equal (via subset)
forall s s' : t, equal s s' = subset s s' && subset s' sforall s s' : t, equal s s' = subset s s' && subset s' sl:treeo:boolr:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rnegb o &&& is_empty l &&& is_empty r = negb o &&& is_empty l &&& is_empty r && truel:treeo:boolr:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rl':treeo':boolr':treeeqb o o' &&& equal l l' &&& equal r r' = (negb o ||| o') &&& subset l l' &&& subset r r' && ((negb o' ||| o) &&& subset l' l &&& subset r' r)l, r:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rnegb true &&& is_empty l &&& is_empty r = negb true &&& is_empty l &&& is_empty r && truel, r:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rnegb false &&& is_empty l &&& is_empty r = negb false &&& is_empty l &&& is_empty r && truel:treeo:boolr:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rl':treeo':boolr':treeeqb o o' &&& equal l l' &&& equal r r' = (negb o ||| o') &&& subset l l' &&& subset r r' && ((negb o' ||| o) &&& subset l' l &&& subset r' r)l, r:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rnegb false &&& is_empty l &&& is_empty r = negb false &&& is_empty l &&& is_empty r && truel:treeo:boolr:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rl':treeo':boolr':treeeqb o o' &&& equal l l' &&& equal r r' = (negb o ||| o') &&& subset l l' &&& subset r r' && ((negb o' ||| o) &&& subset l' l &&& subset r' r)l, r:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rnegb false &&& is_empty l &&& is_empty r = true && (negb false &&& is_empty l &&& is_empty r)l:treeo:boolr:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rl':treeo':boolr':treeeqb o o' &&& equal l l' &&& equal r r' = (negb o ||| o') &&& subset l l' &&& subset r r' && ((negb o' ||| o) &&& subset l' l &&& subset r' r)l:treeo:boolr:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rl':treeo':boolr':treeeqb o o' &&& equal l l' &&& equal r r' = (negb o ||| o') &&& subset l l' &&& subset r r' && ((negb o' ||| o) &&& subset l' l &&& subset r' r)l:treeo:boolr:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rl':treeo':boolr':treeeqb o o' && equal l l' && equal r r' = (negb o ||| o') && subset l l' && subset r r' && ((negb o' ||| o) && subset l' l && subset r' r)l:treeo:boolr:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rl':treeo':boolr':treeeqb o o' && equal l l' && equal r r' = true <-> (negb o ||| o') && subset l l' && subset r r' && ((negb o' ||| o) && subset l' l && subset r' r) = truel:treeo:boolr:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rl':treeo':boolr':tree(o = o' /\ equal l l' = true) /\ equal r r' = true <-> ((negb o ||| o' = true /\ subset l l' = true) /\ subset r r' = true) /\ (negb o' ||| o = true /\ subset l' l = true) /\ subset r' r = truel:treeo:boolr:treeIHl:forall s' : t, equal l s' = subset l s' && subset s' lIHr:forall s' : t, equal r s' = subset r s' && subset s' rl':treeo':boolr':tree(o = o' /\ subset l l' = true /\ subset l' l = true) /\ subset r r' = true /\ subset r' r = true <-> ((negb o ||| o' = true /\ subset l l' = true) /\ subset r r' = true) /\ (negb o' ||| o = true /\ subset l' l = true) /\ subset r' r = truel:treeo:boolr, l':treeo':boolr':tree(o = o' /\ subset l l' = true /\ subset l' l = true) /\ subset r r' = true /\ subset r' r = true <-> ((negb o ||| o' = true /\ subset l l' = true) /\ subset r r' = true) /\ (negb o' ||| o = true /\ subset l' l = true) /\ subset r' r = truel, r, l':treeo':boolr':treeH0:subset r r' = trueH3:subset r' r = trueH1:subset l l' = trueH4:subset l' l = truenegb o' ||| o' = truel, r, l':treeo':boolr':treeH0:subset r r' = trueH3:subset r' r = trueH1:subset l l' = trueH4:subset l' l = truenegb o' ||| o' = truel:treeo:boolr, l':treeo':boolr':treeH2:subset r r' = trueH3:subset r' r = trueH1:negb o ||| o' = trueH4:subset l l' = trueH:negb o' ||| o = trueH5:subset l' l = trueo = o'l, r, l':treeo':boolr':treeH0:subset r r' = trueH3:subset r' r = trueH1:subset l l' = trueH4:subset l' l = truenegb o' ||| o' = truel:treeo:boolr, l':treeo':boolr':treeH2:subset r r' = trueH3:subset r' r = trueH1:negb o ||| o' = trueH4:subset l l' = trueH:negb o' ||| o = trueH5:subset l' l = trueo = o'l:treeo:boolr, l':treeo':boolr':treeH2:subset r r' = trueH3:subset r' r = trueH1:negb o ||| o' = trueH4:subset l l' = trueH:negb o' ||| o = trueH5:subset l' l = trueo = o'destruct o'; trivial. Qed.l, r, l':treeo':boolr':treeH2:subset r r' = trueH3:subset r' r = trueH1:negb false ||| o' = trueH4:subset l l' = trueH:negb o' ||| false = trueH5:subset l' l = truefalse = o'forall s s' : t, s [=] s' <-> equal s s' = trueforall s s' : t, s [=] s' <-> equal s s' = trues, s':ts [=] s' <-> equal s s' = trues, s':ts [=] s' <-> subset s s' && subset s' s = trues, s':ts [=] s' <-> subset s s' = true /\ subset s' s = trues, s':ts [=] s' <-> s [<=] s' /\ s' [<=] sfirstorder. Qed.s, s':t(forall a : elt, In a s <-> In a s') <-> (forall a : elt, In a s -> In a s') /\ (forall a : elt, In a s' -> In a s)forall s s' : t, s [=] s' -> equal s s' = trueforall s s' : t, s [=] s' -> equal s s' = trueapply -> equal_spec; trivial. Qed.s, s':ts [=] s' -> equal s s' = trueforall s s' : t, equal s s' = true -> s [=] s'forall s s' : t, equal s s' = true -> s [=] s'apply <- equal_spec; trivial. Qed.s, s':tequal s s' = true -> s [=] s'forall s s' : t, {eq s s'} + {~ eq s s'}forall s s' : t, {eq s s'} + {~ eq s s'}forall s s' : t, {s [=] s'} + {~ s [=] s'}s, s':t{s [=] s'} + {~ s [=] s'}s, s':tH:equal s s' = true{s [=] s'} + {~ s [=] s'}s, s':tH:equal s s' = false{s [=] s'} + {~ s [=] s'}s, s':tH:equal s s' = trues [=] s's, s':tH:equal s s' = false{s [=] s'} + {~ s [=] s'}s, s':tH:equal s s' = false{s [=] s'} + {~ s [=] s'}abstract (intro H'; rewrite (equal_1 H') in H; discriminate). Defined.s, s':tH:equal s s' = false~ s [=] s'
(Specified) definition of compare
forall u v u' v' : comparison, u = CompOpp u' -> v = CompOpp v' -> lex u v = CompOpp (lex u' v')forall u v u' v' : comparison, u = CompOpp u' -> v = CompOpp v' -> lex u v = CompOpp (lex u' v')case u'; reflexivity. Qed.u', v':comparisonlex (CompOpp u') (CompOpp v') = CompOpp (lex u' v')forall b b' : bool, compare_bool b b' = CompOpp (compare_bool b' b)intros [|] [|]; reflexivity. Qed.forall b b' : bool, compare_bool b b' = CompOpp (compare_bool b' b)forall s s' : t, compare_fun s s' = CompOpp (compare_fun s' s)forall s s' : t, compare_fun s s' = CompOpp (compare_fun s' s)l':treeo':boolr':treecompare_fun Leaf (Node l' o' r') = CompOpp (compare_fun (Node l' o' r') Leaf)l:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = CompOpp (compare_fun s' l)IHr:forall s' : t, compare_fun r s' = CompOpp (compare_fun s' r)compare_fun (Node l o r) Leaf = CompOpp (compare_fun Leaf (Node l o r))l:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = CompOpp (compare_fun s' l)IHr:forall s' : t, compare_fun r s' = CompOpp (compare_fun s' r)l':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = CompOpp (compare_fun (Node l' o' r') (Node l o r))l':treeo':boolr':tree(if is_empty (Node l' o' r') then Eq else Lt) = CompOpp (if is_empty (Node l' o' r') then Eq else Gt)l:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = CompOpp (compare_fun s' l)IHr:forall s' : t, compare_fun r s' = CompOpp (compare_fun s' r)compare_fun (Node l o r) Leaf = CompOpp (compare_fun Leaf (Node l o r))l:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = CompOpp (compare_fun s' l)IHr:forall s' : t, compare_fun r s' = CompOpp (compare_fun s' r)l':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = CompOpp (compare_fun (Node l' o' r') (Node l o r))l:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = CompOpp (compare_fun s' l)IHr:forall s' : t, compare_fun r s' = CompOpp (compare_fun s' r)compare_fun (Node l o r) Leaf = CompOpp (compare_fun Leaf (Node l o r))l:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = CompOpp (compare_fun s' l)IHr:forall s' : t, compare_fun r s' = CompOpp (compare_fun s' r)l':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = CompOpp (compare_fun (Node l' o' r') (Node l o r))l:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = CompOpp (compare_fun s' l)IHr:forall s' : t, compare_fun r s' = CompOpp (compare_fun s' r)(if is_empty (Node l o r) then Eq else Gt) = CompOpp (if is_empty (Node l o r) then Eq else Lt)l:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = CompOpp (compare_fun s' l)IHr:forall s' : t, compare_fun r s' = CompOpp (compare_fun s' r)l':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = CompOpp (compare_fun (Node l' o' r') (Node l o r))l:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = CompOpp (compare_fun s' l)IHr:forall s' : t, compare_fun r s' = CompOpp (compare_fun s' r)l':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = CompOpp (compare_fun (Node l' o' r') (Node l o r))l:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = CompOpp (compare_fun s' l)IHr:forall s' : t, compare_fun r s' = CompOpp (compare_fun s' r)l':treeo':boolr':treelex (compare_bool o o') (lex (compare_fun l l') (compare_fun r r')) = CompOpp (lex (compare_bool o' o) (lex (compare_fun l' l) (compare_fun r' r)))case compare_bool; simpl; trivial; apply lex_Opp; auto. Qed.l:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = CompOpp (compare_fun s' l)IHr:forall s' : t, compare_fun r s' = CompOpp (compare_fun s' r)l':treeo':boolr':treelex (CompOpp (compare_bool o' o)) (lex (compare_fun l l') (compare_fun r r')) = CompOpp (lex (compare_bool o' o) (lex (compare_fun l' l) (compare_fun r' r)))forall u v : comparison, lex u v = Eq <-> u = Eq /\ v = Eqintros u v; destruct u; intuition discriminate. Qed.forall u v : comparison, lex u v = Eq <-> u = Eq /\ v = Eqforall b1 b2 : bool, compare_bool b1 b2 = Eq <-> eqb b1 b2 = trueintros [|] [|]; intuition discriminate. Qed.forall b1 b2 : bool, compare_bool b1 b2 = Eq <-> eqb b1 b2 = trueforall s s' : t, compare_fun s s' = Eq <-> equal s s' = trueforall s s' : t, compare_fun s s' = Eq <-> equal s s' = truecompare_fun Leaf Leaf = Eq <-> equal Leaf Leaf = truel':treeo':boolr':treecompare_fun Leaf (Node l' o' r') = Eq <-> equal Leaf (Node l' o' r') = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truecompare_fun (Node l o r) Leaf = Eq <-> equal (Node l o r) Leaf = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truel':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = Eq <-> equal (Node l o r) (Node l' o' r') = trueEq = Eq <-> true = truel':treeo':boolr':treecompare_fun Leaf (Node l' o' r') = Eq <-> equal Leaf (Node l' o' r') = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truecompare_fun (Node l o r) Leaf = Eq <-> equal (Node l o r) Leaf = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truel':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = Eq <-> equal (Node l o r) (Node l' o' r') = truel':treeo':boolr':treecompare_fun Leaf (Node l' o' r') = Eq <-> equal Leaf (Node l' o' r') = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truecompare_fun (Node l o r) Leaf = Eq <-> equal (Node l o r) Leaf = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truel':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = Eq <-> equal (Node l o r) (Node l' o' r') = truel':treeo':boolr':tree(if is_empty (Node l' o' r') then Eq else Lt) = Eq <-> is_empty (Node l' o' r') = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truecompare_fun (Node l o r) Leaf = Eq <-> equal (Node l o r) Leaf = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truel':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = Eq <-> equal (Node l o r) (Node l' o' r') = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truecompare_fun (Node l o r) Leaf = Eq <-> equal (Node l o r) Leaf = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truel':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = Eq <-> equal (Node l o r) (Node l' o' r') = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = true(if is_empty (Node l o r) then Eq else Gt) = Eq <-> is_empty (Node l o r) = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truel':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = Eq <-> equal (Node l o r) (Node l' o' r') = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truel':treeo':boolr':treecompare_fun (Node l o r) (Node l' o' r') = Eq <-> equal (Node l o r) (Node l' o' r') = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truel':treeo':boolr':treelex (compare_bool o o') (lex (compare_fun l l') (compare_fun r r')) = Eq <-> eqb o o' &&& equal l l' &&& equal r r' = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truel':treeo':boolr':treelex (compare_bool o o') (lex (compare_fun l l') (compare_fun r r')) = Eq <-> (eqb o o' = true /\ equal l l' = true) /\ equal r r' = truel:treeo:boolr:treeIHl:forall s' : t, compare_fun l s' = Eq <-> equal l s' = trueIHr:forall s' : t, compare_fun r s' = Eq <-> equal r s' = truel':treeo':boolr':treelex (compare_bool o o') (lex (compare_fun l l') (compare_fun r r')) = Eq <-> (compare_bool o o' = Eq /\ compare_fun l l' = Eq) /\ compare_fun r r' = Eql:treeo:boolr, l':treeo':boolr':treelex (compare_bool o o') (lex (compare_fun l l') (compare_fun r r')) = Eq <-> (compare_bool o o' = Eq /\ compare_fun l l' = Eq) /\ compare_fun r r' = Eql:treeo:boolr, l':treeo':boolr':treelex (compare_bool o o') (lex (compare_fun l l') (compare_fun r r')) = Eq <-> compare_bool o o' = Eq /\ compare_fun l l' = Eq /\ compare_fun r r' = Eqreflexivity. Qed.l:treeo:boolr, l':treeo':boolr':treelex (compare_bool o o') (lex (compare_fun l l') (compare_fun r r')) = Eq <-> lex (compare_bool o o') (lex (compare_fun l l') (compare_fun r r')) = Eqforall s s' : t, compare_fun s s' = Gt -> lt s' sforall s s' : t, compare_fun s s' = Gt -> lt s' sforall s s' : t, compare_fun s s' = Gt -> compare_fun s' s = Lts, s':tcompare_fun s s' = Gt -> compare_fun s' s = Ltcase compare_fun; trivial; intros; discriminate. Qed.s, s':tCompOpp (compare_fun s' s) = Gt -> compare_fun s' s = Ltforall s s' : t, compare_fun s s' = Eq -> eq s s'forall s s' : t, compare_fun s s' = Eq -> eq s s'forall s s' : t, compare_fun s s' = Eq -> s [=] s's, s':tcompare_fun s s' = Eq -> s [=] s'trivial. Qed.s, s':tequal s s' = true -> equal s s' = trueforall s s' : t, Compare lt eq s s'forall s s' : t, Compare lt eq s s's, s':tCompare lt eq s s's, s':tH:compare_fun s s' = EqCompare lt eq s s's, s':tH:compare_fun s s' = LtCompare lt eq s s's, s':tH:compare_fun s s' = GtCompare lt eq s s's, s':tH:compare_fun s s' = Eqeq s s's, s':tH:compare_fun s s' = LtCompare lt eq s s's, s':tH:compare_fun s s' = GtCompare lt eq s s's, s':tH:compare_fun s s' = LtCompare lt eq s s's, s':tH:compare_fun s s' = GtCompare lt eq s s's, s':tH:compare_fun s s' = Ltlt s s's, s':tH:compare_fun s s' = GtCompare lt eq s s's, s':tH:compare_fun s s' = GtCompare lt eq s s'apply compare_gt, H. Defined. Section lt_spec. Inductive ct: comparison -> comparison -> comparison -> Prop := | ct_xxx: forall x, ct x x x | ct_xex: forall x, ct x Eq x | ct_exx: forall x, ct Eq x x | ct_glx: forall x, ct Gt Lt x | ct_lgx: forall x, ct Lt Gt x.s, s':tH:compare_fun s s' = Gtlt s' sforall x : comparison, ct (CompOpp x) x Eqdestruct x; constructor. Qed.forall x : comparison, ct (CompOpp x) x Eqforall x : comparison, ct x (CompOpp x) Eqdestruct x; constructor. Qed.forall x : comparison, ct x (CompOpp x) Eqforall x : comparison, ct Lt x Ltdestruct x; constructor. Qed.forall x : comparison, ct Lt x Ltforall x : comparison, ct Gt x Gtdestruct x; constructor. Qed.forall x : comparison, ct Gt x Gtforall x : comparison, ct x Lt Ltdestruct x; constructor. Qed.forall x : comparison, ct x Lt Ltforall x : comparison, ct x Gt Gtdestruct x; constructor. Qed. Local Hint Constructors ct: ct. Local Hint Resolve ct_cxe ct_xce ct_lxl ct_xll ct_gxg ct_xgg: ct. Ltac ct := trivial with ct.forall x : comparison, ct x Gt Gtforall u v w u' v' w' : comparison, ct u v w -> ct u' v' w' -> ct (lex u u') (lex v v') (lex w w')forall u v w u' v' w' : comparison, ct u v w -> ct u' v' w' -> ct (lex u u') (lex v v') (lex w w')inversion_clear H; inversion_clear H'; ct; destruct w; ct; destruct w'; ct. Qed.u, v, w, u', v', w':comparisonH:ct u v wH':ct u' v' w'ct (lex u u') (lex v v') (lex w w')forall a b c : bool, ct (compare_bool a b) (compare_bool b c) (compare_bool a c)intros [|] [|] [|]; constructor. Qed.forall a b c : bool, ct (compare_bool a b) (compare_bool b c) (compare_bool a c)forall s : t, compare_fun s Leaf = (if is_empty s then Eq else Gt)forall s : t, compare_fun s Leaf = (if is_empty s then Eq else Gt)s:tcompare_fun s Leaf = (if is_empty s then Eq else Gt)s:tCompOpp (compare_fun Leaf s) = (if is_empty s then Eq else Gt)case (is_empty s); reflexivity. Qed.s:tCompOpp (if is_empty s then Eq else Lt) = (if is_empty s then Eq else Gt)forall a : t, is_empty a = true -> forall b : t, compare_fun a b = (if is_empty b then Eq else Lt)forall a : t, is_empty a = true -> forall b : t, compare_fun a b = (if is_empty b then Eq else Lt)l:treeo:boolr:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)is_empty (Node l o r) = true -> forall b : t, compare_fun (Node l o r) b = (if is_empty b then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)is_empty (Node l true r) = true -> forall b : t, compare_fun (Node l true r) b = (if is_empty b then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)is_empty (Node l false r) = true -> forall b : t, compare_fun (Node l false r) b = (if is_empty b then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)is_empty (Node l false r) = true -> forall b : t, compare_fun (Node l false r) b = (if is_empty b then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)is_empty l &&& is_empty r = true -> forall b : t, compare_fun (Node l false r) b = (if is_empty b then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)is_empty l = true /\ is_empty r = true -> forall b : t, compare_fun (Node l false r) b = (if is_empty b then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)Hl:is_empty l = trueHr:is_empty r = trueforall b : t, compare_fun (Node l false r) b = (if is_empty b then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)Hl:is_empty l = trueHr:is_empty r = true(if is_empty l &&& is_empty r then Eq else Gt) = (if is_empty Leaf then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)Hl:is_empty l = trueHr:is_empty r = truel', r':treelex (compare_fun l l') (compare_fun r r') = (if is_empty (Node l' false r') then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)Hl:is_empty l = trueHr:is_empty r = trueEq = (if is_empty Leaf then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)Hl:is_empty l = trueHr:is_empty r = truel', r':treelex (compare_fun l l') (compare_fun r r') = (if is_empty (Node l' false r') then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)Hl:is_empty l = trueHr:is_empty r = truel', r':treelex (compare_fun l l') (compare_fun r r') = (if is_empty (Node l' false r') then Eq else Lt)l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)Hl:is_empty l = trueHr:is_empty r = truel', r':treelex (if is_empty l' then Eq else Lt) (if is_empty r' then Eq else Lt) = (if is_empty (Node l' false r') then Eq else Lt)case (is_empty l'); case (is_empty r'); trivial. Qed.l, r:treeIHl:is_empty l = true -> forall b : t, compare_fun l b = (if is_empty b then Eq else Lt)IHr:is_empty r = true -> forall b : t, compare_fun r b = (if is_empty b then Eq else Lt)Hl:is_empty l = trueHr:is_empty r = truel', r':treelex (if is_empty l' then Eq else Lt) (if is_empty r' then Eq else Lt) = (if is_empty l' &&& is_empty r' then Eq else Lt)forall a : t, is_empty a = true -> forall b : t, compare_fun b a = (if is_empty b then Eq else Gt)forall a : t, is_empty a = true -> forall b : t, compare_fun b a = (if is_empty b then Eq else Gt)forall a : t, is_empty a = true -> forall b : t, compare_fun b a = compare_fun b Leafa:tH:is_empty a = trueb:tcompare_fun b a = compare_fun b Leafreflexivity. Qed.a:tH:is_empty a = trueb:tCompOpp (if is_empty b then Eq else Lt) = CompOpp (compare_fun Leaf b)forall a b c : t, ct (compare_fun a b) (compare_fun b c) (compare_fun a c)forall a b c : t, ct (compare_fun a b) (compare_fun b c) (compare_fun a c)s', s'':tct (compare_fun Leaf s') (compare_fun s' s'') (compare_fun Leaf s'')l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l':treeo':boolr':treect (compare_fun Leaf (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun Leaf Leaf)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun Leaf (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun Leaf (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l':treeo':boolr':treect (CompOpp (compare_fun (Node l' o' r') Leaf)) (compare_fun (Node l' o' r') Leaf) (compare_fun Leaf Leaf)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun Leaf (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun Leaf (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun Leaf (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun Leaf (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l':treeo':boolr', l'':treeo'':boolr'':treect (if is_empty (Node l' o' r') then Eq else Lt) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun Leaf (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l':treeo':boolr', l'':treeo'':boolr'':treeH':is_empty (Node l' o' r') = truect Eq (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun Leaf (Node l'' o'' r''))l':treeo':boolr', l'':treeo'':boolr'':treeH':is_empty (Node l' o' r') = falsect Lt (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun Leaf (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l':treeo':boolr', l'':treeo'':boolr'':treeH':is_empty (Node l' o' r') = truect Eq (if is_empty (Node l'' o'' r'') then Eq else Lt) (compare_fun Leaf (Node l'' o'' r''))l':treeo':boolr', l'':treeo'':boolr'':treeH':is_empty (Node l' o' r') = falsect Lt (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun Leaf (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l':treeo':boolr', l'':treeo'':boolr'':treeH':is_empty (Node l' o' r') = falsect Lt (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun Leaf (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l':treeo':boolr', l'':treeo'':boolr'':treeH':is_empty (Node l' o' r') = falsect Lt (compare_fun (Node l' o' r') (Node l'' o'' r'')) (if is_empty (Node l'' o'' r'') then Eq else Lt)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l':treeo':boolr', l'':treeo'':boolr'':treeH':is_empty (Node l' o' r') = falseH'':is_empty (Node l'' o'' r'') = truect Lt (compare_fun (Node l' o' r') (Node l'' o'' r'')) Eql':treeo':boolr', l'':treeo'':boolr'':treeH':is_empty (Node l' o' r') = falseH'':is_empty (Node l'' o'' r'') = falsect Lt (compare_fun (Node l' o' r') (Node l'' o'' r'')) Ltl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l':treeo':boolr', l'':treeo'':boolr'':treeH':is_empty (Node l' o' r') = falseH'':is_empty (Node l'' o'' r'') = truect Lt Gt Eql':treeo':boolr', l'':treeo'':boolr'':treeH':is_empty (Node l' o' r') = falseH'':is_empty (Node l'' o'' r'') = falsect Lt (compare_fun (Node l' o' r') (Node l'' o'' r'')) Ltl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l':treeo':boolr', l'':treeo'':boolr'':treeH':is_empty (Node l' o' r') = falseH'':is_empty (Node l'' o'' r'') = falsect Lt (compare_fun (Node l' o' r') (Node l'' o'' r'')) Ltl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)s', s'':tct (compare_fun (Node l o r) s') (compare_fun s' s'') (compare_fun (Node l o r) s'')l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)ct (compare_fun (Node l o r) Leaf) (compare_fun Leaf Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treect (compare_fun (Node l o r) Leaf) (compare_fun Leaf (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treect (compare_fun (Node l o r) Leaf) (compare_fun Leaf (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treect (compare_fun (Node l o r) Leaf) (if is_empty (Node l'' o'' r'') then Eq else Lt) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treect (if is_empty (Node l o r) then Eq else Gt) (if is_empty (Node l'' o'' r'') then Eq else Lt) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treeH:is_empty (Node l o r) = truect Eq (if is_empty (Node l'' o'' r'') then Eq else Lt) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treeH:is_empty (Node l o r) = falsect Gt (if is_empty (Node l'' o'' r'') then Eq else Lt) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treeH:is_empty (Node l o r) = truect Eq (if is_empty (Node l'' o'' r'') then Eq else Lt) (if is_empty (Node l'' o'' r'') then Eq else Lt)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treeH:is_empty (Node l o r) = falsect Gt (if is_empty (Node l'' o'' r'') then Eq else Lt) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treeH:is_empty (Node l o r) = falsect Gt (if is_empty (Node l'' o'' r'') then Eq else Lt) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treeH:is_empty (Node l o r) = falseH'':is_empty (Node l'' o'' r'') = truect Gt Eq (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treeH:is_empty (Node l o r) = falseH'':is_empty (Node l'' o'' r'') = falsect Gt Lt (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treeH:is_empty (Node l o r) = falseH'':is_empty (Node l'' o'' r'') = truect Gt Eq Gtl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treeH:is_empty (Node l o r) = falseH'':is_empty (Node l'' o'' r'') = falsect Gt Lt (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l'':treeo'':boolr'':treeH:is_empty (Node l o r) = falseH'':is_empty (Node l'' o'' r'') = falsect Gt Lt (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') Leaf) (compare_fun (Node l o r) Leaf)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treect (compare_fun (Node l o r) (Node l' o' r')) (if is_empty (Node l' o' r') then Eq else Gt) (if is_empty (Node l o r) then Eq else Gt)l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treeH:is_empty (Node l o r) = truect (compare_fun (Node l o r) (Node l' o' r')) (if is_empty (Node l' o' r') then Eq else Gt) Eql:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treeH:is_empty (Node l o r) = falsect (compare_fun (Node l o r) (Node l' o' r')) (if is_empty (Node l' o' r') then Eq else Gt) Gtl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treeH:is_empty (Node l o r) = truect (CompOpp (if is_empty (Node l' o' r') then Eq else Gt)) (if is_empty (Node l' o' r') then Eq else Gt) Eql:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treeH:is_empty (Node l o r) = falsect (compare_fun (Node l o r) (Node l' o' r')) (if is_empty (Node l' o' r') then Eq else Gt) Gtl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treeH:is_empty (Node l o r) = falsect (compare_fun (Node l o r) (Node l' o' r')) (if is_empty (Node l' o' r') then Eq else Gt) Gtl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treeH:is_empty (Node l o r) = falseH':is_empty (Node l' o' r') = truect (compare_fun (Node l o r) (Node l' o' r')) Eq Gtl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treeH:is_empty (Node l o r) = falseH':is_empty (Node l' o' r') = falsect (compare_fun (Node l o r) (Node l' o' r')) Gt Gtl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treeH:is_empty (Node l o r) = falseH':is_empty (Node l' o' r') = truect Gt Eq Gtl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treeH:is_empty (Node l o r) = falseH':is_empty (Node l' o' r') = falsect (compare_fun (Node l o r) (Node l' o' r')) Gt Gtl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr':treeH:is_empty (Node l o r) = falseH':is_empty (Node l' o' r') = falsect (compare_fun (Node l o r) (Node l' o' r')) Gt Gtl:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_fun (Node l o r) (Node l' o' r')) (compare_fun (Node l' o' r') (Node l'' o'' r'')) (compare_fun (Node l o r) (Node l'' o'' r''))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (lex (compare_bool o o') (lex (compare_fun l l') (compare_fun r r'))) (lex (compare_bool o' o'') (lex (compare_fun l' l'') (compare_fun r' r''))) (lex (compare_bool o o'') (lex (compare_fun l l'') (compare_fun r r'')))l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (compare_bool o o') (compare_bool o' o'') (compare_bool o o'')l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (lex (compare_fun l l') (compare_fun r r')) (lex (compare_fun l' l'') (compare_fun r' r'')) (lex (compare_fun l l'') (compare_fun r r''))apply ct_lex; trivial. Qed. End lt_spec.l:treeo:boolr:treeIHl:forall b c : t, ct (compare_fun l b) (compare_fun b c) (compare_fun l c)IHr:forall b c : t, ct (compare_fun r b) (compare_fun b c) (compare_fun r c)l':treeo':boolr', l'':treeo'':boolr'':treect (lex (compare_fun l l') (compare_fun r r')) (lex (compare_fun l' l'') (compare_fun r' r'')) (lex (compare_fun l l'') (compare_fun r r''))forall s s' s'' : t, lt s s' -> lt s' s'' -> lt s s''forall s s' s'' : t, lt s s' -> lt s' s'' -> lt s s''forall s s' s'' : t, compare_fun s s' = Lt -> compare_fun s' s'' = Lt -> compare_fun s s'' = Lta, b, c:tcompare_fun a b = Lt -> compare_fun b c = Lt -> compare_fun a c = Ltinversion_clear H; trivial; intros; discriminate. Qed.a, b, c:tH:ct (compare_fun a b) (compare_fun b c) (compare_fun a c)compare_fun a b = Lt -> compare_fun b c = Lt -> compare_fun a c = Ltforall s s' : t, lt s s' -> ~ eq s s'forall s s' : t, lt s s' -> ~ eq s s'forall s s' : t, compare_fun s s' = Lt -> ~ s [=] s's, s':tH:compare_fun s s' = LtH':s [=] s'Falsecongruence. Qed.s, s':tH:compare_fun s s' = LtH':compare_fun s s' = EqFalse
Specification of add
forall (x y : elt) (s : t), In y (add x s) <-> x = y \/ In y sforall (x y : elt) (s : t), In y (add x s) <-> x = y \/ In y sinduction x; intros [y|y|] [|l o r]; simpl mem; try (rewrite IHx; clear IHx); rewrite ?mem_Leaf; intuition congruence. Qed.forall (x y : elt) (s : t), mem y (add x s) = true <-> x = y \/ mem y s = trueforall (s : t) (x y : elt), x = y -> In y (add x s)forall (s : t) (x y : elt), x = y -> In y (add x s)s:tx, y:eltH:x = yIn y (add x s)s:tx, y:eltH:x = yx = y \/ In y sassumption. Qed.s:tx, y:eltH:x = yx = yforall (s : t) (x y : elt), In y s -> In y (add x s)forall (s : t) (x y : elt), In y s -> In y (add x s)s:tx, y:eltH:In y sIn y (add x s)s:tx, y:eltH:In y sx = y \/ In y sassumption. Qed.s:tx, y:eltH:In y sIn y sforall (s : t) (x y : elt), x <> y -> In y (add x s) -> In y sforall (s : t) (x y : elt), x <> y -> In y (add x s) -> In y ss:tx, y:eltH:x <> yIn y (add x s) -> In y ss:tx, y:eltH:x <> yx = y \/ In y s -> In y selim H; trivial. Qed.s:ty:eltH:y <> yIn y s
Specification of remove
forall (x y : elt) (s : t), In y (remove x s) <-> x <> y /\ In y sforall (x y : elt) (s : t), In y (remove x s) <-> x <> y /\ In y sinduction x; intros [y|y|] [|l o r]; simpl remove; rewrite ?mem_node; simpl mem; try (rewrite IHx; clear IHx); rewrite ?mem_Leaf; intuition congruence. Qed.forall (x y : elt) (s : t), mem y (remove x s) = true <-> x <> y /\ mem y s = trueforall (s : t) (x y : elt), x = y -> ~ In y (remove x s)forall (s : t) (x y : elt), x = y -> ~ In y (remove x s)s:tx, y:eltH:x = y~ In y (remove x s)tauto. Qed.s:tx, y:eltH:x = y~ (x <> y /\ In y s)forall (s : t) (x y : elt), x <> y -> In y s -> In y (remove x s)forall (s : t) (x y : elt), x <> y -> In y s -> In y (remove x s)s:tx, y:eltH:x <> yH0:In y sIn y (remove x s)split; assumption. Qed.s:tx, y:eltH:x <> yH0:In y sx <> y /\ In y sforall (s : t) (x y : elt), In y (remove x s) -> In y sforall (s : t) (x y : elt), In y (remove x s) -> In y ss:tx, y:eltIn y (remove x s) -> In y stauto. Qed.s:tx, y:eltx <> y /\ In y s -> In y s
Specification of singleton
forall x y : elt, In y (singleton x) -> x = yforall x y : elt, In y (singleton x) -> x = yforall x y : elt, In y (add x empty) -> x = yx, y:eltIn y (add x empty) -> x = yx, y:eltx = y \/ In y empty -> x = yx, y:eltx = y \/ mem y empty = true -> x = yintuition discriminate. Qed.x, y:eltx = y \/ false = true -> x = yforall x y : elt, x = y -> In y (singleton x)forall x y : elt, x = y -> In y (singleton x)forall x y : elt, x = y -> In y (add x empty)x, y:eltH:x = yIn y (add x empty)assumption. Qed.x, y:eltH:x = yx = y
Specification of union
forall (x : elt) (s s' : t), In x (union s s') <-> In x s \/ In x s'forall (x : elt) (s s' : t), In x (union s s') <-> In x s \/ In x s'forall (x : elt) (s s' : t), mem x (union s s') = true <-> mem x s = true \/ mem x s' = trueapply orb_true_iff. Qed.s1:treeb:bools2, s'1:treeb0:bools'2:treeb || b0 = true <-> b = true \/ b0 = trueforall (s s' : t) (x : elt), In x (union s s') -> In x s \/ In x s'forall (s s' : t) (x : elt), In x (union s s') -> In x s \/ In x s's, s':tx:eltH:In x (union s s')In x s \/ In x s'assumption. Qed.s, s':tx:eltH:In x (union s s')In x (union s s')forall (s s' : t) (x : elt), In x s -> In x (union s s')forall (s s' : t) (x : elt), In x s -> In x (union s s')s, s':tx:eltH:In x sIn x (union s s')s, s':tx:eltH:In x sIn x s \/ In x s'assumption. Qed.s, s':tx:eltH:In x sIn x sforall (s s' : t) (x : elt), In x s' -> In x (union s s')forall (s s' : t) (x : elt), In x s' -> In x (union s s')s, s':tx:eltH:In x s'In x (union s s')s, s':tx:eltH:In x s'In x s \/ In x s'assumption. Qed.s, s':tx:eltH:In x s'In x s'
Specification of inter
forall (x : elt) (s s' : t), In x (inter s s') <-> In x s /\ In x s'forall (x : elt) (s s' : t), In x (inter s s') <-> In x s /\ In x s'forall (x : elt) (s s' : t), mem x (inter s s') = true <-> mem x s = true /\ mem x s' = trueapply andb_true_iff. Qed.s1:treeb:bools2, s'1:treeb0:bools'2:treeb && b0 = true <-> b = true /\ b0 = trueforall (s s' : t) (x : elt), In x (inter s s') -> In x sforall (s s' : t) (x : elt), In x (inter s s') -> In x ss, s':tx:eltIn x (inter s s') -> In x stauto. Qed.s, s':tx:eltIn x s /\ In x s' -> In x sforall (s s' : t) (x : elt), In x (inter s s') -> In x s'forall (s s' : t) (x : elt), In x (inter s s') -> In x s's, s':tx:eltIn x (inter s s') -> In x s'tauto. Qed.s, s':tx:eltIn x s /\ In x s' -> In x s'forall (s s' : t) (x : elt), In x s -> In x s' -> In x (inter s s')forall (s s' : t) (x : elt), In x s -> In x s' -> In x (inter s s')s, s':tx:eltH:In x sH0:In x s'In x (inter s s')split; assumption. Qed.s, s':tx:eltH:In x sH0:In x s'In x s /\ In x s'
Specification of diff
forall (x : elt) (s s' : t), In x (diff s s') <-> In x s /\ ~ In x s'forall (x : elt) (s s' : t), In x (diff s s') <-> In x s /\ ~ In x s'forall (x : elt) (s s' : t), mem x (diff s s') = true <-> mem x s = true /\ mem x s' <> trues1:treeb:bools2, l':treeo':boolr':treeb && negb o' = true <-> b = true /\ o' <> truedestruct o'; intuition discriminate. Qed.s1:treeb:bools2, l':treeo':boolr':treeb = true /\ negb o' = true <-> b = true /\ o' <> trueforall (s s' : t) (x : elt), In x (diff s s') -> In x sforall (s s' : t) (x : elt), In x (diff s s') -> In x ss, s':tx:eltIn x (diff s s') -> In x stauto. Qed.s, s':tx:eltIn x s /\ ~ In x s' -> In x sforall (s s' : t) (x : elt), In x (diff s s') -> ~ In x s'forall (s s' : t) (x : elt), In x (diff s s') -> ~ In x s's, s':tx:eltIn x (diff s s') -> ~ In x s'tauto. Qed.s, s':tx:eltIn x s /\ ~ In x s' -> ~ In x s'forall (s s' : t) (x : elt), In x s -> ~ In x s' -> In x (diff s s')forall (s s' : t) (x : elt), In x s -> ~ In x s' -> In x (diff s s')s, s':tx:eltH:In x sH0:~ In x s'In x (diff s s')split; assumption. Qed.s, s':tx:eltH:In x sH0:~ In x s'In x s /\ ~ In x s'
Specification of fold
forall (s : t) (A : Type) (i : A) (f : elt -> A -> A), fold f s i = fold_left (fun (a : A) (e : elt) => f e a) (elements s) iforall (s : t) (A : Type) (i : A) (f : elt -> A -> A), fold f s i = fold_left (fun (a : A) (e : elt) => f e a) (elements s) iforall (s : t) (A : Type) (i : A) (f : elt -> A -> A), xfold f s i 1 = fold_left (fun (a : A) (e : elt) => f e a) (xelements s 1 nil) is:tA:Typei:Af:elt -> A -> Axfold f s i 1 = fold_left (fun (a : A) (e : elt) => f e a) (xelements s 1 nil) iA:Typef:elt -> A -> Aforall (s : t) (i : A), xfold f s i 1 = fold_left (fun (a : A) (e : elt) => f e a) (xelements s 1 nil) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> Aforall (s : t) (i : A), xfold f s i 1 = fold_left f' (xelements s 1 nil) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> Aforall (s : t) (i : A) (j : elt) (acc : list elt), fold_left f' acc (xfold f s i j) = fold_left f' (xelements s j acc) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> AH:forall (s : t) (i : A) (j : elt) (acc : list elt), fold_left f' acc (xfold f s i j) = fold_left f' (xelements s j acc) iforall (s : t) (i : A), xfold f s i 1 = fold_left f' (xelements s 1 nil) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> Al:treeo:boolr:treeIHl:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f l i0 j0) = fold_left f' (xelements l j0 acc0) i0IHr:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f r i0 j0) = fold_left f' (xelements r j0 acc0) i0i:Aj:eltacc:list eltfold_left f' acc (xfold f (Node l o r) i j) = fold_left f' (xelements (Node l o r) j acc) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> AH:forall (s : t) (i : A) (j : elt) (acc : list elt), fold_left f' acc (xfold f s i j) = fold_left f' (xelements s j acc) iforall (s : t) (i : A), xfold f s i 1 = fold_left f' (xelements s 1 nil) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> Al, r:treeIHl:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f l i0 j0) = fold_left f' (xelements l j0 acc0) i0IHr:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f r i0 j0) = fold_left f' (xelements r j0 acc0) i0i:Aj:eltacc:list eltfold_left f' acc (xfold f r (f (rev j) (xfold f l i j~0)) j~1) = fold_left f' (xelements l j~0 (rev j :: xelements r j~1 acc)) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> Al, r:treeIHl:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f l i0 j0) = fold_left f' (xelements l j0 acc0) i0IHr:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f r i0 j0) = fold_left f' (xelements r j0 acc0) i0i:Aj:eltacc:list eltfold_left f' acc (xfold f r (xfold f l i j~0) j~1) = fold_left f' (xelements l j~0 (xelements r j~1 acc)) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> AH:forall (s : t) (i : A) (j : elt) (acc : list elt), fold_left f' acc (xfold f s i j) = fold_left f' (xelements s j acc) iforall (s : t) (i : A), xfold f s i 1 = fold_left f' (xelements s 1 nil) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> Al, r:treeIHl:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f l i0 j0) = fold_left f' (xelements l j0 acc0) i0IHr:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f r i0 j0) = fold_left f' (xelements r j0 acc0) i0i:Aj:eltacc:list eltfold_left f' (xelements r j~1 acc) (f (rev j) (xfold f l i j~0)) = fold_left f' (rev j :: xelements r j~1 acc) (xfold f l i j~0)A:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> Al, r:treeIHl:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f l i0 j0) = fold_left f' (xelements l j0 acc0) i0IHr:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f r i0 j0) = fold_left f' (xelements r j0 acc0) i0i:Aj:eltacc:list eltfold_left f' acc (xfold f r (xfold f l i j~0) j~1) = fold_left f' (xelements l j~0 (xelements r j~1 acc)) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> AH:forall (s : t) (i : A) (j : elt) (acc : list elt), fold_left f' acc (xfold f s i j) = fold_left f' (xelements s j acc) iforall (s : t) (i : A), xfold f s i 1 = fold_left f' (xelements s 1 nil) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> Al, r:treeIHl:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f l i0 j0) = fold_left f' (xelements l j0 acc0) i0IHr:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f r i0 j0) = fold_left f' (xelements r j0 acc0) i0i:Aj:eltacc:list eltfold_left f' acc (xfold f r (xfold f l i j~0) j~1) = fold_left f' (xelements l j~0 (xelements r j~1 acc)) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> AH:forall (s : t) (i : A) (j : elt) (acc : list elt), fold_left f' acc (xfold f s i j) = fold_left f' (xelements s j acc) iforall (s : t) (i : A), xfold f s i 1 = fold_left f' (xelements s 1 nil) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> Al, r:treeIHl:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f l i0 j0) = fold_left f' (xelements l j0 acc0) i0IHr:forall (i0 : A) (j0 : elt) (acc0 : list elt), fold_left f' acc0 (xfold f r i0 j0) = fold_left f' (xelements r j0 acc0) i0i:Aj:eltacc:list eltfold_left f' (xelements r j~1 acc) (xfold f l i j~0) = fold_left f' (xelements l j~0 (xelements r j~1 acc)) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> AH:forall (s : t) (i : A) (j : elt) (acc : list elt), fold_left f' acc (xfold f s i j) = fold_left f' (xelements s j acc) iforall (s : t) (i : A), xfold f s i 1 = fold_left f' (xelements s 1 nil) iA:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> AH:forall (s : t) (i : A) (j : elt) (acc : list elt), fold_left f' acc (xfold f s i j) = fold_left f' (xelements s j acc) iforall (s : t) (i : A), xfold f s i 1 = fold_left f' (xelements s 1 nil) iexact (H s i 1 nil). Qed.A:Typef:elt -> A -> Af':=fun (a : A) (e : elt) => f e a:A -> elt -> AH:forall (s0 : t) (i0 : A) (j : elt) (acc : list elt), fold_left f' acc (xfold f s0 i0 j) = fold_left f' (xelements s0 j acc) i0s:ti:Axfold f s i 1 = fold_left f' (xelements s 1 nil) i
Specification of cardinal
forall s : t, cardinal s = length (elements s)forall s : t, cardinal s = length (elements s)forall s : t, cardinal s = length (xelements s 1 nil)forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)H:forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)forall s : t, cardinal s = length (xelements s 1 nil)l:treeb:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list elt((if b then S (cardinal l + cardinal r) else cardinal l + cardinal r) + length acc)%nat = length (if b then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)forall s : t, cardinal s = length (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list elt(S (cardinal l + cardinal r) + length acc)%nat = length (xelements l j~0 (rev j :: xelements r j~1 acc))l, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list elt(cardinal l + cardinal r + length acc)%nat = length (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)forall s : t, cardinal s = length (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list elt(S (cardinal l + cardinal r) + length acc)%nat = (cardinal l + length (rev j :: xelements r j~1 acc))%natl, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list elt(cardinal l + cardinal r + length acc)%nat = length (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)forall s : t, cardinal s = length (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list eltS (cardinal l + cardinal r + length acc) = (cardinal l + S (length (xelements r j~1 acc)))%natl, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list elt(cardinal l + cardinal r + length acc)%nat = length (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)forall s : t, cardinal s = length (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list eltS (cardinal l + cardinal r + length acc) = (cardinal l + S (cardinal r + length acc))%natl, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list elt(cardinal l + cardinal r + length acc)%nat = length (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)forall s : t, cardinal s = length (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list eltS (cardinal l + cardinal r + length acc) = S (cardinal l + cardinal r + length acc)l, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list elt(cardinal l + cardinal r + length acc)%nat = length (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)forall s : t, cardinal s = length (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list elt(cardinal l + cardinal r + length acc)%nat = length (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)forall s : t, cardinal s = length (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list elt(cardinal l + cardinal r + length acc)%nat = (cardinal l + (cardinal r + length acc))%natH:forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)forall s : t, cardinal s = length (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list elt), (cardinal l + length acc0)%nat = length (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list elt), (cardinal r + length acc0)%nat = length (xelements r j0 acc0)j:eltacc:list elt(cardinal l + cardinal r + length acc)%nat = (cardinal l + cardinal r + length acc)%natH:forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)forall s : t, cardinal s = length (xelements s 1 nil)H:forall (s : t) (j : elt) (acc : list elt), (cardinal s + length acc)%nat = length (xelements s j acc)forall s : t, cardinal s = length (xelements s 1 nil)H:forall (s0 : t) (j : elt) (acc : list elt), (cardinal s0 + length acc)%nat = length (xelements s0 j acc)s:tcardinal s = length (xelements s 1 nil)H:forall (s0 : t) (j : elt) (acc : list elt), (cardinal s0 + length acc)%nat = length (xelements s0 j acc)s:tcardinal s = (cardinal s + length nil)%natH:forall (s0 : t) (j : elt) (acc : list elt), (cardinal s0 + length acc)%nat = length (xelements s0 j acc)s:tcardinal s = (cardinal s + 0)%natreflexivity. Qed.H:forall (s0 : t) (j : elt) (acc : list elt), (cardinal s0 + length acc)%nat = length (xelements s0 j acc)s:tcardinal s = (0 + cardinal s)%nat
Specification of filter
forall (f : elt -> bool) (s : t) (x i : elt), In x (xfilter f s i) <-> In x s /\ f (i @ x) = trueforall (f : elt -> bool) (s : t) (x i : elt), In x (xfilter f s i) <-> In x s /\ f (i @ x) = truef:elt -> boolforall (s : t) (x i : elt), In x (xfilter f s i) <-> In x s /\ f (i @ x) = truef:elt -> boolforall (s : t) (x i : elt), mem x (xfilter f s i) = true <-> mem x s = true /\ f (i @ x) = truef:elt -> boolx, i:eltmem x Leaf = true <-> mem x Leaf = true /\ f (i @ x) = truef:elt -> booll:treeo:boolr:treeIHl:forall x0 i0 : elt, mem x0 (xfilter f l i0) = true <-> mem x0 l = true /\ f (i0 @ x0) = trueIHr:forall x0 i0 : elt, mem x0 (xfilter f r i0) = true <-> mem x0 r = true /\ f (i0 @ x0) = truex, i:eltmem x (node (xfilter f l i~0) (o &&& f (rev i)) (xfilter f r i~1)) = true <-> mem x (Node l o r) = true /\ f (i @ x) = truef:elt -> boolx, i:eltfalse = true <-> false = true /\ f (i @ x) = truef:elt -> booll:treeo:boolr:treeIHl:forall x0 i0 : elt, mem x0 (xfilter f l i0) = true <-> mem x0 l = true /\ f (i0 @ x0) = trueIHr:forall x0 i0 : elt, mem x0 (xfilter f r i0) = true <-> mem x0 r = true /\ f (i0 @ x0) = truex, i:eltmem x (node (xfilter f l i~0) (o &&& f (rev i)) (xfilter f r i~1)) = true <-> mem x (Node l o r) = true /\ f (i @ x) = truef:elt -> booll:treeo:boolr:treeIHl:forall x0 i0 : elt, mem x0 (xfilter f l i0) = true <-> mem x0 l = true /\ f (i0 @ x0) = trueIHr:forall x0 i0 : elt, mem x0 (xfilter f r i0) = true <-> mem x0 r = true /\ f (i0 @ x0) = truex, i:eltmem x (node (xfilter f l i~0) (o &&& f (rev i)) (xfilter f r i~1)) = true <-> mem x (Node l o r) = true /\ f (i @ x) = truef:elt -> booll:treeo:boolr:treeIHl:forall x0 i0 : elt, mem x0 (xfilter f l i0) = true <-> mem x0 l = true /\ f (i0 @ x0) = trueIHr:forall x0 i0 : elt, mem x0 (xfilter f r i0) = true <-> mem x0 r = true /\ f (i0 @ x0) = truex, i:eltmem x (Node (xfilter f l i~0) (o &&& f (rev i)) (xfilter f r i~1)) = true <-> mem x (Node l o r) = true /\ f (i @ x) = truef:elt -> booll:treeo:boolr:treeIHl:forall x0 i0 : elt, mem x0 (xfilter f l i0) = true <-> mem x0 l = true /\ f (i0 @ x0) = trueIHr:forall x0 i0 : elt, mem x0 (xfilter f r i0) = true <-> mem x0 r = true /\ f (i0 @ x0) = truex:positivei:eltmem x (xfilter f r i~1) = true <-> mem x r = true /\ f (i @ x~1) = truef:elt -> booll:treeo:boolr:treeIHl:forall x0 i0 : elt, mem x0 (xfilter f l i0) = true <-> mem x0 l = true /\ f (i0 @ x0) = trueIHr:forall x0 i0 : elt, mem x0 (xfilter f r i0) = true <-> mem x0 r = true /\ f (i0 @ x0) = truex:positivei:eltmem x (xfilter f l i~0) = true <-> mem x l = true /\ f (i @ x~0) = truef:elt -> booll:treeo:boolr:treeIHl:forall x i0 : elt, mem x (xfilter f l i0) = true <-> mem x l = true /\ f (i0 @ x) = trueIHr:forall x i0 : elt, mem x (xfilter f r i0) = true <-> mem x r = true /\ f (i0 @ x) = truei:elto &&& f (rev i) = true <-> o = true /\ f (i @ 1) = truef:elt -> booll:treeo:boolr:treeIHl:forall x0 i0 : elt, mem x0 (xfilter f l i0) = true <-> mem x0 l = true /\ f (i0 @ x0) = trueIHr:forall x0 i0 : elt, mem x0 (xfilter f r i0) = true <-> mem x0 r = true /\ f (i0 @ x0) = truex:positivei:eltmem x r = true /\ f (i~1 @ x) = true <-> mem x r = true /\ f (i @ x~1) = truef:elt -> booll:treeo:boolr:treeIHl:forall x0 i0 : elt, mem x0 (xfilter f l i0) = true <-> mem x0 l = true /\ f (i0 @ x0) = trueIHr:forall x0 i0 : elt, mem x0 (xfilter f r i0) = true <-> mem x0 r = true /\ f (i0 @ x0) = truex:positivei:eltmem x (xfilter f l i~0) = true <-> mem x l = true /\ f (i @ x~0) = truef:elt -> booll:treeo:boolr:treeIHl:forall x i0 : elt, mem x (xfilter f l i0) = true <-> mem x l = true /\ f (i0 @ x) = trueIHr:forall x i0 : elt, mem x (xfilter f r i0) = true <-> mem x r = true /\ f (i0 @ x) = truei:elto &&& f (rev i) = true <-> o = true /\ f (i @ 1) = truef:elt -> booll:treeo:boolr:treeIHl:forall x0 i0 : elt, mem x0 (xfilter f l i0) = true <-> mem x0 l = true /\ f (i0 @ x0) = trueIHr:forall x0 i0 : elt, mem x0 (xfilter f r i0) = true <-> mem x0 r = true /\ f (i0 @ x0) = truex:positivei:eltmem x (xfilter f l i~0) = true <-> mem x l = true /\ f (i @ x~0) = truef:elt -> booll:treeo:boolr:treeIHl:forall x i0 : elt, mem x (xfilter f l i0) = true <-> mem x l = true /\ f (i0 @ x) = trueIHr:forall x i0 : elt, mem x (xfilter f r i0) = true <-> mem x r = true /\ f (i0 @ x) = truei:elto &&& f (rev i) = true <-> o = true /\ f (i @ 1) = truef:elt -> booll:treeo:boolr:treeIHl:forall x0 i0 : elt, mem x0 (xfilter f l i0) = true <-> mem x0 l = true /\ f (i0 @ x0) = trueIHr:forall x0 i0 : elt, mem x0 (xfilter f r i0) = true <-> mem x0 r = true /\ f (i0 @ x0) = truex:positivei:eltmem x l = true /\ f (i~0 @ x) = true <-> mem x l = true /\ f (i @ x~0) = truef:elt -> booll:treeo:boolr:treeIHl:forall x i0 : elt, mem x (xfilter f l i0) = true <-> mem x l = true /\ f (i0 @ x) = trueIHr:forall x i0 : elt, mem x (xfilter f r i0) = true <-> mem x r = true /\ f (i0 @ x) = truei:elto &&& f (rev i) = true <-> o = true /\ f (i @ 1) = truef:elt -> booll:treeo:boolr:treeIHl:forall x i0 : elt, mem x (xfilter f l i0) = true <-> mem x l = true /\ f (i0 @ x) = trueIHr:forall x i0 : elt, mem x (xfilter f r i0) = true <-> mem x r = true /\ f (i0 @ x) = truei:elto &&& f (rev i) = true <-> o = true /\ f (i @ 1) = trueapply andb_true_iff. Qed.f:elt -> booll:treeo:boolr:treeIHl:forall x i0 : elt, mem x (xfilter f l i0) = true <-> mem x l = true /\ f (i0 @ x) = trueIHr:forall x i0 : elt, mem x (xfilter f r i0) = true <-> mem x r = true /\ f (i0 @ x) = truei:elto && f (rev i) = true <-> o = true /\ f (i @ 1) = trueforall (s : t) (x : elt) (f : elt -> bool), compat_bool E.eq f -> In x (filter f s) -> In x sforall (s : t) (x : elt) (f : elt -> bool), compat_bool E.eq f -> In x (filter f s) -> In x sforall (s : t) (x : elt) (f : elt -> bool), compat_bool E.eq f -> In x (xfilter f s 1) -> In x ss:tx:eltf:elt -> boolIn x (xfilter f s 1) -> In x stauto. Qed.s:tx:eltf:elt -> boolIn x s /\ f (1 @ x) = true -> In x sforall (s : t) (x : elt) (f : elt -> bool), compat_bool E.eq f -> In x (filter f s) -> f x = trueforall (s : t) (x : elt) (f : elt -> bool), compat_bool E.eq f -> In x (filter f s) -> f x = trueforall (s : t) (x : elt) (f : elt -> bool), compat_bool E.eq f -> In x (xfilter f s 1) -> f x = trues:tx:eltf:elt -> boolIn x (xfilter f s 1) -> f x = truetauto. Qed.s:tx:eltf:elt -> boolIn x s /\ f (1 @ x) = true -> f x = trueforall (s : t) (x : elt) (f : elt -> bool), compat_bool E.eq f -> In x s -> f x = true -> In x (filter f s)forall (s : t) (x : elt) (f : elt -> bool), compat_bool E.eq f -> In x s -> f x = true -> In x (filter f s)forall (s : t) (x : elt) (f : elt -> bool), compat_bool E.eq f -> In x s -> f x = true -> In x (xfilter f s 1)s:tx:eltf:elt -> boolIn x s -> f x = true -> In x (xfilter f s 1)tauto. Qed.s:tx:eltf:elt -> boolIn x s -> f x = true -> In x s /\ f (1 @ x) = true
Specification of for_all
forall (f : elt -> bool) (s : t) (i : elt), xforall f s i = true <-> For_all (fun x : elt => f (i @ x) = true) sforall (f : elt -> bool) (s : t) (i : elt), xforall f s i = true <-> For_all (fun x : elt => f (i @ x) = true) sforall (f : elt -> bool) (s : t) (i : elt), xforall f s i = true <-> (forall x : elt, mem x s = true -> f (i @ x) = true)f:elt -> boolforall (s : t) (i : elt), xforall f s i = true <-> (forall x : elt, mem x s = true -> f (i @ x) = true)f:elt -> booli:elttrue = true <-> (forall x : elt, false = true -> f (i @ x) = true)f:elt -> booll:treeo:boolr:treeIHl:forall i0 : elt, xforall f l i0 = true <-> (forall x : elt, mem x l = true -> f (i0 @ x) = true)IHr:forall i0 : elt, xforall f r i0 = true <-> (forall x : elt, mem x r = true -> f (i0 @ x) = true)i:elt(negb o ||| f (rev i)) &&& xforall f r i~1 &&& xforall f l i~0 = true <-> (forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true)f:elt -> booll:treeo:boolr:treeIHl:forall i0 : elt, xforall f l i0 = true <-> (forall x : elt, mem x l = true -> f (i0 @ x) = true)IHr:forall i0 : elt, xforall f r i0 = true <-> (forall x : elt, mem x r = true -> f (i0 @ x) = true)i:elt(negb o ||| f (rev i)) &&& xforall f r i~1 &&& xforall f l i~0 = true <-> (forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true)f:elt -> booll:treeo:boolr:treeIHl:forall i0 : elt, xforall f l i0 = true <-> (forall x : elt, mem x l = true -> f (i0 @ x) = true)IHr:forall i0 : elt, xforall f r i0 = true <-> (forall x : elt, mem x r = true -> f (i0 @ x) = true)i:elt(negb o || f (rev i) = true /\ xforall f r i~1 = true) /\ xforall f l i~0 = true <-> (forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true)f:elt -> booll:treeo:boolr:treeIHl:forall i0 : elt, xforall f l i0 = true <-> (forall x : elt, mem x l = true -> f (i0 @ x) = true)IHr:forall i0 : elt, xforall f r i0 = true <-> (forall x : elt, mem x r = true -> f (i0 @ x) = true)i:elt(negb o || f (rev i) = true /\ (forall x : elt, mem x r = true -> f (i~1 @ x) = true)) /\ (forall x : elt, mem x l = true -> f (i~0 @ x) = true) <-> (forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true)f:elt -> booll:treeo:boolr:treei:elt(negb o || f (rev i) = true /\ (forall x : elt, mem x r = true -> f (i~1 @ x) = true)) /\ (forall x : elt, mem x l = true -> f (i~0 @ x) = true) <-> (forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true)f:elt -> booll:treeo:boolr:treei:elt(negb o || f (rev i) = true /\ (forall x : elt, mem x r = true -> f (i~1 @ x) = true)) /\ (forall x : elt, mem x l = true -> f (i~0 @ x) = true) -> forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = truef:elt -> booll:treeo:boolr:treei:elt(forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true) -> (negb o || f (rev i) = true /\ (forall x : elt, mem x r = true -> f (i~1 @ x) = true)) /\ (forall x : elt, mem x l = true -> f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei:eltHi:negb o || f (rev i) = trueHr:forall x0 : elt, mem x0 r = true -> f (i~1 @ x0) = trueHl:forall x0 : elt, mem x0 l = true -> f (i~0 @ x0) = truex:eltmatch x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = truef:elt -> booll:treeo:boolr:treei:elt(forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true) -> (negb o || f (rev i) = true /\ (forall x : elt, mem x r = true -> f (i~1 @ x) = true)) /\ (forall x : elt, mem x l = true -> f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei:eltHi:negb o || f (rev i) = trueHr:forall x0 : elt, mem x0 r = true -> f (i~1 @ x0) = trueHl:forall x0 : elt, mem x0 l = true -> f (i~0 @ x0) = truex:positiveH:mem x r = truef (i @ x~1) = truef:elt -> booll:treeo:boolr:treei:eltHi:negb o || f (rev i) = trueHr:forall x0 : elt, mem x0 r = true -> f (i~1 @ x0) = trueHl:forall x0 : elt, mem x0 l = true -> f (i~0 @ x0) = truex:positiveH:mem x l = truef (i @ x~0) = truef:elt -> booll:treeo:boolr:treei:eltHi:negb o || f (rev i) = trueHr:forall x : elt, mem x r = true -> f (i~1 @ x) = trueHl:forall x : elt, mem x l = true -> f (i~0 @ x) = trueH:o = truef (i @ 1) = truef:elt -> booll:treeo:boolr:treei:elt(forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true) -> (negb o || f (rev i) = true /\ (forall x : elt, mem x r = true -> f (i~1 @ x) = true)) /\ (forall x : elt, mem x l = true -> f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei:eltHi:negb o || f (rev i) = trueHr:forall x0 : elt, mem x0 r = true -> f (i~1 @ x0) = trueHl:forall x0 : elt, mem x0 l = true -> f (i~0 @ x0) = truex:positiveH:mem x l = truef (i @ x~0) = truef:elt -> booll:treeo:boolr:treei:eltHi:negb o || f (rev i) = trueHr:forall x : elt, mem x r = true -> f (i~1 @ x) = trueHl:forall x : elt, mem x l = true -> f (i~0 @ x) = trueH:o = truef (i @ 1) = truef:elt -> booll:treeo:boolr:treei:elt(forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true) -> (negb o || f (rev i) = true /\ (forall x : elt, mem x r = true -> f (i~1 @ x) = true)) /\ (forall x : elt, mem x l = true -> f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei:eltHi:negb o || f (rev i) = trueHr:forall x : elt, mem x r = true -> f (i~1 @ x) = trueHl:forall x : elt, mem x l = true -> f (i~0 @ x) = trueH:o = truef (i @ 1) = truef:elt -> booll:treeo:boolr:treei:elt(forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true) -> (negb o || f (rev i) = true /\ (forall x : elt, mem x r = true -> f (i~1 @ x) = true)) /\ (forall x : elt, mem x l = true -> f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei:eltHi:negb true || f (rev i) = trueHr:forall x : elt, mem x r = true -> f (i~1 @ x) = trueHl:forall x : elt, mem x l = true -> f (i~0 @ x) = trueH:o = truef (i @ 1) = truef:elt -> booll:treeo:boolr:treei:elt(forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true) -> (negb o || f (rev i) = true /\ (forall x : elt, mem x r = true -> f (i~1 @ x) = true)) /\ (forall x : elt, mem x l = true -> f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei:elt(forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = true) -> (negb o || f (rev i) = true /\ (forall x : elt, mem x r = true -> f (i~1 @ x) = true)) /\ (forall x : elt, mem x l = true -> f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei:eltH:forall x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x) = truenegb o || f (rev i) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x r = truef (i~1 @ x) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x l = truef (i~0 @ x) = truef:elt -> booll:treeo:boolr:treei:eltH:match 1 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ 1) = truenegb o || f (rev i) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x r = truef (i~1 @ x) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x l = truef (i~0 @ x) = truef:elt -> booll, r:treei:eltH:true = true -> f (i @ 1) = truenegb true || f (rev i) = truef:elt -> booll, r:treei:eltH:false = true -> f (i @ 1) = truenegb false || f (rev i) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x r = truef (i~1 @ x) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x l = truef (i~0 @ x) = truef:elt -> booll, r:treei:eltH:true = true -> f (i @ 1) = truetrue = truef:elt -> booll, r:treei:eltH:false = true -> f (i @ 1) = truenegb false || f (rev i) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x r = truef (i~1 @ x) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x l = truef (i~0 @ x) = truef:elt -> booll, r:treei:eltH:false = true -> f (i @ 1) = truenegb false || f (rev i) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x r = truef (i~1 @ x) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x l = truef (i~0 @ x) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x r = truef (i~1 @ x) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x l = truef (i~0 @ x) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x r = truemem x r = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x l = truef (i~0 @ x) = truef:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x l = truef (i~0 @ x) = trueassumption. Qed.f:elt -> booll:treeo:boolr:treei:eltH:forall x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true -> f (i @ x0) = truex:eltH0:mem x l = truemem x l = trueforall (s : t) (f : elt -> bool), compat_bool E.eq f -> For_all (fun x : elt => f x = true) s -> for_all f s = trueforall (s : t) (f : elt -> bool), compat_bool E.eq f -> For_all (fun x : elt => f x = true) s -> for_all f s = trues:tf:elt -> boolFor_all (fun x : elt => f x = true) s -> for_all f s = trues:tf:elt -> boolFor_all (fun x : elt => f x = true) s -> xforall f s 1 = truetrivial. Qed.s:tf:elt -> boolFor_all (fun x : elt => f x = true) s -> For_all (fun x : elt => f (1 @ x) = true) sforall (s : t) (f : elt -> bool), compat_bool E.eq f -> for_all f s = true -> For_all (fun x : elt => f x = true) sforall (s : t) (f : elt -> bool), compat_bool E.eq f -> for_all f s = true -> For_all (fun x : elt => f x = true) ss:tf:elt -> boolfor_all f s = true -> For_all (fun x : elt => f x = true) ss:tf:elt -> boolxforall f s 1 = true -> For_all (fun x : elt => f x = true) strivial. Qed.s:tf:elt -> boolFor_all (fun x : elt => f (1 @ x) = true) s -> For_all (fun x : elt => f x = true) s
Specification of ∃
forall (f : elt -> bool) (s : t) (i : elt), xexists f s i = true <-> Exists (fun x : elt => f (i @ x) = true) sforall (f : elt -> bool) (s : t) (i : elt), xexists f s i = true <-> Exists (fun x : elt => f (i @ x) = true) sforall (f : elt -> bool) (s : t) (i : elt), xexists f s i = true <-> (exists x : elt, mem x s = true /\ f (i @ x) = true)f:elt -> boolforall (s : t) (i : elt), xexists f s i = true <-> (exists x : elt, mem x s = true /\ f (i @ x) = true)f:elt -> booli:eltfalse = true <-> (exists x : elt, false = true /\ f (i @ x) = true)f:elt -> booll:treeo:boolr:treeIHl:forall i0 : elt, xexists f l i0 = true <-> (exists x : elt, mem x l = true /\ f (i0 @ x) = true)IHr:forall i0 : elt, xexists f r i0 = true <-> (exists x : elt, mem x r = true /\ f (i0 @ x) = true)i:elto &&& f (rev i) ||| xexists f r i~1 ||| xexists f l i~0 = true <-> (exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true)f:elt -> booll:treeo:boolr:treeIHl:forall i0 : elt, xexists f l i0 = true <-> (exists x : elt, mem x l = true /\ f (i0 @ x) = true)IHr:forall i0 : elt, xexists f r i0 = true <-> (exists x : elt, mem x r = true /\ f (i0 @ x) = true)i:elto &&& f (rev i) ||| xexists f r i~1 ||| xexists f l i~0 = true <-> (exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true)f:elt -> booll:treeo:boolr:treeIHl:forall i0 : elt, xexists f l i0 = true <-> (exists x : elt, mem x l = true /\ f (i0 @ x) = true)IHr:forall i0 : elt, xexists f r i0 = true <-> (exists x : elt, mem x r = true /\ f (i0 @ x) = true)i:elt(o = true /\ f (rev i) = true \/ xexists f r i~1 = true) \/ xexists f l i~0 = true <-> (exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true)f:elt -> booll:treeo:boolr:treeIHl:forall i0 : elt, xexists f l i0 = true <-> (exists x : elt, mem x l = true /\ f (i0 @ x) = true)IHr:forall i0 : elt, xexists f r i0 = true <-> (exists x : elt, mem x r = true /\ f (i0 @ x) = true)i:elt(o = true /\ f (rev i) = true \/ (exists x : elt, mem x r = true /\ f (i~1 @ x) = true)) \/ (exists x : elt, mem x l = true /\ f (i~0 @ x) = true) <-> (exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true)f:elt -> booll:treeo:boolr:treei:elt(o = true /\ f (rev i) = true \/ (exists x : elt, mem x r = true /\ f (i~1 @ x) = true)) \/ (exists x : elt, mem x l = true /\ f (i~0 @ x) = true) <-> (exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true)f:elt -> booll:treeo:boolr:treei:elt(o = true /\ f (rev i) = true \/ (exists x : elt, mem x r = true /\ f (i~1 @ x) = true)) \/ (exists x : elt, mem x l = true /\ f (i~0 @ x) = true) -> exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = truef:elt -> booll:treeo:boolr:treei:elt(exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true) -> (o = true /\ f (rev i) = true \/ (exists x : elt, mem x r = true /\ f (i~1 @ x) = true)) \/ (exists x : elt, mem x l = true /\ f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei:eltHi:o = true /\ f (rev i) = trueexists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = truef:elt -> booll:treeo:boolr:treei, x:eltHr:mem x r = true /\ f (i~1 @ x) = trueexists x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x0) = truef:elt -> booll:treeo:boolr:treei, x:eltHl:mem x l = true /\ f (i~0 @ x) = trueexists x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x0) = truef:elt -> booll:treeo:boolr:treei:elt(exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true) -> (o = true /\ f (rev i) = true \/ (exists x : elt, mem x r = true /\ f (i~1 @ x) = true)) \/ (exists x : elt, mem x l = true /\ f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei:eltHi:o = true /\ f (rev i) = trueo = true /\ f (i @ 1) = truef:elt -> booll:treeo:boolr:treei, x:eltHr:mem x r = true /\ f (i~1 @ x) = trueexists x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x0) = truef:elt -> booll:treeo:boolr:treei, x:eltHl:mem x l = true /\ f (i~0 @ x) = trueexists x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x0) = truef:elt -> booll:treeo:boolr:treei:elt(exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true) -> (o = true /\ f (rev i) = true \/ (exists x : elt, mem x r = true /\ f (i~1 @ x) = true)) \/ (exists x : elt, mem x l = true /\ f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei, x:eltHr:mem x r = true /\ f (i~1 @ x) = trueexists x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x0) = truef:elt -> booll:treeo:boolr:treei, x:eltHl:mem x l = true /\ f (i~0 @ x) = trueexists x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x0) = truef:elt -> booll:treeo:boolr:treei:elt(exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true) -> (o = true /\ f (rev i) = true \/ (exists x : elt, mem x r = true /\ f (i~1 @ x) = true)) \/ (exists x : elt, mem x l = true /\ f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei, x:eltHr:mem x r = true /\ f (i~1 @ x) = truemem x r = true /\ f (i @ x~1) = truef:elt -> booll:treeo:boolr:treei, x:eltHl:mem x l = true /\ f (i~0 @ x) = trueexists x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x0) = truef:elt -> booll:treeo:boolr:treei:elt(exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true) -> (o = true /\ f (rev i) = true \/ (exists x : elt, mem x r = true /\ f (i~1 @ x) = true)) \/ (exists x : elt, mem x l = true /\ f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei, x:eltHl:mem x l = true /\ f (i~0 @ x) = trueexists x0 : elt, match x0 with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x0) = truef:elt -> booll:treeo:boolr:treei:elt(exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true) -> (o = true /\ f (rev i) = true \/ (exists x : elt, mem x r = true /\ f (i~1 @ x) = true)) \/ (exists x : elt, mem x l = true /\ f (i~0 @ x) = true)f:elt -> booll:treeo:boolr:treei, x:eltHl:mem x l = true /\ f (i~0 @ x) = truemem x l = true /\ f (i @ x~0) = truef:elt -> booll:treeo:boolr:treei:elt(exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true) -> (o = true /\ f (rev i) = true \/ (exists x : elt, mem x r = true /\ f (i~1 @ x) = true)) \/ (exists x : elt, mem x l = true /\ f (i~0 @ x) = true)intros [[x|x|] H]; eauto. Qed.f:elt -> booll:treeo:boolr:treei:elt(exists x : elt, match x with | i0~1 => mem i0 r | i0~0 => mem i0 l | 1 => o end = true /\ f (i @ x) = true) -> (o = true /\ f (rev i) = true \/ (exists x : elt, mem x r = true /\ f (i~1 @ x) = true)) \/ (exists x : elt, mem x l = true /\ f (i~0 @ x) = true)forall (s : t) (f : elt -> bool), compat_bool E.eq f -> Exists (fun x : elt => f x = true) s -> exists_ f s = trueforall (s : t) (f : elt -> bool), compat_bool E.eq f -> Exists (fun x : elt => f x = true) s -> exists_ f s = trues:tf:elt -> boolExists (fun x : elt => f x = true) s -> exists_ f s = trues:tf:elt -> boolExists (fun x : elt => f x = true) s -> xexists f s 1 = truetrivial. Qed.s:tf:elt -> boolExists (fun x : elt => f x = true) s -> Exists (fun x : elt => f (1 @ x) = true) sforall (s : t) (f : elt -> bool), compat_bool E.eq f -> exists_ f s = true -> Exists (fun x : elt => f x = true) sforall (s : t) (f : elt -> bool), compat_bool E.eq f -> exists_ f s = true -> Exists (fun x : elt => f x = true) ss:tf:elt -> boolexists_ f s = true -> Exists (fun x : elt => f x = true) ss:tf:elt -> boolxexists f s 1 = true -> Exists (fun x : elt => f x = true) strivial. Qed.s:tf:elt -> boolExists (fun x : elt => f (1 @ x) = true) s -> Exists (fun x : elt => f x = true) s
Specification of partition
forall (s : t) (f : elt -> bool), partition f s = (filter f s, filter (fun x : elt => negb (f x)) s)forall (s : t) (f : elt -> bool), partition f s = (filter f s, filter (fun x : elt => negb (f x)) s)forall (s : t) (f : elt -> bool), xpartition f s 1 = (xfilter f s 1, xfilter (fun x : elt => negb (f x)) s 1)s:tf:elt -> boolxpartition f s 1 = (xfilter f s 1, xfilter (fun x : elt => negb (f x)) s 1)s:tf:elt -> boolforall j : positive, xpartition f s j = (xfilter f s j, xfilter (fun x : elt => negb (f x)) s j)f:elt -> boolj:positivexpartition f Leaf j = (xfilter f Leaf j, xfilter (fun x : elt => negb (f x)) Leaf j)l:treeo:boolr:treef:elt -> boolIHl:forall j0 : positive, xpartition f l j0 = (xfilter f l j0, xfilter (fun x : elt => negb (f x)) l j0)IHr:forall j0 : positive, xpartition f r j0 = (xfilter f r j0, xfilter (fun x : elt => negb (f x)) r j0)j:positivexpartition f (Node l o r) j = (xfilter f (Node l o r) j, xfilter (fun x : elt => negb (f x)) (Node l o r) j)destruct o; simpl; rewrite IHl, IHr; reflexivity. Qed.l:treeo:boolr:treef:elt -> boolIHl:forall j0 : positive, xpartition f l j0 = (xfilter f l j0, xfilter (fun x : elt => negb (f x)) l j0)IHr:forall j0 : positive, xpartition f r j0 = (xfilter f r j0, xfilter (fun x : elt => negb (f x)) r j0)j:positivexpartition f (Node l o r) j = (xfilter f (Node l o r) j, xfilter (fun x : elt => negb (f x)) (Node l o r) j)forall (s : t) (f : elt -> bool), compat_bool E.eq f -> fst (partition f s) [=] filter f sforall (s : t) (f : elt -> bool), compat_bool E.eq f -> fst (partition f s) [=] filter f ss:tf:elt -> boolH:compat_bool E.eq ffst (partition f s) [=] filter f sapply eq_refl. Qed.s:tf:elt -> boolH:compat_bool E.eq ffst (filter f s, filter (fun x : elt => negb (f x)) s) [=] filter f sforall (s : t) (f : elt -> bool), compat_bool E.eq f -> snd (partition f s) [=] filter (fun x : elt => negb (f x)) sforall (s : t) (f : elt -> bool), compat_bool E.eq f -> snd (partition f s) [=] filter (fun x : elt => negb (f x)) ss:tf:elt -> boolH:compat_bool E.eq fsnd (partition f s) [=] filter (fun x : elt => negb (f x)) sapply eq_refl. Qed.s:tf:elt -> boolH:compat_bool E.eq fsnd (filter f s, filter (fun x : elt => negb (f x)) s) [=] filter (fun x : elt => negb (f x)) s
Specification of elements
Notation InL := (InA E.eq).forall (s : t) (j : elt) (acc : list elt) (y : positive), InL y (xelements s j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x s = true)forall (s : t) (j : elt) (acc : list elt) (y : positive), InL y (xelements s j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x s = true)forall (j : elt) (acc : list elt) (y : positive), InL y acc <-> InL y acc \/ (exists x : elt, y = j @ x /\ false = true)l:treeo:boolr:treeIHl:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements l j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x l = true)IHr:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements r j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x r = true)forall (j : elt) (acc : list elt) (y : positive), InL y (if o then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = true)j:eltacc:list elty:positiveInL y acc <-> InL y acc \/ (exists x : elt, y = j @ x /\ false = true)l:treeo:boolr:treeIHl:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements l j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x l = true)IHr:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements r j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x r = true)forall (j : elt) (acc : list elt) (y : positive), InL y (if o then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = true)j:eltacc:list elty:positiveH:InL y accInL y acc \/ (exists x : elt, y = j @ x /\ false = true)j:eltacc:list elty:positiveH:InL y acc \/ (exists x : elt, y = j @ x /\ false = true)InL y accl:treeo:boolr:treeIHl:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements l j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x l = true)IHr:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements r j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x r = true)forall (j : elt) (acc : list elt) (y : positive), InL y (if o then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = true)j:eltacc:list elty:positiveH:InL y accInL y accj:eltacc:list elty:positiveH:InL y acc \/ (exists x : elt, y = j @ x /\ false = true)InL y accl:treeo:boolr:treeIHl:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements l j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x l = true)IHr:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements r j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x r = true)forall (j : elt) (acc : list elt) (y : positive), InL y (if o then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = true)j:eltacc:list elty:positiveH:InL y acc \/ (exists x : elt, y = j @ x /\ false = true)InL y accl:treeo:boolr:treeIHl:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements l j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x l = true)IHr:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements r j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x r = true)forall (j : elt) (acc : list elt) (y : positive), InL y (if o then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = true)j:eltacc:list elty:positiveH:InL y accInL y accj:eltacc:list elty:positivex:eltHx:y = j @ xHx':false = trueInL y accl:treeo:boolr:treeIHl:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements l j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x l = true)IHr:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements r j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x r = true)forall (j : elt) (acc : list elt) (y : positive), InL y (if o then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = true)j:eltacc:list elty:positivex:eltHx:y = j @ xHx':false = trueInL y accl:treeo:boolr:treeIHl:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements l j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x l = true)IHr:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements r j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x r = true)forall (j : elt) (acc : list elt) (y : positive), InL y (if o then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = true)l:treeo:boolr:treeIHl:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements l j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x l = true)IHr:forall (j : elt) (acc : list elt) (y : positive), InL y (xelements r j acc) <-> InL y acc \/ (exists x : elt, y = j @ x /\ mem x r = true)forall (j : elt) (acc : list elt) (y : positive), InL y (if o then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (if o then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (rev j :: xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (rev j :: xelements r j~1 acc) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positive(E.eq y (rev j) \/ InL y (xelements r j~1 acc)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positive(E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positive(E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positive(E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true) -> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true) -> (E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x r = trueInL (j~1 @ x) acc \/ (exists x0 : elt, j~1 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = trueInL (j~0 @ x) acc \/ (exists x0 : elt, j~0 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true) -> (E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x r = trueexists x0 : elt, j~1 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => true end = truel:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = trueInL (j~0 @ x) acc \/ (exists x0 : elt, j~0 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true) -> (E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x r = truej~1 @ x = j @ x~1 /\ mem x r = truel:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = trueInL (j~0 @ x) acc \/ (exists x0 : elt, j~0 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true) -> (E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = trueInL (j~0 @ x) acc \/ (exists x0 : elt, j~0 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true) -> (E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = trueexists x0 : elt, j~0 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => true end = truel:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true) -> (E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = truej~0 @ x = j @ x~0 /\ mem x l = truel:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true) -> (E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true) -> (E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positiveH:InL y acc(E.eq y (rev j) \/ InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true(E.eq (j @ x) (rev j) \/ InL (j @ x) acc \/ (exists x0 : elt, j @ x = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:match x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = true(E.eq (j @ x) (rev j) \/ InL (j @ x) acc \/ (exists x0 : elt, j @ x = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x r = true(E.eq (j @ x~1) (rev j) \/ InL (j @ x~1) acc \/ (exists x0 : elt, j @ x~1 = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x~1 = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x l = true(E.eq (j @ x~0) (rev j) \/ InL (j @ x~0) acc \/ (exists x0 : elt, j @ x~0 = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x~0 = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltH:true = true(E.eq (j @ 1) (rev j) \/ InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x r = trueE.eq (j @ x~1) (rev j) \/ InL (j @ x~1) acc \/ (exists x0 : elt, j @ x~1 = j~1 @ x0 /\ mem x0 r = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x l = true(E.eq (j @ x~0) (rev j) \/ InL (j @ x~0) acc \/ (exists x0 : elt, j @ x~0 = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x~0 = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltH:true = true(E.eq (j @ 1) (rev j) \/ InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x r = trueInL (j @ x~1) acc \/ (exists x0 : elt, j @ x~1 = j~1 @ x0 /\ mem x0 r = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x l = true(E.eq (j @ x~0) (rev j) \/ InL (j @ x~0) acc \/ (exists x0 : elt, j @ x~0 = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x~0 = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltH:true = true(E.eq (j @ 1) (rev j) \/ InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x r = trueexists x0 : elt, j @ x~1 = j~1 @ x0 /\ mem x0 r = truel:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x l = true(E.eq (j @ x~0) (rev j) \/ InL (j @ x~0) acc \/ (exists x0 : elt, j @ x~0 = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x~0 = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltH:true = true(E.eq (j @ 1) (rev j) \/ InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x l = true(E.eq (j @ x~0) (rev j) \/ InL (j @ x~0) acc \/ (exists x0 : elt, j @ x~0 = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x~0 = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltH:true = true(E.eq (j @ 1) (rev j) \/ InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x l = trueexists x0 : elt, j @ x~0 = j~0 @ x0 /\ mem x0 l = truel:treeo:boolr:treej:eltacc:list eltH:true = true(E.eq (j @ 1) (rev j) \/ InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltH:true = true(E.eq (j @ 1) (rev j) \/ InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltH:true = trueE.eq (j @ 1) (rev j) \/ InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltH:true = trueE.eq (j @ 1) (rev j)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positiveInL y (xelements l j~0 (xelements r j~1 acc)) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements l j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x l = true)IHr:forall (j0 : elt) (acc0 : list elt) (y0 : positive), InL y0 (xelements r j0 acc0) <-> InL y0 acc0 \/ (exists x : elt, y0 = j0 @ x /\ mem x r = true)j:eltacc:list elty:positive(InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positive(InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true) <-> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positive(InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true) -> InL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true) -> (InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list elty:positiveH:InL y accInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x r = trueInL (j~1 @ x) acc \/ (exists x0 : elt, j~1 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = trueInL (j~0 @ x) acc \/ (exists x0 : elt, j~0 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true) -> (InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x r = trueInL (j~1 @ x) acc \/ (exists x0 : elt, j~1 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = trueInL (j~0 @ x) acc \/ (exists x0 : elt, j~0 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true) -> (InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x r = trueexists x0 : elt, j~1 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = trueInL (j~0 @ x) acc \/ (exists x0 : elt, j~0 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true) -> (InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x r = truej~1 @ x = j @ x~1 /\ mem x r = truel:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = trueInL (j~0 @ x) acc \/ (exists x0 : elt, j~0 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true) -> (InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = trueInL (j~0 @ x) acc \/ (exists x0 : elt, j~0 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true) -> (InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = trueexists x0 : elt, j~0 @ x = j @ x0 /\ match x0 with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true) -> (InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:mem x l = truej~0 @ x = j @ x~0 /\ mem x l = truel:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true) -> (InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list elty:positiveInL y acc \/ (exists x : elt, y = j @ x /\ match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true) -> (InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list elty:positiveH:InL y acc(InL y acc \/ (exists x : elt, y = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, y = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true(InL (j @ x) acc \/ (exists x0 : elt, j @ x = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltx:eltH:match x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = true(InL (j @ x) acc \/ (exists x0 : elt, j @ x = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x r = true(InL (j @ x~1) acc \/ (exists x0 : elt, j @ x~1 = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x~1 = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x l = true(InL (j @ x~0) acc \/ (exists x0 : elt, j @ x~0 = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x~0 = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltH:false = true(InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x r = trueInL (j @ x~1) acc \/ (exists x0 : elt, j @ x~1 = j~1 @ x0 /\ mem x0 r = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x l = true(InL (j @ x~0) acc \/ (exists x0 : elt, j @ x~0 = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x~0 = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltH:false = true(InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x r = trueexists x0 : elt, j @ x~1 = j~1 @ x0 /\ mem x0 r = truel:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x l = true(InL (j @ x~0) acc \/ (exists x0 : elt, j @ x~0 = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x~0 = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltH:false = true(InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x l = true(InL (j @ x~0) acc \/ (exists x0 : elt, j @ x~0 = j~1 @ x0 /\ mem x0 r = true)) \/ (exists x0 : elt, j @ x~0 = j~0 @ x0 /\ mem x0 l = true)l:treeo:boolr:treej:eltacc:list eltH:false = true(InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)l:treeo:boolr:treej:eltacc:list eltx:positiveH:mem x l = trueexists x0 : elt, j @ x~0 = j~0 @ x0 /\ mem x0 l = truel:treeo:boolr:treej:eltacc:list eltH:false = true(InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)discriminate. Qed.l:treeo:boolr:treej:eltacc:list eltH:false = true(InL (j @ 1) acc \/ (exists x : elt, j @ 1 = j~1 @ x /\ mem x r = true)) \/ (exists x : elt, j @ 1 = j~0 @ x /\ mem x l = true)forall (s : t) (x : elt), In x s -> InL x (elements s)forall (s : t) (x : elt), In x s -> InL x (elements s)forall (s : t) (x : elt), mem x s = true -> InL x (xelements s 1 nil)s:tx:eltH:mem x s = trueInL x (xelements s 1 nil)s:tx:eltH:mem x s = trueInL x nil \/ (exists x0 : elt, x = 1 @ x0 /\ mem x0 s = true)s:tx:eltH:mem x s = trueexists x0 : elt, x = 1 @ x0 /\ mem x0 s = trueauto. Qed.s:tx:eltH:mem x s = truex = 1 @ x /\ mem x s = trueforall (s : t) (x : positive), InL x (elements s) -> In x sforall (s : t) (x : positive), InL x (elements s) -> In x sforall (s : t) (x : positive), InL x (xelements s 1 nil) -> mem x s = trues:tx:positiveH:InL x (xelements s 1 nil)mem x s = trues:tx:positiveH:InL x nil \/ (exists x0 : elt, x = 1 @ x0 /\ mem x0 s = true)mem x s = trues:tx:positiveH:InL x nilmem x s = trues:tx:positivey:eltH:x = 1 @ yH':mem y s = truemem x s = trues:tx:positivey:eltH:x = 1 @ yH':mem y s = truemem x s = trueassumption. Qed.s:tx:positivey:eltH:x = 1 @ yH':mem y s = truemem (1 @ y) s = trueforall (j : elt) (x y : positive), E.lt x y -> E.lt (j @ x) (j @ y)induction j; intros; simpl; auto. Qed.forall (j : elt) (x y : positive), E.lt x y -> E.lt (j @ x) (j @ y)forall s : t, Sorted E.lt (elements s)forall s : t, Sorted E.lt (elements s)forall s : t, Sorted E.lt (xelements s 1 nil)forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l:treeo:boolr:treeIHl:forall (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x l -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements l j acc)IHr:forall (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x r -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements r j acc)forall (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x (Node l o r) -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (if o then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l:treeo:boolr:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l o r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (if o then xelements l j~0 (rev j :: xelements r j~1 acc) else xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (rev j :: xelements r j~1 acc))l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (rev j :: xelements r j~1 acc)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements r j~1 acc)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yHdRel E.lt (rev j) (xelements r j~1 acc)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) ySorted E.lt accl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x r -> InL y acc -> E.lt (j~1 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yHdRel E.lt (rev j) (xelements r j~1 acc)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x r -> InL y acc -> E.lt (j~1 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yHdRel E.lt (rev j) (xelements r j~1 acc)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x rHy:InL y accE.lt (j~1 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yHdRel E.lt (rev j) (xelements r j~1 acc)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yHdRel E.lt (rev j) (xelements r j~1 acc)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yxelements r j~1 acc = nil -> HdRel E.lt (rev j) nill, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (e : elt) (l0 : list elt), xelements r j~1 acc = e :: l0 -> HdRel E.lt (rev j) (e :: l0)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (e : elt) (l0 : list elt), xelements r j~1 acc = e :: l0 -> HdRel E.lt (rev j) (e :: l0)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltq:list eltH:xelements r j~1 acc = z :: qHdRel E.lt (rev j) (z :: q)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltq:list eltH:xelements r j~1 acc = z :: qE.lt (rev j) zl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltq:list eltH:xelements r j~1 acc = z :: qInL z (xelements r j~1 acc)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltq:list eltH:xelements r j~1 acc = z :: qH':InL z (xelements r j~1 acc)E.lt (rev j) zl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltq:list eltH:xelements r j~1 acc = z :: qInL z (z :: q)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltq:list eltH:xelements r j~1 acc = z :: qH':InL z (xelements r j~1 acc)E.lt (rev j) zl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltq:list eltH:xelements r j~1 acc = z :: qE.eq z zl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltq:list eltH:xelements r j~1 acc = z :: qH':InL z (xelements r j~1 acc)E.lt (rev j) zl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltq:list eltH:xelements r j~1 acc = z :: qH':InL z (xelements r j~1 acc)E.lt (rev j) zl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltH':InL z (xelements r j~1 acc)E.lt (rev j) zl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltH':InL z acc \/ (exists x : elt, z = j~1 @ x /\ mem x r = true)E.lt (rev j) zl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltHy:InL z accE.lt (rev j) zl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l true r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:mem x r = trueE.lt (rev j) (j~1 @ x)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yz:eltHy:InL z accIn 1 (Node l true r)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l true r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:mem x r = trueE.lt (rev j) (j~1 @ x)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l true r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:mem x r = trueE.lt (rev j) (j~1 @ x)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l true r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:mem x r = trueE.lt (rev j) (j @ x~1)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l true r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:mem x r = trueE.lt 1 x~1l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l true r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (rev j :: xelements r j~1 acc) -> E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lHy:InL y (rev j :: xelements r j~1 acc)E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lH:E.eq y (rev j)E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lH:InL y (xelements r j~1 acc)E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lH:E.eq y (rev j)E.lt (j~0 @ x) (rev j)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lH:InL y (xelements r j~1 acc)E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lH:E.eq y (rev j)E.lt (j @ x~0) (rev j)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lH:InL y (xelements r j~1 acc)E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lH:E.eq y (rev j)E.lt x~0 1l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lH:InL y (xelements r j~1 acc)E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lH:InL y (xelements r j~1 acc)E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lH:InL y acc \/ (exists x0 : elt, y = j~1 @ x0 /\ mem x0 r = true)E.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l true r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lHy:InL y accE.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l true r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:In x lz:eltHy:mem z r = trueE.lt (j~0 @ x) (j~1 @ z)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l true r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:In x lz:eltHy:mem z r = trueE.lt (j~0 @ x) (j~1 @ z)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l true r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:In x lz:eltHy:mem z r = trueE.lt (j @ x~0) (j @ z~1)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l true r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:In x lz:eltHy:mem z r = trueE.lt x~0 z~1l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements l j~0 (xelements r j~1 acc))H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt (xelements r j~1 acc)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (xelements r j~1 acc) -> E.lt (j~0 @ x) yH:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) ySorted E.lt accl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x r -> InL y acc -> E.lt (j~1 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (xelements r j~1 acc) -> E.lt (j~0 @ x) yH:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x r -> InL y acc -> E.lt (j~1 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (xelements r j~1 acc) -> E.lt (j~0 @ x) yH:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l false r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x rHy:InL y accE.lt (j~1 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (xelements r j~1 acc) -> E.lt (j~0 @ x) yH:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x l -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x : elt) (y : positive), In x r -> InL y acc0 -> E.lt (j0 @ x) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x : elt) (y : positive), In x (Node l false r) -> InL y acc -> E.lt (j @ x) yforall (x : elt) (y : positive), In x l -> InL y (xelements r j~1 acc) -> E.lt (j~0 @ x) yH:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l false r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lHy:InL y (xelements r j~1 acc)E.lt (j~0 @ x) yH:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l false r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lHy:InL y acc \/ (exists x0 : elt, y = j~1 @ x0 /\ mem x0 r = true)E.lt (j~0 @ x) yH:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 l -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y0 : positive), In x0 r -> InL y0 acc0 -> E.lt (j0 @ x0) y0) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y0 : positive), In x0 (Node l false r) -> InL y0 acc -> E.lt (j @ x0) y0x:elty:positiveHx:In x lHy:InL y accE.lt (j~0 @ x) yl, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l false r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:In x lz:eltHy:mem z r = trueE.lt (j~0 @ x) (j~1 @ z)H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l false r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:In x lz:eltHy:mem z r = trueE.lt (j~0 @ x) (j~1 @ z)H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l false r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:In x lz:eltHy:mem z r = trueE.lt (j @ x~0) (j @ z~1)H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)l, r:treeIHl:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 l -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements l j0 acc0)IHr:forall (j0 : elt) (acc0 : list positive), Sorted E.lt acc0 -> (forall (x0 : elt) (y : positive), In x0 r -> InL y acc0 -> E.lt (j0 @ x0) y) -> Sorted E.lt (xelements r j0 acc0)j:eltacc:list positiveHacc:Sorted E.lt accHsacc:forall (x0 : elt) (y : positive), In x0 (Node l false r) -> InL y acc -> E.lt (j @ x0) yx:eltHx:In x lz:eltHy:mem z r = trueE.lt x~0 z~1H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)H:forall (s : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s j acc)forall s : t, Sorted E.lt (xelements s 1 nil)H:forall (s0 : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s0 -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s0 j acc)s:tSorted E.lt (xelements s 1 nil)H:forall (s0 : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s0 -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s0 j acc)s:tSorted E.lt nilH:forall (s0 : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s0 -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s0 j acc)s:tforall (x : elt) (y : positive), In x s -> InL y nil -> E.lt (1 @ x) yH:forall (s0 : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x : elt) (y : positive), In x s0 -> InL y acc -> E.lt (j @ x) y) -> Sorted E.lt (xelements s0 j acc)s:tforall (x : elt) (y : positive), In x s -> InL y nil -> E.lt (1 @ x) yinversion H'. Qed.H:forall (s0 : t) (j : elt) (acc : list positive), Sorted E.lt acc -> (forall (x0 : elt) (y0 : positive), In x0 s0 -> InL y0 acc -> E.lt (j @ x0) y0) -> Sorted E.lt (xelements s0 j acc)s:tx:elty:positiveH':InL y nilE.lt (1 @ x) yforall s : t, NoDupA E.eq (elements s)forall s : t, NoDupA E.eq (elements s)s:tNoDupA E.eq (elements s)s:tEquivalence E.eqs:tStrictOrder E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tReflexive E.eqs:tSymmetric E.eqs:tTransitive E.eqs:tStrictOrder E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tx:positiveE.eq x xs:tSymmetric E.eqs:tTransitive E.eqs:tStrictOrder E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tSymmetric E.eqs:tTransitive E.eqs:tStrictOrder E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tx:positiveforall y : positive, E.eq x y -> E.eq y xs:tTransitive E.eqs:tStrictOrder E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tTransitive E.eqs:tStrictOrder E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tx:positiveforall y z : positive, E.eq x y -> E.eq y z -> E.eq x zs:tStrictOrder E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tStrictOrder E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tIrreflexive E.lts:tTransitive E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tx:positiveH:E.lt x xFalses:tTransitive E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tx:positiveH:~ E.eq x xFalses:tTransitive E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tx:positiveH:~ E.eq x xE.eq x xs:tTransitive E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tTransitive E.lts:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tx:positiveforall y z : positive, E.lt x y -> E.lt y z -> E.lt x zs:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)s:tProper (E.eq ==> E.eq ==> iff) E.lts:tSorted E.lt (elements s)apply elements_3. Qed.s:tSorted E.lt (elements s)
Specification of choose
forall (s : t) (x : elt), choose s = Some x -> In x sforall (s : t) (x : elt), choose s = Some x -> In x sforall x : elt, None = Some x -> In x Leafl:treeo:boolr:treeIHl:forall x : elt, choose l = Some x -> In x lIHr:forall x : elt, choose r = Some x -> In x rforall x : elt, (if o then Some 1 else match choose l with | Some i => Some i~0 | None => omap xI (choose r) end) = Some x -> In x (Node l o r)x:eltH:None = Some xIn x Leafl:treeo:boolr:treeIHl:forall x : elt, choose l = Some x -> In x lIHr:forall x : elt, choose r = Some x -> In x rforall x : elt, (if o then Some 1 else match choose l with | Some i => Some i~0 | None => omap xI (choose r) end) = Some x -> In x (Node l o r)l:treeo:boolr:treeIHl:forall x : elt, choose l = Some x -> In x lIHr:forall x : elt, choose r = Some x -> In x rforall x : elt, (if o then Some 1 else match choose l with | Some i => Some i~0 | None => omap xI (choose r) end) = Some x -> In x (Node l o r)l, r:treeIHl:forall x : elt, choose l = Some x -> In x lIHr:forall x : elt, choose r = Some x -> In x rforall x : elt, Some 1 = Some x -> In x (Node l true r)l, r:treeIHl:forall x : elt, choose l = Some x -> In x lIHr:forall x : elt, choose r = Some x -> In x rforall x : elt, match choose l with | Some i => Some i~0 | None => omap xI (choose r) end = Some x -> In x (Node l false r)l, r:treeIHl:forall x0 : elt, choose l = Some x0 -> In x0 lIHr:forall x0 : elt, choose r = Some x0 -> In x0 rx:eltH:Some 1 = Some xIn x (Node l true r)l, r:treeIHl:forall x : elt, choose l = Some x -> In x lIHr:forall x : elt, choose r = Some x -> In x rforall x : elt, match choose l with | Some i => Some i~0 | None => omap xI (choose r) end = Some x -> In x (Node l false r)l, r:treeIHl:forall x : elt, choose l = Some x -> In x lIHr:forall x : elt, choose r = Some x -> In x rH:Some 1 = Some 1In 1 (Node l true r)l, r:treeIHl:forall x : elt, choose l = Some x -> In x lIHr:forall x : elt, choose r = Some x -> In x rforall x : elt, match choose l with | Some i => Some i~0 | None => omap xI (choose r) end = Some x -> In x (Node l false r)l, r:treeIHl:forall x : elt, choose l = Some x -> In x lIHr:forall x : elt, choose r = Some x -> In x rforall x : elt, match choose l with | Some i => Some i~0 | None => omap xI (choose r) end = Some x -> In x (Node l false r)l, r:treeIHr:forall x : elt, choose r = Some x -> In x r(forall x : elt, choose l = Some x -> In x l) -> forall x : elt, match choose l with | Some i => Some i~0 | None => omap xI (choose r) end = Some x -> In x (Node l false r)l, r:treeIHr:forall x : elt, choose r = Some x -> In x rforall e : elt, (forall x : elt, Some e = Some x -> In x l) -> forall x : elt, Some e~0 = Some x -> In x (Node l false r)l, r:treeIHr:forall x : elt, choose r = Some x -> In x r(forall x : elt, None = Some x -> In x l) -> forall x : elt, omap xI (choose r) = Some x -> In x (Node l false r)l, r:treeIHr:forall x0 : elt, choose r = Some x0 -> In x0 rp:eltHp:forall x0 : elt, Some p = Some x0 -> In x0 lx:eltH:Some p~0 = Some xIn x (Node l false r)l, r:treeIHr:forall x : elt, choose r = Some x -> In x r(forall x : elt, None = Some x -> In x l) -> forall x : elt, omap xI (choose r) = Some x -> In x (Node l false r)l, r:treeIHr:forall x : elt, choose r = Some x -> In x rp:eltHp:forall x : elt, Some p = Some x -> In x lIn p~0 (Node l false r)l, r:treeIHr:forall x : elt, choose r = Some x -> In x r(forall x : elt, None = Some x -> In x l) -> forall x : elt, omap xI (choose r) = Some x -> In x (Node l false r)l, r:treeIHr:forall x : elt, choose r = Some x -> In x rp:eltHp:forall x : elt, Some p = Some x -> In x lSome p = Some pl, r:treeIHr:forall x : elt, choose r = Some x -> In x r(forall x : elt, None = Some x -> In x l) -> forall x : elt, omap xI (choose r) = Some x -> In x (Node l false r)l, r:treeIHr:forall x : elt, choose r = Some x -> In x r(forall x : elt, None = Some x -> In x l) -> forall x : elt, omap xI (choose r) = Some x -> In x (Node l false r)l, r:treeIHr:forall x0 : elt, choose r = Some x0 -> In x0 rx:eltomap xI (choose r) = Some x -> In x (Node l false r)l, r:treex:elt(forall x0 : elt, choose r = Some x0 -> In x0 r) -> omap xI (choose r) = Some x -> In x (Node l false r)l, r:treex:eltforall e : elt, (forall x0 : elt, Some e = Some x0 -> In x0 r) -> omap xI (Some e) = Some x -> In x (Node l false r)l, r:treex:elt(forall x0 : elt, None = Some x0 -> In x0 r) -> omap xI None = Some x -> In x (Node l false r)l, r:treex, p:eltHp:forall x0 : elt, Some p = Some x0 -> In x0 rH:omap xI (Some p) = Some xIn x (Node l false r)l, r:treex:elt(forall x0 : elt, None = Some x0 -> In x0 r) -> omap xI None = Some x -> In x (Node l false r)l, r:treep:eltHp:forall x : elt, Some p = Some x -> In x rIn p~1 (Node l false r)l, r:treex:elt(forall x0 : elt, None = Some x0 -> In x0 r) -> omap xI None = Some x -> In x (Node l false r)l, r:treep:eltHp:forall x : elt, Some p = Some x -> In x rSome p = Some pl, r:treex:elt(forall x0 : elt, None = Some x0 -> In x0 r) -> omap xI None = Some x -> In x (Node l false r)l, r:treex:elt(forall x0 : elt, None = Some x0 -> In x0 r) -> omap xI None = Some x -> In x (Node l false r)discriminate. Qed.l, r:treex:eltIHr:forall x0 : elt, None = Some x0 -> In x0 rH:omap xI None = Some xIn x (Node l false r)forall s : t, choose s = None -> Empty sforall s : t, choose s = None -> Empty sforall s : t, choose s = None -> forall a : elt, mem a s <> trues:tH:choose s = Noneforall a : elt, mem a s <> trueH:choose Leaf = Noneforall a : elt, mem a Leaf <> truel:treeo:boolr:treeH:choose (Node l o r) = NoneIHl:choose l = None -> forall a : elt, mem a l <> trueIHr:choose r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l o r) <> trueH:choose Leaf = Nonea:eltmem a Leaf <> truel:treeo:boolr:treeH:choose (Node l o r) = NoneIHl:choose l = None -> forall a : elt, mem a l <> trueIHr:choose r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l o r) <> truel:treeo:boolr:treeH:choose (Node l o r) = NoneIHl:choose l = None -> forall a : elt, mem a l <> trueIHr:choose r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l o r) <> truel, r:treeH:choose (Node l true r) = NoneIHl:choose l = None -> forall a : elt, mem a l <> trueIHr:choose r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l true r) <> truel, r:treeH:choose (Node l false r) = NoneIHl:choose l = None -> forall a : elt, mem a l <> trueIHr:choose r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l false r) <> truel, r:treeH:choose (Node l false r) = NoneIHl:choose l = None -> forall a : elt, mem a l <> trueIHr:choose r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l false r) <> truel, r:treeH:match choose l with | Some i => Some i~0 | None => omap xI (choose r) end = NoneIHl:choose l = None -> forall a : elt, mem a l <> trueIHr:choose r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l false r) <> truel, r:treee:eltH:Some e~0 = NoneIHl:Some e = None -> forall a : elt, mem a l <> trueIHr:choose r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l false r) <> truel, r:treeH:omap xI (choose r) = NoneIHl:None = None -> forall a : elt, mem a l <> trueIHr:choose r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l false r) <> truel, r:treeH:omap xI (choose r) = NoneIHl:None = None -> forall a : elt, mem a l <> trueIHr:choose r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l false r) <> truel, r:treee:eltH:omap xI (Some e) = NoneIHl:None = None -> forall a : elt, mem a l <> trueIHr:Some e = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l false r) <> truel, r:treeH:omap xI None = NoneIHl:None = None -> forall a : elt, mem a l <> trueIHr:None = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l false r) <> truel, r:treeH:omap xI None = NoneIHl:None = None -> forall a : elt, mem a l <> trueIHr:None = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l false r) <> truel, r:treeH:omap xI None = NoneIHl:None = None -> forall a0 : elt, mem a0 l <> trueIHr:None = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~1 (Node l false r) <> truel, r:treeH:omap xI None = NoneIHl:None = None -> forall a0 : elt, mem a0 l <> trueIHr:None = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l false r) <> truel, r:treeH:omap xI None = NoneIHl:None = None -> forall a : elt, mem a l <> trueIHr:None = None -> forall a : elt, mem a r <> truemem 1 (Node l false r) <> truel, r:treeH:omap xI None = NoneIHl:None = None -> forall a0 : elt, mem a0 l <> trueIHr:None = None -> forall a0 : elt, mem a0 r <> truea:positiveNone = Nonel, r:treeH:omap xI None = NoneIHl:None = None -> forall a0 : elt, mem a0 l <> trueIHr:None = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l false r) <> truel, r:treeH:omap xI None = NoneIHl:None = None -> forall a : elt, mem a l <> trueIHr:None = None -> forall a : elt, mem a r <> truemem 1 (Node l false r) <> truel, r:treeH:omap xI None = NoneIHl:None = None -> forall a0 : elt, mem a0 l <> trueIHr:None = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l false r) <> truel, r:treeH:omap xI None = NoneIHl:None = None -> forall a : elt, mem a l <> trueIHr:None = None -> forall a : elt, mem a r <> truemem 1 (Node l false r) <> truel, r:treeH:omap xI None = NoneIHl:None = None -> forall a0 : elt, mem a0 l <> trueIHr:None = None -> forall a0 : elt, mem a0 r <> truea:positiveNone = Nonel, r:treeH:omap xI None = NoneIHl:None = None -> forall a : elt, mem a l <> trueIHr:None = None -> forall a : elt, mem a r <> truemem 1 (Node l false r) <> truediscriminate. Qed.l, r:treeH:omap xI None = NoneIHl:None = None -> forall a : elt, mem a l <> trueIHr:None = None -> forall a : elt, mem a r <> truemem 1 (Node l false r) <> trueforall s : t, is_empty s = true -> choose s = Noneforall s : t, is_empty s = true -> choose s = Nones:tHs:is_empty s = truechoose s = Nones:tHs:is_empty s = trueforall e : elt, choose s = Some e -> Some e = Nones:tHs:is_empty s = truep:eltHp:choose s = Some pSome p = Nones:tHs:is_empty s = truep:eltHp:In p sSome p = Noneelim (Hs _ Hp). Qed.s:tHs:Empty sp:eltHp:In p sSome p = Noneforall s s' : t, s [=] s' -> choose s = choose s'forall s s' : t, s [=] s' -> choose s = choose s'forall s s' : t, equal s s' = true -> choose s = choose s'forall s' : t, equal Leaf s' = true -> choose Leaf = choose s'l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'forall s' : t, equal (Node l o r) s' = true -> choose (Node l o r) = choose s's':tH:equal Leaf s' = truechoose Leaf = choose s'l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'forall s' : t, equal (Node l o r) s' = true -> choose (Node l o r) = choose s's':tH:equal Leaf s' = truechoose s' = choose Leafl:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'forall s' : t, equal (Node l o r) s' = true -> choose (Node l o r) = choose s's':tH:equal Leaf s' = trueis_empty s' = truel:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'forall s' : t, equal (Node l o r) s' = true -> choose (Node l o r) = choose s'l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'forall s' : t, equal (Node l o r) s' = true -> choose (Node l o r) = choose s'l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'equal (Node l o r) Leaf = true -> choose (Node l o r) = choose Leafl:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l':treeo':boolr':treeequal (Node l o r) (Node l' o' r') = true -> choose (Node l o r) = choose (Node l' o' r')l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'forall s : tree, equal s Leaf = true -> choose s = choose Leafl:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l':treeo':boolr':treeequal (Node l o r) (Node l' o' r') = true -> choose (Node l o r) = choose (Node l' o' r')l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'forall s : tree, equal s Leaf = true -> choose s = Nonel:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l':treeo':boolr':treeequal (Node l o r) (Node l' o' r') = true -> choose (Node l o r) = choose (Node l' o' r')l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s's:treeH:equal s Leaf = truechoose s = Nonel:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l':treeo':boolr':treeequal (Node l o r) (Node l' o' r') = true -> choose (Node l o r) = choose (Node l' o' r')l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s's:treeH:equal s Leaf = trueis_empty s = truel:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l':treeo':boolr':treeequal (Node l o r) (Node l' o' r') = true -> choose (Node l o r) = choose (Node l' o' r')l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s's:treeH:s [=] Leafis_empty s = truel:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l':treeo':boolr':treeequal (Node l o r) (Node l' o' r') = true -> choose (Node l o r) = choose (Node l' o' r')l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s's:treeH:eq Leaf sis_empty s = truel:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l':treeo':boolr':treeequal (Node l o r) (Node l' o' r') = true -> choose (Node l o r) = choose (Node l' o' r')l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s's:treeH:equal Leaf s = trueis_empty s = truel:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l':treeo':boolr':treeequal (Node l o r) (Node l' o' r') = true -> choose (Node l o r) = choose (Node l' o' r')l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l':treeo':boolr':treeequal (Node l o r) (Node l' o' r') = true -> choose (Node l o r) = choose (Node l' o' r')l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l':treeo':boolr':treeeqb o o' &&& equal l l' &&& equal r r' = true -> (if o then Some 1 else match choose l with | Some i => Some i~0 | None => omap xI (choose r) end) = (if o' then Some 1 else match choose l' with | Some i => Some i~0 | None => omap xI (choose r') end)l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l':treeo':boolr':tree(o = o' /\ equal l l' = true) /\ equal r r' = true -> (if o then Some 1 else match choose l with | Some i => Some i~0 | None => omap xI (choose r) end) = (if o' then Some 1 else match choose l' with | Some i => Some i~0 | None => omap xI (choose r') end)l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l', r':treeHl:equal l l' = trueHr:equal r r' = true(if o then Some 1 else match choose l with | Some i => Some i~0 | None => omap xI (choose r) end) = (if o then Some 1 else match choose l' with | Some i => Some i~0 | None => omap xI (choose r') end)reflexivity. Qed.l:treeo:boolr:treeIHl:forall s' : t, equal l s' = true -> choose l = choose s'IHr:forall s' : t, equal r s' = true -> choose r = choose s'l', r':treeHl:equal l l' = trueHr:equal r r' = true(if o then Some 1 else match choose l' with | Some i => Some i~0 | None => omap xI (choose r') end) = (if o then Some 1 else match choose l' with | Some i => Some i~0 | None => omap xI (choose r') end)forall (s s' : t) (x y : elt), choose s = Some x -> choose s' = Some y -> s [=] s' -> E.eq x yforall (s s' : t) (x y : elt), choose s = Some x -> choose s' = Some y -> s [=] s' -> E.eq x ys, s':tx, y:eltHx:choose s = Some xHy:choose s' = Some yH:s [=] s'E.eq x ycongruence. Qed.s, s':tx, y:eltHx:choose s = Some xHy:choose s' = Some yH:choose s = choose s'E.eq x y
Specification of min_elt
forall (s : t) (x : elt), min_elt s = Some x -> In x sforall (s : t) (x : elt), min_elt s = Some x -> In x sforall (s : t) (x : elt), min_elt s = Some x -> mem x s = trueforall x : elt, None = Some x -> false = truel:treeo:boolr:treeIHl:forall x : elt, min_elt l = Some x -> mem x l = trueIHr:forall x : elt, min_elt r = Some x -> mem x r = trueforall x : elt, match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some x -> match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truex:eltH:None = Some xfalse = truel:treeo:boolr:treeIHl:forall x : elt, min_elt l = Some x -> mem x l = trueIHr:forall x : elt, min_elt r = Some x -> mem x r = trueforall x : elt, match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some x -> match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treeIHl:forall x : elt, min_elt l = Some x -> mem x l = trueIHr:forall x : elt, min_elt r = Some x -> mem x r = trueforall x : elt, match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some x -> match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treeIHl:forall x0 : elt, min_elt l = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, min_elt r = Some x0 -> mem x0 r = truex:eltmatch min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some x -> match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treee:eltIHl:forall x0 : elt, Some e = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, min_elt r = Some x0 -> mem x0 r = truex:eltH:Some e~0 = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, min_elt r = Some x0 -> mem x0 r = truex:eltH:(if o then Some 1 else omap xI (min_elt r)) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treee:eltIHl:forall x : elt, Some e = Some x -> mem x l = trueIHr:forall x : elt, min_elt r = Some x -> mem x r = truemem e l = truel:treeo:boolr:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, min_elt r = Some x0 -> mem x0 r = truex:eltH:(if o then Some 1 else omap xI (min_elt r)) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treee:eltIHl:forall x : elt, Some e = Some x -> mem x l = trueIHr:forall x : elt, min_elt r = Some x -> mem x r = trueSome e = Some el:treeo:boolr:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, min_elt r = Some x0 -> mem x0 r = truex:eltH:(if o then Some 1 else omap xI (min_elt r)) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, min_elt r = Some x0 -> mem x0 r = truex:eltH:(if o then Some 1 else omap xI (min_elt r)) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, min_elt r = Some x0 -> mem x0 r = truex:eltH:Some 1 = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = truel, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, min_elt r = Some x0 -> mem x0 r = truex:eltH:omap xI (min_elt r) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treeIHl:forall x : elt, None = Some x -> mem x l = trueIHr:forall x : elt, min_elt r = Some x -> mem x r = truetrue = truel, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, min_elt r = Some x0 -> mem x0 r = truex:eltH:omap xI (min_elt r) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, min_elt r = Some x0 -> mem x0 r = truex:eltH:omap xI (min_elt r) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = truee:eltIHr:forall x0 : elt, Some e = Some x0 -> mem x0 r = truex:eltH:Some e~1 = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:None = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treeIHl:forall x : elt, None = Some x -> mem x l = truee:eltIHr:forall x : elt, Some e = Some x -> mem x r = truemem e r = truel, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:None = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treeIHl:forall x : elt, None = Some x -> mem x l = truee:eltIHr:forall x : elt, Some e = Some x -> mem x r = trueSome e = Some el, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:None = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truediscriminate. Qed.l, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:None = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = trueforall s : t, min_elt s = None -> Empty sforall s : t, min_elt s = None -> Empty sforall s : t, min_elt s = None -> forall a : elt, mem a s <> trues:tH:min_elt s = Noneforall a : elt, mem a s <> trueH:min_elt Leaf = Noneforall a : elt, mem a Leaf <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l o r) <> trueH:min_elt Leaf = Nonea:eltmem a Leaf <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~1 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemin_elt r = Nonel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemin_elt (Node l o r) = None -> min_elt r = Nonel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treemin_elt (Node l o r) = None -> min_elt r = Nonel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treematch min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = None -> min_elt r = Nonel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treee:eltmatch min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (Some e) end = None -> Some e = Nonel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treee:eltH:(if o then Some 1 else omap xI (Some e)) = NoneSome e = Nonel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemin_elt l = Nonel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeIHl:min_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:min_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemin_elt (Node l o r) = None -> min_elt l = Nonel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treemin_elt (Node l o r) = None -> min_elt l = Nonel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treematch min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = None -> min_elt l = Nonel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treee:eltSome e~0 = None -> Some e = Nonel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeH:min_elt (Node l o r) = NoneIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeIHl:min_elt l = None -> forall a : elt, mem a l <> trueIHr:min_elt r = None -> forall a : elt, mem a r <> truemin_elt (Node l o r) = None -> mem 1 (Node l o r) <> truel:treeo:boolr:treemin_elt (Node l o r) = None -> mem 1 (Node l o r) <> truel:treeo:boolr:treematch min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = None -> o <> truedestruct o; discriminate. Qed.l:treeo:boolr:treeH:(if o then Some 1 else omap xI (min_elt r)) = Noneo <> trueforall (s : t) (x y : elt), min_elt s = Some x -> In y s -> ~ E.lt y xforall (s : t) (x y : elt), min_elt s = Some x -> In y s -> ~ E.lt y xforall (s : t) (x y : elt), min_elt s = Some x -> mem y s = true -> ~ E.lt y xx, y:eltH:min_elt Leaf = Some xH':mem y Leaf = true~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:min_elt (Node l o r) = Some xH':mem y (Node l o r) = true~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:min_elt (Node l o r) = Some xH':mem y (Node l o r) = true~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some xH':mem y (Node l o r) = true~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some xH':mem y (Node l o r) = trueforall e : elt, min_elt l = Some e -> ~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some xH':mem y (Node l o r) = truemin_elt l = None -> ~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some xH':mem y (Node l o r) = truep:eltHp:min_elt l = Some p~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some xH':mem y (Node l o r) = truemin_elt l = None -> ~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y, p:eltH:Some p~0 = Some xH':mem y (Node l o r) = trueHp:min_elt l = Some p~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some xH':mem y (Node l o r) = truemin_elt l = None -> ~ E.lt y xl:treeo:boolr:treeIHl:forall x y0 : elt, min_elt l = Some x -> mem y0 l = true -> ~ E.lt y0 xIHr:forall x y0 : elt, min_elt r = Some x -> mem y0 r = true -> ~ E.lt y0 xy, p:eltH':mem y (Node l o r) = trueHp:min_elt l = Some p~ E.lt y p~0l:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some xH':mem y (Node l o r) = truemin_elt l = None -> ~ E.lt y xl:treeo:boolr:treeIHl:forall x y : elt, min_elt l = Some x -> mem y l = true -> ~ E.lt y xIHr:forall x y : elt, min_elt r = Some x -> mem y r = true -> ~ E.lt y xz:positivep:eltH':mem z~0 (Node l o r) = trueHp:min_elt l = Some pH:E.lt z pFalsel:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some xH':mem y (Node l o r) = truemin_elt l = None -> ~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:match min_elt l with | Some i => Some i~0 | None => if o then Some 1 else omap xI (min_elt r) end = Some xH':mem y (Node l o r) = truemin_elt l = None -> ~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:(if o then Some 1 else omap xI (min_elt r)) = Some xH':mem y (Node l o r) = trueHp:min_elt l = None~ E.lt y xl:treeo:boolr:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:(if o then Some 1 else omap xI (min_elt r)) = Some xH':mem y (Node l o r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:Some 1 = Some xH':mem y (Node l true r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI (min_elt r) = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x y0 : elt, min_elt l = Some x -> mem y0 l = true -> ~ E.lt y0 xIHr:forall x y0 : elt, min_elt r = Some x -> mem y0 r = true -> ~ E.lt y0 xy:eltH':mem y (Node l true r) = trueHp:Empty l~ E.lt y 1l, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI (min_elt r) = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x y0 : elt, min_elt l = Some x -> mem y0 l = true -> ~ E.lt y0 xIHr:forall x y0 : elt, min_elt r = Some x -> mem y0 r = true -> ~ E.lt y0 xy:eltH':mem y (Node l true r) = trueHp:Empty lHl:E.lt y 1Falsel, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI (min_elt r) = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x y : elt, min_elt l = Some x -> mem y l = true -> ~ E.lt y xIHr:forall x y : elt, min_elt r = Some x -> mem y r = true -> ~ E.lt y xz:positiveH':mem z~0 (Node l true r) = trueHp:Empty lHl:E.lt z~0 1Falsel, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI (min_elt r) = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, min_elt r = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI (min_elt r) = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0e:eltIHr:forall x0 y0 : elt, Some e = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI (Some e) = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, None = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI None = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x y0 : elt, min_elt l = Some x -> mem y0 l = true -> ~ E.lt y0 xe:eltIHr:forall x y0 : elt, Some e = Some x -> mem y0 r = true -> ~ E.lt y0 xy:eltH':mem y (Node l false r) = trueHp:Empty l~ E.lt y e~1l, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, None = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI None = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x y : elt, min_elt l = Some x -> mem y l = true -> ~ E.lt y xe:eltIHr:forall x y : elt, Some e = Some x -> mem y r = true -> ~ E.lt y xz:positiveH':mem z~1 (Node l false r) = trueHp:Empty l~ E.lt z~1 e~1l, r:treeIHl:forall x y : elt, min_elt l = Some x -> mem y l = true -> ~ E.lt y xe:eltIHr:forall x y : elt, Some e = Some x -> mem y r = true -> ~ E.lt y xz:positiveH':mem z~0 (Node l false r) = trueHp:Empty l~ E.lt z~0 e~1l, r:treeIHl:forall x y : elt, min_elt l = Some x -> mem y l = true -> ~ E.lt y xe:eltIHr:forall x y : elt, Some e = Some x -> mem y r = true -> ~ E.lt y xH':mem 1 (Node l false r) = trueHp:Empty l~ E.lt 1 e~1l, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, None = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI None = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x y : elt, min_elt l = Some x -> mem y l = true -> ~ E.lt y xe:eltIHr:forall x y : elt, Some e = Some x -> mem y r = true -> ~ E.lt y xz:positiveH':mem z~0 (Node l false r) = trueHp:Empty l~ E.lt z~0 e~1l, r:treeIHl:forall x y : elt, min_elt l = Some x -> mem y l = true -> ~ E.lt y xe:eltIHr:forall x y : elt, Some e = Some x -> mem y r = true -> ~ E.lt y xH':mem 1 (Node l false r) = trueHp:Empty l~ E.lt 1 e~1l, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, None = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI None = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y xl, r:treeIHl:forall x y : elt, min_elt l = Some x -> mem y l = true -> ~ E.lt y xe:eltIHr:forall x y : elt, Some e = Some x -> mem y r = true -> ~ E.lt y xH':mem 1 (Node l false r) = trueHp:Empty l~ E.lt 1 e~1l, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, None = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI None = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y xdiscriminate. Qed.l, r:treeIHl:forall x0 y0 : elt, min_elt l = Some x0 -> mem y0 l = true -> ~ E.lt y0 x0IHr:forall x0 y0 : elt, None = Some x0 -> mem y0 r = true -> ~ E.lt y0 x0x, y:eltH:omap xI None = Some xH':mem y (Node l false r) = trueHp:Empty l~ E.lt y x
Specification of max_elt
forall (s : t) (x : elt), max_elt s = Some x -> In x sforall (s : t) (x : elt), max_elt s = Some x -> In x sforall (s : t) (x : elt), max_elt s = Some x -> mem x s = trueforall x : elt, None = Some x -> false = truel:treeo:boolr:treeIHl:forall x : elt, max_elt l = Some x -> mem x l = trueIHr:forall x : elt, max_elt r = Some x -> mem x r = trueforall x : elt, match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some x -> match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truex:eltH:None = Some xfalse = truel:treeo:boolr:treeIHl:forall x : elt, max_elt l = Some x -> mem x l = trueIHr:forall x : elt, max_elt r = Some x -> mem x r = trueforall x : elt, match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some x -> match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treeIHl:forall x : elt, max_elt l = Some x -> mem x l = trueIHr:forall x : elt, max_elt r = Some x -> mem x r = trueforall x : elt, match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some x -> match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treeIHl:forall x0 : elt, max_elt l = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, max_elt r = Some x0 -> mem x0 r = truex:eltmatch max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some x -> match x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treeIHl:forall x0 : elt, max_elt l = Some x0 -> mem x0 l = truee:eltIHr:forall x0 : elt, Some e = Some x0 -> mem x0 r = truex:eltH:Some e~1 = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treeIHl:forall x0 : elt, max_elt l = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:(if o then Some 1 else omap xO (max_elt l)) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treeIHl:forall x : elt, max_elt l = Some x -> mem x l = truee:eltIHr:forall x : elt, Some e = Some x -> mem x r = truemem e r = truel:treeo:boolr:treeIHl:forall x0 : elt, max_elt l = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:(if o then Some 1 else omap xO (max_elt l)) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treeIHl:forall x : elt, max_elt l = Some x -> mem x l = truee:eltIHr:forall x : elt, Some e = Some x -> mem x r = trueSome e = Some el:treeo:boolr:treeIHl:forall x0 : elt, max_elt l = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:(if o then Some 1 else omap xO (max_elt l)) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel:treeo:boolr:treeIHl:forall x0 : elt, max_elt l = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:(if o then Some 1 else omap xO (max_elt l)) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => o end = truel, r:treeIHl:forall x0 : elt, max_elt l = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:Some 1 = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => true end = truel, r:treeIHl:forall x0 : elt, max_elt l = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:omap xO (max_elt l) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treeIHl:forall x : elt, max_elt l = Some x -> mem x l = trueIHr:forall x : elt, None = Some x -> mem x r = truetrue = truel, r:treeIHl:forall x0 : elt, max_elt l = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:omap xO (max_elt l) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treeIHl:forall x0 : elt, max_elt l = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:omap xO (max_elt l) = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treee:eltIHl:forall x0 : elt, Some e = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:Some e~0 = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:None = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treee:eltIHl:forall x : elt, Some e = Some x -> mem x l = trueIHr:forall x : elt, None = Some x -> mem x r = truemem e l = truel, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:None = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truel, r:treee:eltIHl:forall x : elt, Some e = Some x -> mem x l = trueIHr:forall x : elt, None = Some x -> mem x r = trueSome e = Some el, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:None = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = truediscriminate. Qed.l, r:treeIHl:forall x0 : elt, None = Some x0 -> mem x0 l = trueIHr:forall x0 : elt, None = Some x0 -> mem x0 r = truex:eltH:None = Some xmatch x with | i~1 => mem i r | i~0 => mem i l | 1 => false end = trueforall s : t, max_elt s = None -> Empty sforall s : t, max_elt s = None -> Empty sforall s : t, max_elt s = None -> forall a : elt, mem a s <> trues:tH:max_elt s = Noneforall a : elt, mem a s <> trueH:max_elt Leaf = Noneforall a : elt, mem a Leaf <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l o r) <> trueH:max_elt Leaf = Nonea:eltmem a Leaf <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> trueforall a : elt, mem a (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~1 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemax_elt r = Nonel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemax_elt (Node l o r) = None -> max_elt r = Nonel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treemax_elt (Node l o r) = None -> max_elt r = Nonel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treematch max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = None -> max_elt r = Nonel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treee:eltSome e~1 = None -> Some e = Nonel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemem a~0 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemax_elt l = Nonel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeIHl:max_elt l = None -> forall a0 : elt, mem a0 l <> trueIHr:max_elt r = None -> forall a0 : elt, mem a0 r <> truea:positivemax_elt (Node l o r) = None -> max_elt l = Nonel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treemax_elt (Node l o r) = None -> max_elt l = Nonel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treematch max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = None -> max_elt l = Nonel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treee:eltmatch max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (Some e) end = None -> Some e = Nonel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treee:eltH:(if o then Some 1 else omap xO (Some e)) = NoneSome e = Nonel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeH:max_elt (Node l o r) = NoneIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemem 1 (Node l o r) <> truel:treeo:boolr:treeIHl:max_elt l = None -> forall a : elt, mem a l <> trueIHr:max_elt r = None -> forall a : elt, mem a r <> truemax_elt (Node l o r) = None -> mem 1 (Node l o r) <> truel:treeo:boolr:treemax_elt (Node l o r) = None -> mem 1 (Node l o r) <> truel:treeo:boolr:treematch max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = None -> o <> truedestruct o; discriminate. Qed.l:treeo:boolr:treeH:(if o then Some 1 else omap xO (max_elt l)) = Noneo <> trueforall (s : t) (x y : elt), max_elt s = Some x -> In y s -> ~ E.lt x yforall (s : t) (x y : elt), max_elt s = Some x -> In y s -> ~ E.lt x yforall (s : t) (x y : elt), max_elt s = Some x -> mem y s = true -> ~ E.lt x yx, y:eltH:max_elt Leaf = Some xH':mem y Leaf = true~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:max_elt (Node l o r) = Some xH':mem y (Node l o r) = true~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:max_elt (Node l o r) = Some xH':mem y (Node l o r) = true~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some xH':mem y (Node l o r) = true~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some xH':mem y (Node l o r) = trueforall e : elt, max_elt r = Some e -> ~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some xH':mem y (Node l o r) = truemax_elt r = None -> ~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some xH':mem y (Node l o r) = truep:eltHp:max_elt r = Some p~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some xH':mem y (Node l o r) = truemax_elt r = None -> ~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y, p:eltH:Some p~1 = Some xH':mem y (Node l o r) = trueHp:max_elt r = Some p~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some xH':mem y (Node l o r) = truemax_elt r = None -> ~ E.lt x yl:treeo:boolr:treeIHl:forall x y0 : elt, max_elt l = Some x -> mem y0 l = true -> ~ E.lt x y0IHr:forall x y0 : elt, max_elt r = Some x -> mem y0 r = true -> ~ E.lt x y0y, p:eltH':mem y (Node l o r) = trueHp:max_elt r = Some p~ E.lt p~1 yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some xH':mem y (Node l o r) = truemax_elt r = None -> ~ E.lt x yl:treeo:boolr:treeIHl:forall x y : elt, max_elt l = Some x -> mem y l = true -> ~ E.lt x yIHr:forall x y : elt, max_elt r = Some x -> mem y r = true -> ~ E.lt x yz:positivep:eltH':mem z~1 (Node l o r) = trueHp:max_elt r = Some pH:E.lt p zFalsel:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some xH':mem y (Node l o r) = truemax_elt r = None -> ~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:match max_elt r with | Some i => Some i~1 | None => if o then Some 1 else omap xO (max_elt l) end = Some xH':mem y (Node l o r) = truemax_elt r = None -> ~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:(if o then Some 1 else omap xO (max_elt l)) = Some xH':mem y (Node l o r) = trueHp:max_elt r = None~ E.lt x yl:treeo:boolr:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:(if o then Some 1 else omap xO (max_elt l)) = Some xH':mem y (Node l o r) = trueHp:Empty r~ E.lt x yl, r:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:Some 1 = Some xH':mem y (Node l true r) = trueHp:Empty r~ E.lt x yl, r:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO (max_elt l) = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x yl, r:treeIHl:forall x y0 : elt, max_elt l = Some x -> mem y0 l = true -> ~ E.lt x y0IHr:forall x y0 : elt, max_elt r = Some x -> mem y0 r = true -> ~ E.lt x y0y:eltH':mem y (Node l true r) = trueHp:Empty r~ E.lt 1 yl, r:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO (max_elt l) = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x yl, r:treeIHl:forall x y0 : elt, max_elt l = Some x -> mem y0 l = true -> ~ E.lt x y0IHr:forall x y0 : elt, max_elt r = Some x -> mem y0 r = true -> ~ E.lt x y0y:eltH':mem y (Node l true r) = trueHp:Empty rHl:E.lt 1 yFalsel, r:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO (max_elt l) = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x yl, r:treeIHl:forall x y : elt, max_elt l = Some x -> mem y l = true -> ~ E.lt x yIHr:forall x y : elt, max_elt r = Some x -> mem y r = true -> ~ E.lt x yz:positiveH':mem z~1 (Node l true r) = trueHp:Empty rHl:E.lt 1 z~1Falsel, r:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO (max_elt l) = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x yl, r:treeIHl:forall x0 y0 : elt, max_elt l = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO (max_elt l) = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x yl, r:treee:eltIHl:forall x0 y0 : elt, Some e = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO (Some e) = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x yl, r:treeIHl:forall x0 y0 : elt, None = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO None = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x yl, r:treee:eltIHl:forall x y0 : elt, Some e = Some x -> mem y0 l = true -> ~ E.lt x y0IHr:forall x y0 : elt, max_elt r = Some x -> mem y0 r = true -> ~ E.lt x y0y:eltH':mem y (Node l false r) = trueHp:Empty r~ E.lt e~0 yl, r:treeIHl:forall x0 y0 : elt, None = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO None = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x yl, r:treee:eltIHl:forall x y : elt, Some e = Some x -> mem y l = true -> ~ E.lt x yIHr:forall x y : elt, max_elt r = Some x -> mem y r = true -> ~ E.lt x yz:positiveH':mem z~1 (Node l false r) = trueHp:Empty r~ E.lt e~0 z~1l, r:treee:eltIHl:forall x y : elt, Some e = Some x -> mem y l = true -> ~ E.lt x yIHr:forall x y : elt, max_elt r = Some x -> mem y r = true -> ~ E.lt x yz:positiveH':mem z~0 (Node l false r) = trueHp:Empty r~ E.lt e~0 z~0l, r:treee:eltIHl:forall x y : elt, Some e = Some x -> mem y l = true -> ~ E.lt x yIHr:forall x y : elt, max_elt r = Some x -> mem y r = true -> ~ E.lt x yH':mem 1 (Node l false r) = trueHp:Empty r~ E.lt e~0 1l, r:treeIHl:forall x0 y0 : elt, None = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO None = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x yl, r:treee:eltIHl:forall x y : elt, Some e = Some x -> mem y l = true -> ~ E.lt x yIHr:forall x y : elt, max_elt r = Some x -> mem y r = true -> ~ E.lt x yz:positiveH':mem z~0 (Node l false r) = trueHp:Empty r~ E.lt e~0 z~0l, r:treee:eltIHl:forall x y : elt, Some e = Some x -> mem y l = true -> ~ E.lt x yIHr:forall x y : elt, max_elt r = Some x -> mem y r = true -> ~ E.lt x yH':mem 1 (Node l false r) = trueHp:Empty r~ E.lt e~0 1l, r:treeIHl:forall x0 y0 : elt, None = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO None = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x yl, r:treee:eltIHl:forall x y : elt, Some e = Some x -> mem y l = true -> ~ E.lt x yIHr:forall x y : elt, max_elt r = Some x -> mem y r = true -> ~ E.lt x yH':mem 1 (Node l false r) = trueHp:Empty r~ E.lt e~0 1l, r:treeIHl:forall x0 y0 : elt, None = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO None = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x ydiscriminate. Qed. End PositiveSet.l, r:treeIHl:forall x0 y0 : elt, None = Some x0 -> mem y0 l = true -> ~ E.lt x0 y0IHr:forall x0 y0 : elt, max_elt r = Some x0 -> mem y0 r = true -> ~ E.lt x0 y0x, y:eltH:omap xO None = Some xH':mem y (Node l false r) = trueHp:Empty r~ E.lt x y