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) *)
(************************************************************************)
This file proposes an implementation of the non-dependent interface
FMapInterface.S using lists of pairs ordered (increasing) with respect to
left projection.
Require Import FunInd FMapInterface. Set Implicit Arguments. Unset Strict Implicit. Module Raw (X:OrderedType). Module Import MX := OrderedTypeFacts X. Module Import PX := KeyOrderedType X. Definition key := X.t. Definition t (elt:Type) := list (X.t * elt). Section Elt. Variable elt : Type. Notation eqk := (eqk (elt:=elt)). Notation eqke := (eqke (elt:=elt)). Notation ltk := (ltk (elt:=elt)). Notation MapsTo := (MapsTo (elt:=elt)). Notation In := (In (elt:=elt)). Notation Sort := (sort ltk). Notation Inf := (lelistA (ltk)).
Definition empty : t elt := nil. Definition Empty m := forall (a : key)(e:elt) , ~ MapsTo a e m.elt:TypeEmpty emptyelt:TypeEmpty emptyelt:Typeforall (a : key) (e : elt), ~ MapsTo a e nilelt:Typea:keye:elt~ MapsTo a e nilinversion abs. Qed. Hint Resolve empty_1 : core.elt:Typea:keye:eltabs:MapsTo a e nilFalseelt:TypeSort emptyunfold empty; auto. Qed.elt:TypeSort empty
Definition is_empty (l : t elt) : bool := if l then true else false.elt:Typeforall m : list (X.t * elt), Empty m -> is_empty m = trueelt:Typeforall m : list (X.t * elt), Empty m -> is_empty m = trueelt:Typeforall m : list (X.t * elt), (forall (a : key) (e : elt), ~ InA eqke (a, e) m) -> is_empty m = trueelt:Typem:list (X.t * elt)(forall (a : key) (e : elt), ~ InA eqke (a, e) m) -> is_empty m = trueelt:Typem:list (X.t * elt)forall (p : X.t * elt) (l : list (X.t * elt)), (forall (a : key) (e : elt), ~ InA eqke (a, e) (p :: l)) -> is_empty (p :: l) = trueabsurd (InA eqke (k, e) ((k, e) :: l));auto. Qed.elt:Typem:list (X.t * elt)k:X.te:eltl:list (X.t * elt)inlist:forall (a : key) (e0 : elt), ~ InA eqke (a, e0) ((k, e) :: l)is_empty ((k, e) :: l) = trueelt:Typeforall m : t elt, is_empty m = true -> Empty melt:Typeforall m : t elt, is_empty m = true -> Empty melt:Typem:t eltis_empty m = true -> Empty melt:Typem:t eltforall (p : X.t * elt) (l : list (X.t * elt)), is_empty (p :: l) = true -> Empty (p :: l)inversion abs. Qed.elt:Typem:t eltp:(X.t * elt)%typel:list (X.t * elt)abs:is_empty (p :: l) = trueEmpty (p :: l)
Function mem (k : key) (s : t elt) {struct s} : bool := match s with | nil => false | (k',_) :: l => match X.compare k k' with | LT _ => false | EQ _ => true | GT _ => mem k l end end.elt:Typeforall m : list (X.t * elt), Sort m -> forall x : X.t, In x m -> mem x m = trueelt:Typeforall m : list (X.t * elt), Sort m -> forall x : X.t, In x m -> mem x m = trueelt:Typem:list (X.t * elt)x:X.tSort m -> In x m -> mem x m = trueelt:Typex:X.tsorted:Sort nilbelong1:In x nilfalse = trueelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt x k'e0:X.compare x k' = LT _x0sorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)false = trueelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> In x l -> mem x l = truesorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)mem x l = trueelt:Typex:X.tsorted:Sort nilbelong1:In x nilx0:eltH:MapsTo x x0 nilfalse = trueelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt x k'e0:X.compare x k' = LT _x0sorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)false = trueelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> In x l -> mem x l = truesorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)mem x l = trueelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt x k'e0:X.compare x k' = LT _x0sorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)false = trueelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> In x l -> mem x l = truesorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)mem x l = trueelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt x k'e0:X.compare x k' = LT _x0sorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)~ In x ((k', _x) :: l)elt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> In x l -> mem x l = truesorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)mem x l = trueelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> In x l -> mem x l = truesorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)mem x l = trueelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> In x l -> mem x l = truesorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)Sort lelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> In x l -> mem x l = truesorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)In x lelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> In x l -> mem x l = truesorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)In x lelt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> In x l -> mem x l = truesorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)X.eq x k' -> In x labsurd (X.eq x k');auto. Qed.elt:Typex, k':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> In x l -> mem x l = truesorted:Sort ((k', _x) :: l)belong1:In x ((k', _x) :: l)abs:X.eq x k'In x lelt:Typeforall m : list (X.t * elt), Sort m -> forall x : key, mem x m = true -> In x melt:Typeforall m : list (X.t * elt), Sort m -> forall x : key, mem x m = true -> In x melt:Typem:list (X.t * elt)x:keySort m -> mem x m = true -> exists e : elt, InA eqke (x, e) melt:Typex:keyk':X.t_x:eltl:list (X.t * elt)_x0:X.eq x k'e0:X.compare x k' = EQ _x0sorted:Sort ((k', _x) :: l)hyp:true = trueexists e : elt, InA eqke (x, e) ((k', _x) :: l)elt:Typex:keyk':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> mem x l = true -> exists e : elt, InA eqke (x, e) lsorted:Sort ((k', _x) :: l)hyp:mem x l = trueexists e : elt, InA eqke (x, e) ((k', _x) :: l)elt:Typex:keyk':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0IHb:Sort l -> mem x l = true -> exists e : elt, InA eqke (x, e) lsorted:Sort ((k', _x) :: l)hyp:mem x l = trueexists e : elt, InA eqke (x, e) ((k', _x) :: l)elt:Typex:keyk':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0sorted:Sort ((k', _x) :: l)hyp:mem x l = truex0:eltH:InA eqke (x, x0) lexists e : elt, InA eqke (x, e) ((k', _x) :: l)elt:Typex:keyk':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0sorted:Sort ((k', _x) :: l)hyp:mem x l = trueSort linversion_clear sorted; auto. Qed.elt:Typex:keyk':X.t_x:eltl:list (X.t * elt)_x0:X.lt k' xe0:X.compare x k' = GT _x0sorted:Sort ((k', _x) :: l)hyp:mem x l = trueSort l
Function find (k:key) (s: t elt) {struct s} : option elt := match s with | nil => None | (k',x)::s' => match X.compare k k' with | LT _ => None | EQ _ => Some x | GT _ => find k s' end end.elt:Typeforall (m : t elt) (x : key) (e : elt), find x m = Some e -> MapsTo x e melt:Typeforall (m : t elt) (x : key) (e : elt), find x m = Some e -> MapsTo x e melt:Typem:t eltx:keyforall e : elt, find x m = Some e -> MapsTo x e mfunctional induction (find x m);simpl;intros e' eqfind; inversion eqfind; auto. Qed.elt:Typem:t eltx:keyforall e : elt, find x m = Some e -> InA eqke (x, e) melt:Typeforall m : list (X.t * elt), Sort m -> forall (x : X.t) (e : elt), MapsTo x e m -> find x m = Some eelt:Typeforall m : list (X.t * elt), Sort m -> forall (x : X.t) (e : elt), MapsTo x e m -> find x m = Some eelt:Typem:list (X.t * elt)x:X.te:eltSort m -> InA eqke (x, e) m -> find x m = Some eelt:Typex:X.te:eltSort nil -> InA eqke (x, e) nil -> None = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt x k'e1:X.compare x k' = LT _xSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> None = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.eq x k'e1:X.compare x k' = EQ _xSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> Some x0 = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt k' xe1:X.compare x k' = GT _xIHo:Sort s' -> InA eqke (x, e) s' -> find x s' = Some eSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> find x s' = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt x k'e1:X.compare x k' = LT _xSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> None = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.eq x k'e1:X.compare x k' = EQ _xSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> Some x0 = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt k' xe1:X.compare x k' = GT _xIHo:Sort s' -> InA eqke (x, e) s' -> find x s' = Some eSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> find x s' = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt x k'e1:X.compare x k' = LT _xHm:Sort ((k', x0) :: s')H0:eqke (x, e) (k', x0)None = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt x k'e1:X.compare x k' = LT _xHm:Sort ((k', x0) :: s')H0:InA eqke (x, e) s'None = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.eq x k'e1:X.compare x k' = EQ _xSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> Some x0 = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt k' xe1:X.compare x k' = GT _xIHo:Sort s' -> InA eqke (x, e) s' -> find x s' = Some eSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> find x s' = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt x k'e1:X.compare x k' = LT _xHm:Sort ((k', x0) :: s')H0:InA eqke (x, e) s'None = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.eq x k'e1:X.compare x k' = EQ _xSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> Some x0 = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt k' xe1:X.compare x k' = GT _xIHo:Sort s' -> InA eqke (x, e) s' -> find x s' = Some eSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> find x s' = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.eq x k'e1:X.compare x k' = EQ _xSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> Some x0 = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt k' xe1:X.compare x k' = GT _xIHo:Sort s' -> InA eqke (x, e) s' -> find x s' = Some eSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> find x s' = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.eq x k'Hm:Sort ((k', x0) :: s')H0:eqke (x, e) (k', x0)Some x0 = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.eq x k'Hm:Sort ((k', x0) :: s')H0:InA eqke (x, e) s'Some x0 = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt k' xe1:X.compare x k' = GT _xIHo:Sort s' -> InA eqke (x, e) s' -> find x s' = Some eSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> find x s' = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.eq x k'Hm:Sort ((k', x0) :: s')H0:InA eqke (x, e) s'Some x0 = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt k' xe1:X.compare x k' = GT _xIHo:Sort s' -> InA eqke (x, e) s' -> find x s' = Some eSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> find x s' = Some eelt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt k' xe1:X.compare x k' = GT _xIHo:Sort s' -> InA eqke (x, e) s' -> find x s' = Some eSort ((k', x0) :: s') -> InA eqke (x, e) ((k', x0) :: s') -> find x s' = Some ecompute in H2; destruct H2; order. Qed.elt:Typex:X.te:eltk':X.tx0:elts':list (X.t * elt)_x:X.lt k' xIHo:Sort s' -> InA eqke (x, e) s' -> find x s' = Some eH:Sort s'H0:Inf (k', x0) s'H2:eqke (x, e) (k', x0)find x s' = Some e
Function add (k : key) (x : elt) (s : t elt) {struct s} : t elt := match s with | nil => (k,x) :: nil | (k',y) :: l => match X.compare k k' with | LT _ => (k,x)::s | EQ _ => (k,x)::l | GT _ => (k',y) :: add k x l end end.elt:Typeforall (m : t elt) (x y : X.t) (e : elt), X.eq x y -> MapsTo y e (add x e m)elt:Typeforall (m : t elt) (x y : X.t) (e : elt), X.eq x y -> MapsTo y e (add x e m)elt:Typem:t eltx:X.te:eltforall y : X.t, X.eq x y -> MapsTo y e (add x e m)functional induction (add x e m);simpl;auto. Qed.elt:Typem:t eltx:X.te:eltforall y : X.t, X.eq x y -> InA eqke (y, e) (add x e m)elt:Typeforall (m : list (X.t * elt)) (x y : X.t) (e e' : elt), ~ X.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m)elt:Typeforall (m : list (X.t * elt)) (x y : X.t) (e e' : elt), ~ X.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m)elt:Typem:list (X.t * elt)x, y:X.te, e':elt~ X.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m)elt:Typem:list (X.t * elt)x:X.te':eltforall (y : X.t) (e : elt), ~ X.eq x y -> InA eqke (y, e) m -> InA eqke (y, e) (add x e' m)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.eq x k'forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) ((k', y) :: l) -> InA eqke (y0, e) ((x, e') :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt k' xIHt0:forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) l -> InA eqke (y0, e) (add x e' l)forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) ((k', y) :: l) -> InA eqke (y0, e) ((k', y) :: add x e' l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.eq x k'forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) ((k', y) :: l) -> InA eqke (y0, e) ((x, e') :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt k' xIHt0:forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) l -> InA eqke (y0, e) (add x e' l)forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) ((k', y) :: l) -> InA eqke (y0, e) ((k', y) :: add x e' l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.eq x k'y':X.te'':elteqky':~ X.eq x y'H:X.eq y' k'H0:e'' = yInA eqke (y', e'') ((x, e') :: l)elt:Typex:X.te':eltk':X.ty:elt_x:X.eq x k'y':X.te'':elteqky':~ X.eq x y'y0:(X.t * elt)%typel:list (X.t * elt)H:eqke (y', e'') y0InA eqke (y', e'') ((x, e') :: y0 :: l)elt:Typex:X.te':eltk':X.ty:elt_x:X.eq x k'y':X.te'':elteqky':~ X.eq x y'y0:(X.t * elt)%typel:list (X.t * elt)H0:InA eqke (y', e'') lInA eqke (y', e'') ((x, e') :: y0 :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt k' xIHt0:forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) l -> InA eqke (y0, e) (add x e' l)forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) ((k', y) :: l) -> InA eqke (y0, e) ((k', y) :: add x e' l)elt:Typex:X.te':eltk':X.ty:elt_x:X.eq x k'y':X.te'':elteqky':~ X.eq x y'y0:(X.t * elt)%typel:list (X.t * elt)H:eqke (y', e'') y0InA eqke (y', e'') ((x, e') :: y0 :: l)elt:Typex:X.te':eltk':X.ty:elt_x:X.eq x k'y':X.te'':elteqky':~ X.eq x y'y0:(X.t * elt)%typel:list (X.t * elt)H0:InA eqke (y', e'') lInA eqke (y', e'') ((x, e') :: y0 :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt k' xIHt0:forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) l -> InA eqke (y0, e) (add x e' l)forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) ((k', y) :: l) -> InA eqke (y0, e) ((k', y) :: add x e' l)elt:Typex:X.te':eltk':X.ty:elt_x:X.eq x k'y':X.te'':elteqky':~ X.eq x y'y0:(X.t * elt)%typel:list (X.t * elt)H0:InA eqke (y', e'') lInA eqke (y', e'') ((x, e') :: y0 :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt k' xIHt0:forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) l -> InA eqke (y0, e) (add x e' l)forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) ((k', y) :: l) -> InA eqke (y0, e) ((k', y) :: add x e' l)intros y' e'' eqky'; inversion_clear 1; intuition. Qed.elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt k' xIHt0:forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) l -> InA eqke (y0, e) (add x e' l)forall (y0 : X.t) (e : elt), ~ X.eq x y0 -> InA eqke (y0, e) ((k', y) :: l) -> InA eqke (y0, e) ((k', y) :: add x e' l)elt:Typeforall (m : t elt) (x y : X.t) (e e' : elt), ~ X.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e melt:Typeforall (m : t elt) (x y : X.t) (e e' : elt), ~ X.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e melt:Typem:t eltx, y:X.te, e':elt~ X.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e melt:Typem:t eltx:X.te':eltforall (y : X.t) (e : elt), ~ X.eq x y -> InA eqke (y, e) (add x e' m) -> InA eqke (y, e) melt:Typex:X.te':elty:X.te:eltH:~ X.eq x yH0:InA eqke (y, e) ((x, e') :: nil)InA eqke (y, e) nilelt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt x k'e0:X.compare x k' = LT _xy0:X.te:eltH:~ X.eq x y0H0:InA eqke (y0, e) ((x, e') :: (k', y) :: l)InA eqke (y0, e) ((k', y) :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.eq x k'e0:X.compare x k' = EQ _xy0:X.te:eltH:~ X.eq x y0H0:InA eqke (y0, e) ((x, e') :: l)InA eqke (y0, e) ((k', y) :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt k' xe0:X.compare x k' = GT _xIHt0:forall (y1 : X.t) (e1 : elt), ~ X.eq x y1 -> InA eqke (y1, e1) (add x e' l) -> InA eqke (y1, e1) ly0:X.te:eltH:~ X.eq x y0H0:InA eqke (y0, e) ((k', y) :: add x e' l)InA eqke (y0, e) ((k', y) :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt x k'e0:X.compare x k' = LT _xy0:X.te:eltH:~ X.eq x y0H0:InA eqke (y0, e) ((x, e') :: (k', y) :: l)InA eqke (y0, e) ((k', y) :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.eq x k'e0:X.compare x k' = EQ _xy0:X.te:eltH:~ X.eq x y0H0:InA eqke (y0, e) ((x, e') :: l)InA eqke (y0, e) ((k', y) :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt k' xe0:X.compare x k' = GT _xIHt0:forall (y1 : X.t) (e1 : elt), ~ X.eq x y1 -> InA eqke (y1, e1) (add x e' l) -> InA eqke (y1, e1) ly0:X.te:eltH:~ X.eq x y0H0:InA eqke (y0, e) ((k', y) :: add x e' l)InA eqke (y0, e) ((k', y) :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.eq x k'e0:X.compare x k' = EQ _xy0:X.te:eltH:~ X.eq x y0H0:InA eqke (y0, e) ((x, e') :: l)InA eqke (y0, e) ((k', y) :: l)elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt k' xe0:X.compare x k' = GT _xIHt0:forall (y1 : X.t) (e1 : elt), ~ X.eq x y1 -> InA eqke (y1, e1) (add x e' l) -> InA eqke (y1, e1) ly0:X.te:eltH:~ X.eq x y0H0:InA eqke (y0, e) ((k', y) :: add x e' l)InA eqke (y0, e) ((k', y) :: l)inversion_clear H0; auto. Qed.elt:Typex:X.te':eltk':X.ty:eltl:list (X.t * elt)_x:X.lt k' xe0:X.compare x k' = GT _xIHt0:forall (y1 : X.t) (e1 : elt), ~ X.eq x y1 -> InA eqke (y1, e1) (add x e' l) -> InA eqke (y1, e1) ly0:X.te:eltH:~ X.eq x y0H0:InA eqke (y0, e) ((k', y) :: add x e' l)InA eqke (y0, e) ((k', y) :: l)elt:Typeforall (m : t elt) (x x' : key) (e e' : elt), Inf (x', e') m -> ltk (x', e') (x, e) -> Inf (x', e') (add x e m)elt:Typeforall (m : t elt) (x x' : key) (e e' : elt), Inf (x', e') m -> ltk (x', e') (x, e) -> Inf (x', e') (add x e m)elt:Typeforall (x x' : key) (e e' : elt), Inf (x', e') nil -> ltk (x', e') (x, e) -> Inf (x', e') (add x e nil)elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:forall (x x' : key) (e e' : elt), Inf (x', e') m -> ltk (x', e') (x, e) -> Inf (x', e') (add x e m)forall (x x' : key) (e e' : elt), Inf (x', e') (a :: m) -> ltk (x', e') (x, e) -> Inf (x', e') (add x e (a :: m))elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:forall (x x' : key) (e e' : elt), Inf (x', e') m -> ltk (x', e') (x, e) -> Inf (x', e') (add x e m)forall (x x' : key) (e e' : elt), Inf (x', e') (a :: m) -> ltk (x', e') (x, e) -> Inf (x', e') (add x e (a :: m))elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:forall (x0 x'0 : key) (e0 e'0 : elt), Inf (x'0, e'0) m -> ltk (x'0, e'0) (x0, e0) -> Inf (x'0, e'0) (add x0 e0 m)x, x':keye, e':eltH:Inf (x', e') (a :: m)H0:ltk (x', e') (x, e)Inf (x', e') (add x e (a :: m))elt:Typex'':X.te'':eltm:list (X.t * elt)IHm:forall (x0 x'0 : key) (e0 e'0 : elt), Inf (x'0, e'0) m -> ltk (x'0, e'0) (x0, e0) -> Inf (x'0, e'0) (add x0 e0 m)x, x':keye, e':eltH:Inf (x', e') ((x'', e'') :: m)H0:ltk (x', e') (x, e)Inf (x', e') (add x e ((x'', e'') :: m))elt:Typex'':X.te'':eltm:list (X.t * elt)IHm:forall (x0 x'0 : key) (e0 e'0 : elt), Inf (x'0, e'0) m -> ltk (x'0, e'0) (x0, e0) -> Inf (x'0, e'0) (add x0 e0 m)x, x':keye, e':eltH0:ltk (x', e') (x, e)H1:ltk (x', e') (x'', e'')Inf (x', e') (add x e ((x'', e'') :: m))simpl; case (X.compare x x''); intuition. Qed. Hint Resolve add_Inf : core.elt:Typex'':X.te'':eltm:list (X.t * elt)IHm:forall (x0 x'0 : key) (e0 e'0 : elt), Inf (x'0, e'0) m -> ltk (x'0, e'0) (x0, e0) -> Inf (x'0, e'0) (add x0 e0 m)x, x':keye, e':eltH0:X.lt x' xH1:X.lt x' x''Inf (x', e') (add x e ((x'', e'') :: m))elt:Typeforall m : list (X.t * elt), Sort m -> forall (x : key) (e : elt), Sort (add x e m)elt:Typeforall m : list (X.t * elt), Sort m -> forall (x : key) (e : elt), Sort (add x e m)elt:TypeSort nil -> forall (x : key) (e : elt), Sort (add x e nil)elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sort m -> forall (x : key) (e : elt), Sort (add x e m)Sort (a :: m) -> forall (x : key) (e : elt), Sort (add x e (a :: m))elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sort m -> forall (x : key) (e : elt), Sort (add x e m)Sort (a :: m) -> forall (x : key) (e : elt), Sort (add x e (a :: m))elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sort m -> forall (x0 : key) (e0 : elt), Sort (add x0 e0 m)Hm:Sort (a :: m)x:keye:eltSort (add x e (a :: m))elt:Typex':X.te':eltm:list (X.t * elt)IHm:Sort m -> forall (x0 : key) (e0 : elt), Sort (add x0 e0 m)Hm:Sort ((x', e') :: m)x:keye:eltSort (add x e ((x', e') :: m))elt:Typex':X.te':eltm:list (X.t * elt)IHm:Sort m -> forall (x0 : key) (e1 : elt), Sort (add x0 e1 m)x:keye:elte0:X.eq x x'H:Sort mH0:Inf (x', e') mSort ((x, e) :: m)apply Inf_eq with (x',e'); auto. Qed.elt:Typex':X.te':eltm:list (X.t * elt)IHm:Sort m -> forall (x0 : key) (e1 : elt), Sort (add x0 e1 m)x:keye:elte0:X.eq x x'H:Sort mH0:Inf (x', e') mInf (x, e) m
Function remove (k : key) (s : t elt) {struct s} : t elt := match s with | nil => nil | (k',x) :: l => match X.compare k k' with | LT _ => s | EQ _ => l | GT _ => (k',x) :: remove k l end end.elt:Typeforall m : list (X.t * elt), Sort m -> forall x y : X.t, X.eq x y -> ~ In y (remove x m)elt:Typeforall m : list (X.t * elt), Sort m -> forall x y : X.t, X.eq x y -> ~ In y (remove x m)elt:Typem:list (X.t * elt)x, y:X.tSort m -> X.eq x y -> ~ In y (remove x m)elt:Typex, y:X.tHm:Sort nilH:X.eq x y~ In y nilelt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt x k'e0:X.compare x k' = LT _xHm:Sort ((k', x0) :: l)H:X.eq x y~ In y ((k', x0) :: l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.eq x k'e0:X.compare x k' = EQ _xHm:Sort ((k', x0) :: l)H:X.eq x y~ In y lelt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xe0:X.compare x k' = GT _xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)Hm:Sort ((k', x0) :: l)H:X.eq x y~ In y ((k', x0) :: remove x l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt x k'e0:X.compare x k' = LT _xHm:Sort ((k', x0) :: l)H:X.eq x y~ In y ((k', x0) :: l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.eq x k'e0:X.compare x k' = EQ _xHm:Sort ((k', x0) :: l)H:X.eq x y~ In y lelt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xe0:X.compare x k' = GT _xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)Hm:Sort ((k', x0) :: l)H:X.eq x y~ In y ((k', x0) :: remove x l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt x k'e0:X.compare x k' = LT _xHm:Sort ((k', x0) :: l)H:X.eq x yInf (y, x0) ((k', x0) :: l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.eq x k'e0:X.compare x k' = EQ _xHm:Sort ((k', x0) :: l)H:X.eq x y~ In y lelt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xe0:X.compare x k' = GT _xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)Hm:Sort ((k', x0) :: l)H:X.eq x y~ In y ((k', x0) :: remove x l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.eq x k'e0:X.compare x k' = EQ _xHm:Sort ((k', x0) :: l)H:X.eq x y~ In y lelt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xe0:X.compare x k' = GT _xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)Hm:Sort ((k', x0) :: l)H:X.eq x y~ In y ((k', x0) :: remove x l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.eq x k'H:X.eq x yH0:Sort lH1:Inf (k', x0) l~ In y lelt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xe0:X.compare x k' = GT _xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)Hm:Sort ((k', x0) :: l)H:X.eq x y~ In y ((k', x0) :: remove x l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.eq x k'H:X.eq x yH0:Sort lH1:Inf (k', x0) lInf (y, x0) lelt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xe0:X.compare x k' = GT _xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)Hm:Sort ((k', x0) :: l)H:X.eq x y~ In y ((k', x0) :: remove x l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xe0:X.compare x k' = GT _xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)Hm:Sort ((k', x0) :: l)H:X.eq x y~ In y ((k', x0) :: remove x l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)H:X.eq x yH0:Sort lH1:Inf (k', x0) l~ In y ((k', x0) :: remove x l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)H:X.eq x yH0:Sort lH1:Inf (k', x0) lnotin:~ In y (remove x l)~ In y ((k', x0) :: remove x l)elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)H:X.eq x yH0:Sort lH1:Inf (k', x0) lnotin:~ In y (remove x l)x1:eltabs:MapsTo y x1 ((k', x0) :: remove x l)Falseelt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)H:X.eq x yH0:Sort lH1:Inf (k', x0) lnotin:~ In y (remove x l)x1:eltH2:eqke (y, x1) (k', x0)Falseelt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)H:X.eq x yH0:Sort lH1:Inf (k', x0) lnotin:~ In y (remove x l)x1:eltH2:InA eqke (y, x1) (remove x l)Falseapply notin; exists x1; auto. Qed.elt:Typex, y, k':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xIHt0:Sort l -> X.eq x y -> ~ In y (remove x l)H:X.eq x yH0:Sort lH1:Inf (k', x0) lnotin:~ In y (remove x l)x1:eltH2:InA eqke (y, x1) (remove x l)Falseelt:Typeforall m : list (X.t * elt), Sort m -> forall (x y : X.t) (e : elt), ~ X.eq x y -> MapsTo y e m -> MapsTo y e (remove x m)elt:Typeforall m : list (X.t * elt), Sort m -> forall (x y : X.t) (e : elt), ~ X.eq x y -> MapsTo y e m -> MapsTo y e (remove x m)elt:Typem:list (X.t * elt)x, y:X.te:eltSort m -> ~ X.eq x y -> InA eqke (y, e) m -> InA eqke (y, e) (remove x m)elt:Typex, y:X.te:eltk':X.tx0:eltl:list (X.t * elt)_x:X.eq x k'Sort ((k', x0) :: l) -> ~ X.eq x y -> InA eqke (y, e) ((k', x0) :: l) -> InA eqke (y, e) lelt:Typex, y:X.te:eltk':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xIHt0:Sort l -> ~ X.eq x y -> InA eqke (y, e) l -> InA eqke (y, e) (remove x l)Sort ((k', x0) :: l) -> ~ X.eq x y -> InA eqke (y, e) ((k', x0) :: l) -> InA eqke (y, e) ((k', x0) :: remove x l)elt:Typex, y:X.te:eltk':X.tx0:eltl:list (X.t * elt)_x:X.eq x k'Hm:Sort ((k', x0) :: l)H:~ X.eq x yH1:eqke (y, e) (k', x0)InA eqke (y, e) lelt:Typex, y:X.te:eltk':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xIHt0:Sort l -> ~ X.eq x y -> InA eqke (y, e) l -> InA eqke (y, e) (remove x l)Sort ((k', x0) :: l) -> ~ X.eq x y -> InA eqke (y, e) ((k', x0) :: l) -> InA eqke (y, e) ((k', x0) :: remove x l)inversion_clear 1; inversion_clear 2; auto. Qed.elt:Typex, y:X.te:eltk':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xIHt0:Sort l -> ~ X.eq x y -> InA eqke (y, e) l -> InA eqke (y, e) (remove x l)Sort ((k', x0) :: l) -> ~ X.eq x y -> InA eqke (y, e) ((k', x0) :: l) -> InA eqke (y, e) ((k', x0) :: remove x l)elt:Typeforall m : list (X.t * elt), Sort m -> forall (x : key) (y : X.t) (e : elt), MapsTo y e (remove x m) -> MapsTo y e melt:Typeforall m : list (X.t * elt), Sort m -> forall (x : key) (y : X.t) (e : elt), MapsTo y e (remove x m) -> MapsTo y e melt:Typem:list (X.t * elt)x:keyy:X.te:eltSort m -> InA eqke (y, e) (remove x m) -> InA eqke (y, e) minversion_clear 1; inversion_clear 1; auto. Qed.elt:Typex:keyy:X.te:eltk':X.tx0:eltl:list (X.t * elt)_x:X.lt k' xe1:X.compare x k' = GT _xIHt0:Sort l -> InA eqke (y, e) (remove x l) -> InA eqke (y, e) lSort ((k', x0) :: l) -> InA eqke (y, e) ((k', x0) :: remove x l) -> InA eqke (y, e) ((k', x0) :: l)elt:Typeforall m : t elt, Sort m -> forall (x x' : key) (e' : elt), Inf (x', e') m -> Inf (x', e') (remove x m)elt:Typeforall m : t elt, Sort m -> forall (x x' : key) (e' : elt), Inf (x', e') m -> Inf (x', e') (remove x m)elt:TypeSort nil -> forall (x x' : key) (e' : elt), Inf (x', e') nil -> Inf (x', e') (remove x nil)elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sort m -> forall (x x' : key) (e' : elt), Inf (x', e') m -> Inf (x', e') (remove x m)Sort (a :: m) -> forall (x x' : key) (e' : elt), Inf (x', e') (a :: m) -> Inf (x', e') (remove x (a :: m))elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sort m -> forall (x x' : key) (e' : elt), Inf (x', e') m -> Inf (x', e') (remove x m)Sort (a :: m) -> forall (x x' : key) (e' : elt), Inf (x', e') (a :: m) -> Inf (x', e') (remove x (a :: m))elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sort m -> forall (x0 x'0 : key) (e'0 : elt), Inf (x'0, e'0) m -> Inf (x'0, e'0) (remove x0 m)Hm:Sort (a :: m)x, x':keye':eltH:Inf (x', e') (a :: m)Inf (x', e') (remove x (a :: m))elt:Typex'':X.te'':eltm:list (X.t * elt)IHm:Sort m -> forall (x0 x'0 : key) (e'0 : elt), Inf (x'0, e'0) m -> Inf (x'0, e'0) (remove x0 m)Hm:Sort ((x'', e'') :: m)x, x':keye':eltH:Inf (x', e') ((x'', e'') :: m)Inf (x', e') (remove x ((x'', e'') :: m))elt:Typex'':X.te'':eltm:list (X.t * elt)IHm:Sort m -> forall (x0 x'0 : key) (e'0 : elt), Inf (x'0, e'0) m -> Inf (x'0, e'0) (remove x0 m)Hm:Sort ((x'', e'') :: m)x, x':keye':eltH0:ltk (x', e') (x'', e'')Inf (x', e') (remove x ((x'', e'') :: m))elt:Typex'':X.te'':eltm:list (X.t * elt)IHm:Sort m -> forall (x0 x'0 : key) (e'0 : elt), Inf (x'0, e'0) m -> Inf (x'0, e'0) (remove x0 m)Hm:Sort ((x'', e'') :: m)x, x':keye':eltH0:X.lt x' x''Inf (x', e') (remove x ((x'', e'') :: m))elt:Typex'':X.te'':eltm:list (X.t * elt)IHm:Sort m -> forall (x0 x'0 : key) (e'0 : elt), Inf (x'0, e'0) m -> Inf (x'0, e'0) (remove x0 m)Hm:Sort ((x'', e'') :: m)x, x':keye':eltH0:X.lt x' x''e:X.eq x x''Inf (x', e') mapply Inf_lt with (x'',e''); auto. Qed. Hint Resolve remove_Inf : core.elt:Typex'':X.te'':eltm:list (X.t * elt)IHm:Sort m -> forall (x0 x'0 : key) (e'0 : elt), Inf (x'0, e'0) m -> Inf (x'0, e'0) (remove x0 m)x, x':keye':eltH0:X.lt x' x''e:X.eq x x''H:Sort mH1:Inf (x'', e'') mInf (x', e') melt:Typeforall m : list (X.t * elt), Sort m -> forall x : key, Sort (remove x m)elt:Typeforall m : list (X.t * elt), Sort m -> forall x : key, Sort (remove x m)elt:TypeSort nil -> forall x : key, Sort (remove x nil)elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sort m -> forall x : key, Sort (remove x m)Sort (a :: m) -> forall x : key, Sort (remove x (a :: m))elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sort m -> forall x : key, Sort (remove x m)Sort (a :: m) -> forall x : key, Sort (remove x (a :: m))elt:Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sort m -> forall x0 : key, Sort (remove x0 m)Hm:Sort (a :: m)x:keySort (remove x (a :: m))simpl; case (X.compare x x'); intuition; inversion_clear Hm; auto. Qed.elt:Typex':X.te':eltm:list (X.t * elt)IHm:Sort m -> forall x0 : key, Sort (remove x0 m)Hm:Sort ((x', e') :: m)x:keySort (remove x ((x', e') :: m))
Definition elements (m: t elt) := m.elt:Typeforall (m : list (X.t * elt)) (x : X.t) (e : elt), MapsTo x e m -> InA eqke (x, e) (elements m)auto. Qed.elt:Typeforall (m : list (X.t * elt)) (x : X.t) (e : elt), MapsTo x e m -> InA eqke (x, e) (elements m)elt:Typeforall (m : t elt) (x : X.t) (e : elt), InA eqke (x, e) (elements m) -> MapsTo x e mauto. Qed.elt:Typeforall (m : t elt) (x : X.t) (e : elt), InA eqke (x, e) (elements m) -> MapsTo x e melt:Typeforall m : list (X.t * elt), Sort m -> Sort (elements m)auto. Qed.elt:Typeforall m : list (X.t * elt), Sort m -> Sort (elements m)elt:Typeforall m : list (X.t * elt), Sort m -> NoDupA eqk (elements m)elt:Typeforall m : list (X.t * elt), Sort m -> NoDupA eqk (elements m)elt:Typem:list (X.t * elt)Hm:Sort mNoDupA eqk (elements m)apply elements_3; auto. Qed.elt:Typem:list (X.t * elt)Hm:Sort mSort (elements m)
Function fold (A:Type)(f:key->elt->A->A)(m:t elt) (acc:A) {struct m} : A := match m with | nil => acc | (k,e)::m' => fold f m' (f k e acc) end.elt:Typeforall (m : t elt) (A : Type) (i : A) (f : key -> elt -> A -> A), fold f m i = fold_left (fun (a : A) (p : key * elt) => f (fst p) (snd p) a) (elements m) iintros; functional induction (fold f m i); auto. Qed.elt:Typeforall (m : t elt) (A : Type) (i : A) (f : key -> elt -> A -> A), fold f m i = fold_left (fun (a : A) (p : key * elt) => f (fst p) (snd p) a) (elements m) i
Function equal (cmp:elt->elt->bool)(m m' : t elt) {struct m} : bool := match m, m' with | nil, nil => true | (x,e)::l, (x',e')::l' => match X.compare x x' with | EQ _ => cmp e e' && equal cmp l l' | _ => false end | _, _ => false end. Definition Equivb cmp m m' := (forall k, In k m <-> In k m') /\ (forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).elt:Typeforall m : list (X.t * elt), Sort m -> forall m' : list (X.t * elt), Sort m' -> forall cmp : elt -> elt -> bool, Equivb cmp m m' -> equal cmp m m' = trueelt:Typeforall m : list (X.t * elt), Sort m -> forall m' : list (X.t * elt), Sort m' -> forall cmp : elt -> elt -> bool, Equivb cmp m m' -> equal cmp m m' = trueelt:Typem, m':list (X.t * elt)cmp:elt -> elt -> boolSort m -> Sort m' -> Equivb cmp m m' -> equal cmp m m' = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'e2:X.compare x x' = EQ _xIHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp e e' && equal cmp l l' = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp e e' && equal cmp l l' = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp e e' = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truecmp e e' && equal cmp l l' = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truecmp e e' && equal cmp l l' = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = trueequal cmp l l' = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = trueSort lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = trueSort l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = trueEquivb cmp l l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = trueSort l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = trueEquivb cmp l l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = trueEquivb cmp l l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k lIn k l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k lH2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)In k l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k lH2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k lH2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x, e) :: l)In k l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k lH2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x, e) :: l)In k l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k lH2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x, e) :: l)H5:X.eq k x'In k l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k lH2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x, e) :: l)H5:X.eq k x'H6:Sort lH7:Inf (x, e) lIn k l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k lH2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x, e) :: l)H5:X.eq k x'H6:Sort lH7:Inf (x, e) lIn x lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.te'':elthyp:MapsTo k e'' lH2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x, e) :: l)H5:X.eq k x'H6:Sort lH7:Inf (x, e) lMapsTo x e'' lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'H2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'H2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)In k ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'H2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x', e') :: l')In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'H2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x', e') :: l')In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'H2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x', e') :: l')H5:X.eq k xIn k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'H2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x', e') :: l')H5:X.eq k xH6:Sort l'H7:Inf (x', e') l'In k lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.tH:In k l'H2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x', e') :: l')H5:X.eq k xH6:Sort l'H7:Inf (x', e') l'In x' l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 ((x, e) :: l) -> MapsTo k0 e'0 ((x', e') :: l') -> cmp e0 e'0 = truecmp_e_e':cmp e e' = truek:X.te'':elthyp:MapsTo k e'' l'H2:In k ((x, e) :: l) -> In k ((x', e') :: l')H3:In k ((x', e') :: l') -> In k ((x, e) :: l)H4:In k ((x', e') :: l')H5:X.eq k xH6:Sort l'H7:Inf (x', e') l'MapsTo x' e'' l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> Equivb cmp l l' -> equal cmp l l' = trueHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k0 : X.t, In k0 ((x, e) :: l) <-> In k0 ((x', e') :: l')H1:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 ((x, e) :: l) -> MapsTo k0 e'1 ((x', e') :: l') -> cmp e1 e'1 = truecmp_e_e':cmp e e' = truek:X.te0, e'0:eltH:MapsTo k e0 lH2:MapsTo k e'0 l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)y:match X.compare x x' with | EQ _ => False | _ => True endHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)false = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)In x ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)H3:In x ((x', e') :: l')false = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)In x ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)H3:In x ((x', e') :: l')false = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)H3:In x ((x', e') :: l')false = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)H3:In x ((x', e') :: l')H4:X.eq x x'false = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)H3:In x ((x', e') :: l')H4:In x l'false = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)H3:In x ((x', e') :: l')H4:In x l'false = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)H3:In x ((x', e') :: l')H4:In x l'H5:Sort l'H6:Inf (x', e') l'false = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)H3:In x ((x', e') :: l')H4:In x l'H5:Sort l'H6:Inf (x', e') l'Inf (x, e) l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)H3:In x ((x', e') :: l')H4:In x l'H5:Sort l'H6:Inf (x', e') l'H7:Inf (x, e) l'false = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x x'Hm:Sort ((x, e) :: l)H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x ((x, e) :: l) -> In x ((x', e') :: l')H2:In x ((x', e') :: l') -> In x ((x, e) :: l)H3:In x ((x', e') :: l')H4:In x l'H5:Sort l'H6:Inf (x', e') l'H7:Inf (x, e) l'false = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)false = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)In x' ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)H3:In x' ((x, e) :: l)false = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)In x' ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)H3:In x' ((x, e) :: l)false = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)H3:In x' ((x, e) :: l)false = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)H3:In x' ((x, e) :: l)H4:X.eq x' xfalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)H3:In x' ((x, e) :: l)H4:In x' lfalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)H3:In x' ((x, e) :: l)H4:In x' lfalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)H3:In x' ((x, e) :: l)H4:In x' lH5:Sort lH6:Inf (x, e) lfalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)H3:In x' ((x, e) :: l)H4:In x' lH5:Sort lH6:Inf (x, e) lInf (x', e') lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)H3:In x' ((x, e) :: l)H4:In x' lH5:Sort lH6:Inf (x, e) lH7:Inf (x', e') lfalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)Hlt:X.lt x' xHm':Sort ((x', e') :: l')H0:forall k : X.t, In k ((x, e) :: l) <-> In k ((x', e') :: l')H1:forall (k : X.t) (e0 e'0 : elt), MapsTo k e0 ((x, e) :: l) -> MapsTo k e'0 ((x', e') :: l') -> cmp e0 e'0 = trueH:In x' ((x, e) :: l) -> In x' ((x', e') :: l')H2:In x' ((x', e') :: l') -> In x' ((x, e) :: l)H3:In x' ((x, e) :: l)H4:In x' lH5:Sort lH6:Inf (x, e) lH7:Inf (x', e') lfalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolm, m':t elty:match m with | nil => match m' with | nil => False | _ :: _ => True end | (_, _) :: _ => match m' with | nil => True | (_, _) :: _ => False end endHm:Sort mHm':Sort m'H0:forall k : X.t, In k m <-> In k m'H1:forall (k : X.t) (e e' : elt), MapsTo k e m -> MapsTo k e' m' -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem':list (X.t * elt)y:TrueHm:Sort nilHm':Sort (p :: m')H0:forall k : X.t, In k nil <-> In k (p :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e nil -> MapsTo k e' (p :: m') -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)y:let (_, _) := p in TrueHm:Sort (p :: m)Hm':Sort nilH0:forall k : X.t, In k (p :: m) <-> In k nilH1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' nil -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)p0:(X.t * elt)%typem':list (X.t * elt)y:let (_, _) := p in let (_, _) := p0 in FalseHm:Sort (p :: m)Hm':Sort (p0 :: m')H0:forall k : X.t, In k (p :: m) <-> In k (p0 :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' (p0 :: m') -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolk:X.te:eltm':list (X.t * elt)y:TrueHm:Sort nilHm':Sort ((k, e) :: m')H0:forall k0 : X.t, In k0 nil <-> In k0 ((k, e) :: m')false = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)y:let (_, _) := p in TrueHm:Sort (p :: m)Hm':Sort nilH0:forall k : X.t, In k (p :: m) <-> In k nilH1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' nil -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)p0:(X.t * elt)%typem':list (X.t * elt)y:let (_, _) := p in let (_, _) := p0 in FalseHm:Sort (p :: m)Hm':Sort (p0 :: m')H0:forall k : X.t, In k (p :: m) <-> In k (p0 :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' (p0 :: m') -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolk:X.te:eltm':list (X.t * elt)y:TrueHm:Sort nilHm':Sort ((k, e) :: m')H0:forall k0 : X.t, In k0 nil <-> In k0 ((k, e) :: m')H:In k nil -> In k ((k, e) :: m')H1:In k ((k, e) :: m') -> In k nilfalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)y:let (_, _) := p in TrueHm:Sort (p :: m)Hm':Sort nilH0:forall k : X.t, In k (p :: m) <-> In k nilH1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' nil -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)p0:(X.t * elt)%typem':list (X.t * elt)y:let (_, _) := p in let (_, _) := p0 in FalseHm:Sort (p :: m)Hm':Sort (p0 :: m')H0:forall k : X.t, In k (p :: m) <-> In k (p0 :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' (p0 :: m') -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolk:X.te:eltm':list (X.t * elt)y:TrueHm:Sort nilHm':Sort ((k, e) :: m')H0:forall k0 : X.t, In k0 nil <-> In k0 ((k, e) :: m')H:In k nil -> In k ((k, e) :: m')In k ((k, e) :: m')elt:Typecmp:elt -> elt -> boolk:X.te:eltm':list (X.t * elt)y:TrueHm:Sort nilHm':Sort ((k, e) :: m')H0:forall k0 : X.t, In k0 nil <-> In k0 ((k, e) :: m')H:In k nil -> In k ((k, e) :: m')x:eltH1:MapsTo k x nilfalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)y:let (_, _) := p in TrueHm:Sort (p :: m)Hm':Sort nilH0:forall k : X.t, In k (p :: m) <-> In k nilH1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' nil -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)p0:(X.t * elt)%typem':list (X.t * elt)y:let (_, _) := p in let (_, _) := p0 in FalseHm:Sort (p :: m)Hm':Sort (p0 :: m')H0:forall k : X.t, In k (p :: m) <-> In k (p0 :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' (p0 :: m') -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolk:X.te:eltm':list (X.t * elt)y:TrueHm:Sort nilHm':Sort ((k, e) :: m')H0:forall k0 : X.t, In k0 nil <-> In k0 ((k, e) :: m')H:In k nil -> In k ((k, e) :: m')x:eltH1:MapsTo k x nilfalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)y:let (_, _) := p in TrueHm:Sort (p :: m)Hm':Sort nilH0:forall k : X.t, In k (p :: m) <-> In k nilH1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' nil -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)p0:(X.t * elt)%typem':list (X.t * elt)y:let (_, _) := p in let (_, _) := p0 in FalseHm:Sort (p :: m)Hm':Sort (p0 :: m')H0:forall k : X.t, In k (p :: m) <-> In k (p0 :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' (p0 :: m') -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)y:let (_, _) := p in TrueHm:Sort (p :: m)Hm':Sort nilH0:forall k : X.t, In k (p :: m) <-> In k nilH1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' nil -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)p0:(X.t * elt)%typem':list (X.t * elt)y:let (_, _) := p in let (_, _) := p0 in FalseHm:Sort (p :: m)Hm':Sort (p0 :: m')H0:forall k : X.t, In k (p :: m) <-> In k (p0 :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' (p0 :: m') -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltm:list (X.t * elt)y:TrueHm:Sort ((x, e) :: m)Hm':Sort nilH0:forall k : X.t, In k ((x, e) :: m) <-> In k nilH1:forall (k : X.t) (e0 e' : elt), MapsTo k e0 ((x, e) :: m) -> MapsTo k e' nil -> cmp e0 e' = truefalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)p0:(X.t * elt)%typem':list (X.t * elt)y:let (_, _) := p in let (_, _) := p0 in FalseHm:Sort (p :: m)Hm':Sort (p0 :: m')H0:forall k : X.t, In k (p :: m) <-> In k (p0 :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' (p0 :: m') -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltm:list (X.t * elt)y:TrueHm:Sort ((x, e) :: m)Hm':Sort nilH0:forall k : X.t, In k ((x, e) :: m) <-> In k nilH1:forall (k : X.t) (e0 e' : elt), MapsTo k e0 ((x, e) :: m) -> MapsTo k e' nil -> cmp e0 e' = trueH:In x ((x, e) :: m) -> In x nilH2:In x nil -> In x ((x, e) :: m)false = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)p0:(X.t * elt)%typem':list (X.t * elt)y:let (_, _) := p in let (_, _) := p0 in FalseHm:Sort (p :: m)Hm':Sort (p0 :: m')H0:forall k : X.t, In k (p :: m) <-> In k (p0 :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' (p0 :: m') -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltm:list (X.t * elt)y:TrueHm:Sort ((x, e) :: m)Hm':Sort nilH0:forall k : X.t, In k ((x, e) :: m) <-> In k nilH1:forall (k : X.t) (e0 e' : elt), MapsTo k e0 ((x, e) :: m) -> MapsTo k e' nil -> cmp e0 e' = trueH2:In x nil -> In x ((x, e) :: m)In x ((x, e) :: m)elt:Typecmp:elt -> elt -> boolx:X.te:eltm:list (X.t * elt)y:TrueHm:Sort ((x, e) :: m)Hm':Sort nilH0:forall k : X.t, In k ((x, e) :: m) <-> In k nilH1:forall (k : X.t) (e0 e' : elt), MapsTo k e0 ((x, e) :: m) -> MapsTo k e' nil -> cmp e0 e' = trueH2:In x nil -> In x ((x, e) :: m)x0:eltH:MapsTo x x0 nilfalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)p0:(X.t * elt)%typem':list (X.t * elt)y:let (_, _) := p in let (_, _) := p0 in FalseHm:Sort (p :: m)Hm':Sort (p0 :: m')H0:forall k : X.t, In k (p :: m) <-> In k (p0 :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' (p0 :: m') -> cmp e e' = truefalse = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltm:list (X.t * elt)y:TrueHm:Sort ((x, e) :: m)Hm':Sort nilH0:forall k : X.t, In k ((x, e) :: m) <-> In k nilH1:forall (k : X.t) (e0 e' : elt), MapsTo k e0 ((x, e) :: m) -> MapsTo k e' nil -> cmp e0 e' = trueH2:In x nil -> In x ((x, e) :: m)x0:eltH:MapsTo x x0 nilfalse = trueelt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)p0:(X.t * elt)%typem':list (X.t * elt)y:let (_, _) := p in let (_, _) := p0 in FalseHm:Sort (p :: m)Hm':Sort (p0 :: m')H0:forall k : X.t, In k (p :: m) <-> In k (p0 :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' (p0 :: m') -> cmp e e' = truefalse = truedestruct p;destruct p0;contradiction. Qed.elt:Typecmp:elt -> elt -> boolp:(X.t * elt)%typem:list (X.t * elt)p0:(X.t * elt)%typem':list (X.t * elt)y:let (_, _) := p in let (_, _) := p0 in FalseHm:Sort (p :: m)Hm':Sort (p0 :: m')H0:forall k : X.t, In k (p :: m) <-> In k (p0 :: m')H1:forall (k : X.t) (e e' : elt), MapsTo k e (p :: m) -> MapsTo k e' (p0 :: m') -> cmp e e' = truefalse = trueelt:Typeforall m : list (X.t * elt), Sort m -> forall m' : list (X.t * elt), Sort m' -> forall cmp : elt -> elt -> bool, equal cmp m m' = true -> Equivb cmp m m'elt:Typeforall m : list (X.t * elt), Sort m -> forall m' : list (X.t * elt), Sort m' -> forall cmp : elt -> elt -> bool, equal cmp m m' = true -> Equivb cmp m m'elt:Typem, m':list (X.t * elt)cmp:elt -> elt -> boolSort m -> Sort m' -> equal cmp m m' = true -> Equivb cmp m m'elt:Typecmp:elt -> elt -> boolHm, Hm':Sort nilH:true = truek:X.te, e':eltH0:MapsTo k e nilH1:MapsTo k e' nilcmp e e' = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x, e) :: l)In k ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x', e') :: l')In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x, e) :: l)In k ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x', e') :: l')In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x, e) :: l)H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'In k ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x', e') :: l')In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x, e) :: l)H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueIn k ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x', e') :: l')In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x, e) :: l)H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueIn k ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x', e') :: l')In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x, e) :: l)H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueH8:X.eq k xIn k ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x, e) :: l)H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueH8:In k lIn k ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x', e') :: l')In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x, e) :: l)H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueH8:In k lIn k ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x', e') :: l')In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x, e) :: l)H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueH8:In k lH9:In k l -> In k l'H10:In k l' -> In k lIn k ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x', e') :: l')In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x, e) :: l)H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueH8:In k lH9:In k l -> In k l'H10:In k l' -> In k le'':elthyp:MapsTo k e'' l'In k ((x', e') :: l')elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x', e') :: l')In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x', e') :: l')In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'H:cmp e e' && equal cmp l l' = truek:X.tH0:In k ((x', e') :: l')H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x', e') :: l')H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueIn k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x', e') :: l')H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueIn k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x', e') :: l')H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueH8:X.eq k x'In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x', e') :: l')H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueH8:In k l'In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x', e') :: l')H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueH8:In k l'In k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x', e') :: l')H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueH8:In k l'H9:In k l -> In k l'H10:In k l' -> In k lIn k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.tH0:In k ((x', e') :: l')H1:Sort lH2:Inf (x, e) lH3:Sort l'H4:Inf (x', e') l'H5:cmp e e' = trueH6:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H7:forall (k0 : X.t) (e0 e'0 : elt), MapsTo k0 e0 l -> MapsTo k0 e'0 l' -> cmp e0 e'0 = trueH8:In k l'H9:In k l -> In k l'H10:In k l' -> In k le'':elthyp:MapsTo k e'' lIn k ((x, e) :: l)elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'Hm:Sort ((x, e) :: l)Hm':Sort ((x', e') :: l')H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'H:cmp e e' && equal cmp l l' = truek:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')H2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')H2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = truecmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH0:MapsTo k e0 ((x, e) :: l)H1:MapsTo k e'0 ((x', e') :: l')H2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = truecmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH1:MapsTo k e'0 ((x', e') :: l')H2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:eqke (k, e0) (x, e)cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH1:MapsTo k e'0 ((x', e') :: l')H2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:InA eqke (k, e0) lcmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te'0:eltH1:MapsTo k e'0 ((x', e') :: l')H2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e0 e'1 : elt), MapsTo k0 e0 l -> MapsTo k0 e'1 l' -> cmp e0 e'1 = trueH0:X.eq k xcmp e e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH1:MapsTo k e'0 ((x', e') :: l')H2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:InA eqke (k, e0) lcmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te'0:eltH2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e0 e'1 : elt), MapsTo k0 e0 l -> MapsTo k0 e'1 l' -> cmp e0 e'1 = trueH0:X.eq k xH9:eqke (k, e'0) (x', e')cmp e e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te'0:eltH2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e0 e'1 : elt), MapsTo k0 e0 l -> MapsTo k0 e'1 l' -> cmp e0 e'1 = trueH0:X.eq k xH9:InA eqke (k, e'0) l'cmp e e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH1:MapsTo k e'0 ((x', e') :: l')H2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:InA eqke (k, e0) lcmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te'0:eltH2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e0 e'1 : elt), MapsTo k0 e0 l -> MapsTo k0 e'1 l' -> cmp e0 e'1 = trueH0:X.eq k xH9:InA eqke (k, e'0) l'cmp e e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH1:MapsTo k e'0 ((x', e') :: l')H2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:InA eqke (k, e0) lcmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te'0:eltH2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e0 e'1 : elt), MapsTo k0 e0 l -> MapsTo k0 e'1 l' -> cmp e0 e'1 = trueH0:X.eq k xH9:InA eqke (k, e'0) l'In x' l'elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH1:MapsTo k e'0 ((x', e') :: l')H2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:InA eqke (k, e0) lcmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH1:MapsTo k e'0 ((x', e') :: l')H2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:InA eqke (k, e0) lcmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:InA eqke (k, e0) lH0:eqke (k, e'0) (x', e')cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:InA eqke (k, e0) lH0:InA eqke (k, e'0) l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0:eltH2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'0 : elt), MapsTo k0 e1 l -> MapsTo k0 e'0 l' -> cmp e1 e'0 = trueH9:InA eqke (k, e0) lH0:X.eq k x'cmp e0 e' = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:InA eqke (k, e0) lH0:InA eqke (k, e'0) l'cmp e0 e'0 = trueelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0:eltH2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'0 : elt), MapsTo k0 e1 l -> MapsTo k0 e'0 l' -> cmp e1 e'0 = trueH9:InA eqke (k, e0) lH0:X.eq k x'In x lelt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:InA eqke (k, e0) lH0:InA eqke (k, e'0) l'cmp e0 e'0 = trueapply H8 with k; auto. Qed.elt:Typecmp:elt -> elt -> boolx:X.te:eltl:list (X.t * elt)x':X.te':eltl':list (X.t * elt)_x:X.eq x x'IHb:Sort l -> Sort l' -> equal cmp l l' = true -> Equivb cmp l l'k:X.te0, e'0:eltH2:Sort lH3:Inf (x, e) lH4:Sort l'H5:Inf (x', e') l'H6:cmp e e' = trueH7:equal cmp l l' = trueH:forall k0 : X.t, In k0 l <-> In k0 l'H8:forall (k0 : X.t) (e1 e'1 : elt), MapsTo k0 e1 l -> MapsTo k0 e'1 l' -> cmp e1 e'1 = trueH9:InA eqke (k, e0) lH0:InA eqke (k, e'0) l'cmp e0 e'0 = true
This lemma isn't part of the spec of Equivb, but is used in FMapAVL
elt:Typeforall (cmp : elt -> elt -> bool) (l1 l2 : list (X.t * elt)) (x y : X.t * elt), Sort (x :: l1) -> Sort (y :: l2) -> eqk x y -> cmp (snd x) (snd y) = true -> Equivb cmp l1 l2 <-> Equivb cmp (x :: l1) (y :: l2)elt:Typeforall (cmp : elt -> elt -> bool) (l1 l2 : list (X.t * elt)) (x y : X.t * elt), Sort (x :: l1) -> Sort (y :: l2) -> eqk x y -> cmp (snd x) (snd y) = true -> Equivb cmp l1 l2 <-> Equivb cmp (x :: l1) (y :: l2)elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)x, y:(X.t * elt)%typeH:Sort (x :: l1)H0:Sort (y :: l2)H1:eqk x yH2:cmp (snd x) (snd y) = trueEquivb cmp l1 l2 <-> Equivb cmp (x :: l1) (y :: l2)elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)x, y:(X.t * elt)%typeH:Sort (x :: l1)H0:Sort (y :: l2)H1:eqk x yH2:cmp (snd x) (snd y) = trueH5:Sort l1H6:Inf x l1Equivb cmp l1 l2 <-> Equivb cmp (x :: l1) (y :: l2)elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)x, y:(X.t * elt)%typeH:Sort (x :: l1)H0:Sort (y :: l2)H1:eqk x yH2:cmp (snd x) (snd y) = trueH5:Sort l1H6:Inf x l1H7:Sort l2H8:Inf y l2Equivb cmp l1 l2 <-> Equivb cmp (x :: l1) (y :: l2)elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2Equivb cmp l1 l2 <-> Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp l1 l2Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)Equivb cmp l1 l2elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp l1 l2equal cmp ((t0, e) :: l1) ((t1, e0) :: l2) = trueelt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)Equivb cmp l1 l2elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp l1 l2match X.compare t0 t1 with | EQ _ => cmp e e0 && equal cmp l1 l2 | _ => false end = trueelt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)Equivb cmp l1 l2elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp l1 l2H4:X.eq t0 t1cmp e e0 && equal cmp l1 l2 = trueelt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)Equivb cmp l1 l2elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp l1 l2H4:X.eq t0 t1equal cmp l1 l2 = trueelt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)Equivb cmp l1 l2elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)Equivb cmp l1 l2elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)equal cmp l1 l2 = trueelt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)equal cmp ((t0, e) :: l1) ((t1, e0) :: l2) = true -> equal cmp l1 l2 = trueelt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)match X.compare t0 t1 with | EQ _ => cmp e e0 && equal cmp l1 l2 | _ => false end = true -> equal cmp l1 l2 = truerewrite H2; simpl; auto. Qed. Variable elt':Type.elt:Typecmp:elt -> elt -> booll1, l2:list (X.t * elt)t0:X.te:eltt1:X.te0:eltH:Sort ((t0, e) :: l1)H0:Sort ((t1, e0) :: l2)H1:X.eq t0 t1H2:cmp e e0 = trueH5:Sort l1H6:Inf (t0, e) l1H7:Sort l2H8:Inf (t1, e0) l2H3:Equivb cmp ((t0, e) :: l1) ((t1, e0) :: l2)H4:X.eq t0 t1cmp e e0 && equal cmp l1 l2 = true -> equal cmp l1 l2 = true
Fixpoint map (f:elt -> elt') (m:t elt) : t elt' := match m with | nil => nil | (k,e)::m' => (k,f e) :: map f m' end. Fixpoint mapi (f: key -> elt -> elt') (m:t elt) : t elt' := match m with | nil => nil | (k,e)::m' => (k,f k e) :: mapi f m' end. End Elt. Section Elt2. (* A new section is necessary for previous definitions to work with different [elt], especially [MapsTo]... *) Variable elt elt' : Type.
Specification of map
elt, elt':Typeforall (m : t elt) (x : key) (e : elt) (f : elt -> elt'), MapsTo x e m -> MapsTo x (f e) (map f m)elt, elt':Typeforall (m : t elt) (x : key) (e : elt) (f : elt -> elt'), MapsTo x e m -> MapsTo x (f e) (map f m)(* functional induction map elt elt' f m. *) (* Marche pas ??? *)elt, elt':Typem:t eltx:keye:eltf:elt -> elt'MapsTo x e m -> MapsTo x (f e) (map f m)elt, elt':Typex:keye:eltf:elt -> elt'MapsTo x e nil -> MapsTo x (f e) (map f nil)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)x:keye:eltf:elt -> elt'IHm:MapsTo x e m -> MapsTo x (f e) (map f m)MapsTo x e (a :: m) -> MapsTo x (f e) (map f (a :: m))elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)x:keye:eltf:elt -> elt'IHm:MapsTo x e m -> MapsTo x (f e) (map f m)MapsTo x e (a :: m) -> MapsTo x (f e) (map f (a :: m))elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:elt -> elt'IHm:MapsTo x e m -> MapsTo x (f e) (map f m)MapsTo x e ((x', e') :: m) -> MapsTo x (f e) (map f ((x', e') :: m))elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:elt -> elt'IHm:MapsTo x e m -> MapsTo x (f e) (map f m)MapsTo x e ((x', e') :: m) -> MapsTo x (f e) ((x', f e') :: map f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:elt -> elt'IHm:MapsTo x e m -> MapsTo x (f e) (map f m)H0:eqke (x, e) (x', e')MapsTo x (f e) ((x', f e') :: map f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:elt -> elt'IHm:MapsTo x e m -> MapsTo x (f e) (map f m)H0:InA (eqke (elt:=elt)) (x, e) mMapsTo x (f e) ((x', f e') :: map f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:elt -> elt'IHm:MapsTo x e m -> MapsTo x (f e) (map f m)H0:eqke (x, e) (x', e')eqke (x, f e) (x', f e')elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:elt -> elt'IHm:MapsTo x e m -> MapsTo x (f e) (map f m)H0:InA (eqke (elt:=elt)) (x, e) mMapsTo x (f e) ((x', f e') :: map f m)unfold MapsTo in *; auto. Qed.elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:elt -> elt'IHm:MapsTo x e m -> MapsTo x (f e) (map f m)H0:InA (eqke (elt:=elt)) (x, e) mMapsTo x (f e) ((x', f e') :: map f m)elt, elt':Typeforall (m : t elt) (x : key) (f : elt -> elt'), In x (map f m) -> In x melt, elt':Typeforall (m : t elt) (x : key) (f : elt -> elt'), In x (map f m) -> In x m(* functional induction map elt elt' f m. *) (* Marche pas ??? *)elt, elt':Typem:t eltx:keyf:elt -> elt'In x (map f m) -> In x melt, elt':Typex:keyf:elt -> elt'In x nil -> In x nilelt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x mIn x (let (k, e) := a in (k, f e) :: map f m) -> In x (a :: m)elt, elt':Typex:keyf:elt -> elt'e:elt'abs:MapsTo x e nilIn x nilelt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x mIn x (let (k, e) := a in (k, f e) :: map f m) -> In x (a :: m)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x mIn x (let (k, e) := a in (k, f e) :: map f m) -> In x (a :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x mIn x ((x', f e) :: map f m) -> In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x mhyp:In x ((x', f e) :: map f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x mhyp:In x ((x', f e) :: map f m)x0:elt'H:MapsTo x x0 ((x', f e) :: map f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x mx0:elt'H:MapsTo x x0 ((x', f e) :: map f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x me':elt'H:MapsTo x e' ((x', f e) :: map f m)H1:eqke (x, e') (x', f e)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x me':elt'H:MapsTo x e' ((x', f e) :: map f m)H1:InA (eqke (elt:=elt')) (x, e') (map f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x me':elt'H:MapsTo x e' ((x', f e) :: map f m)H1:eqke (x, e') (x', f e)eqke (x, e) (x', e)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x me':elt'H:MapsTo x e' ((x', f e) :: map f m)H1:InA (eqke (elt:=elt')) (x, e') (map f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'IHm:In x (map f m) -> In x me':elt'H:MapsTo x e' ((x', f e) :: map f m)H1:InA (eqke (elt:=elt')) (x, e') (map f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'e':elt'H:MapsTo x e' ((x', f e) :: map f m)H1:InA (eqke (elt:=elt')) (x, e') (map f m)In x (map f m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'e':elt'H:MapsTo x e' ((x', f e) :: map f m)H1:InA (eqke (elt:=elt')) (x, e') (map f m)e'':elthyp:MapsTo x e'' mIn x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'e':elt'H:MapsTo x e' ((x', f e) :: map f m)H1:InA (eqke (elt:=elt')) (x, e') (map f m)e'':elthyp:MapsTo x e'' mIn x ((x', e) :: m)constructor 2; auto. Qed.elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:elt -> elt'e':elt'H:MapsTo x e' ((x', f e) :: map f m)H1:InA (eqke (elt:=elt')) (x, e') (map f m)e'':elthyp:MapsTo x e'' mMapsTo x e'' ((x', e) :: m)elt, elt':Typeforall (m : t elt) (x : key) (e : elt) (e' : elt') (f : elt -> elt'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') (map f m)elt, elt':Typeforall (m : t elt) (x : key) (e : elt) (e' : elt') (f : elt -> elt'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') (map f m)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)IHm:forall (x : key) (e : elt) (e' : elt') (f : elt -> elt'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') (map f m)forall (x : key) (e : elt) (e' : elt') (f : elt -> elt'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') (let (k, e0) := a in (k, f e0) :: map f m)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)IHm:forall (x0 : key) (e0 : elt) (e'0 : elt') (f0 : elt -> elt'), HdRel (ltk (elt:=elt)) (x0, e0) m -> HdRel (ltk (elt:=elt')) (x0, e'0) (map f0 m)x:keye:elte':elt'f:elt -> elt'H:HdRel (ltk (elt:=elt)) (x, e) (a :: m)HdRel (ltk (elt:=elt')) (x, e') (let (k, e0) := a in (k, f e0) :: map f m)inversion_clear H; auto. Qed. Hint Resolve map_lelistA : core.elt, elt':Typex0:X.te0:eltm:list (X.t * elt)IHm:forall (x1 : key) (e1 : elt) (e'0 : elt') (f0 : elt -> elt'), HdRel (ltk (elt:=elt)) (x1, e1) m -> HdRel (ltk (elt:=elt')) (x1, e'0) (map f0 m)x:keye:elte':elt'f:elt -> elt'H:HdRel (ltk (elt:=elt)) (x, e) ((x0, e0) :: m)HdRel (ltk (elt:=elt')) (x, e') ((x0, f e0) :: map f m)elt, elt':Typeforall m : t elt, Sorted (ltk (elt:=elt)) m -> forall f : elt -> elt', Sorted (ltk (elt:=elt')) (map f m)elt, elt':Typeforall m : t elt, Sorted (ltk (elt:=elt)) m -> forall f : elt -> elt', Sorted (ltk (elt:=elt')) (map f m)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall f : elt -> elt', Sorted (ltk (elt:=elt')) (map f m)Sorted (ltk (elt:=elt)) (a :: m) -> forall f : elt -> elt', Sorted (ltk (elt:=elt')) (let (k, e) := a in (k, f e) :: map f m)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall f0 : elt -> elt', Sorted (ltk (elt:=elt')) (map f0 m)Hm:Sorted (ltk (elt:=elt)) (a :: m)f:elt -> elt'Sorted (ltk (elt:=elt')) (let (k, e) := a in (k, f e) :: map f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall f0 : elt -> elt', Sorted (ltk (elt:=elt')) (map f0 m)Hm:Sorted (ltk (elt:=elt)) ((x', e') :: m)f:elt -> elt'Sorted (ltk (elt:=elt')) ((x', f e') :: map f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall f0 : elt -> elt', Sorted (ltk (elt:=elt')) (map f0 m)f:elt -> elt'H:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (x', e') mSorted (ltk (elt:=elt')) ((x', f e') :: map f m)exact (map_lelistA _ _ H0). Qed.elt, elt':Typex':X.te':eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall f0 : elt -> elt', Sorted (ltk (elt:=elt')) (map f0 m)f:elt -> elt'H:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (x', e') mHdRel (ltk (elt:=elt')) (x', f e') (map f m)
Specification of mapi
elt, elt':Typeforall (m : t elt) (x : key) (e : elt) (f : key -> elt -> elt'), MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)elt, elt':Typeforall (m : t elt) (x : key) (e : elt) (f : key -> elt -> elt'), MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)(* functional induction mapi elt elt' f m. *) (* Marche pas ??? *)elt, elt':Typem:t eltx:keye:eltf:key -> elt -> elt'MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)elt, elt':Typex:keye:eltf:key -> elt -> elt'MapsTo x e nil -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f nil)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)MapsTo x e (a :: m) -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f (a :: m))elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)MapsTo x e (a :: m) -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f (a :: m))elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)MapsTo x e ((x', e') :: m) -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f ((x', e') :: m))elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)MapsTo x e ((x', e') :: m) -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) ((x', f x' e') :: mapi f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)H0:eqke (x, e) (x', e')exists y : X.t, X.eq y x /\ MapsTo x (f y e) ((x', f x' e') :: mapi f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)H0:InA (eqke (elt:=elt)) (x, e) mexists y : X.t, X.eq y x /\ MapsTo x (f y e) ((x', f x' e') :: mapi f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)H0:eqke (x, e) (x', e')X.eq x' x /\ MapsTo x (f x' e) ((x', f x' e') :: mapi f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)H0:InA (eqke (elt:=elt)) (x, e) mexists y : X.t, X.eq y x /\ MapsTo x (f y e) ((x', f x' e') :: mapi f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)H:X.eq x x'H0:e = e'X.eq x' x /\ MapsTo x (f x' e) ((x', f x' e') :: mapi f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)H0:InA (eqke (elt:=elt)) (x, e) mexists y : X.t, X.eq y x /\ MapsTo x (f y e) ((x', f x' e') :: mapi f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)H:X.eq x x'H0:e = e'MapsTo x (f x' e) ((x', f x' e') :: mapi f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)H0:InA (eqke (elt:=elt)) (x, e) mexists y : X.t, X.eq y x /\ MapsTo x (f y e) ((x', f x' e') :: mapi f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)H:X.eq x x'H0:e = e'eqke (x, f x' e) (x', f x' e')elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)H0:InA (eqke (elt:=elt)) (x, e) mexists y : X.t, X.eq y x /\ MapsTo x (f y e) ((x', f x' e') :: mapi f m)elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'IHm:MapsTo x e m -> exists y : X.t, X.eq y x /\ MapsTo x (f y e) (mapi f m)H0:InA (eqke (elt:=elt)) (x, e) mexists y : X.t, X.eq y x /\ MapsTo x (f y e) ((x', f x' e') :: mapi f m)exists y; intuition. Qed.elt, elt':Typex':X.te':eltm:list (X.t * elt)x:keye:eltf:key -> elt -> elt'H0:InA (eqke (elt:=elt)) (x, e) my:X.thyp:X.eq y x /\ MapsTo x (f y e) (mapi f m)exists y0 : X.t, X.eq y0 x /\ MapsTo x (f y0 e) ((x', f x' e') :: mapi f m)elt, elt':Typeforall (m : t elt) (x : key) (f : key -> elt -> elt'), In x (mapi f m) -> In x melt, elt':Typeforall (m : t elt) (x : key) (f : key -> elt -> elt'), In x (mapi f m) -> In x m(* functional induction mapi elt elt' f m. *) (* Marche pas ??? *)elt, elt':Typem:t eltx:keyf:key -> elt -> elt'In x (mapi f m) -> In x melt, elt':Typex:keyf:key -> elt -> elt'In x nil -> In x nilelt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x mIn x (let (k, e) := a in (k, f k e) :: mapi f m) -> In x (a :: m)elt, elt':Typex:keyf:key -> elt -> elt'e:elt'abs:MapsTo x e nilIn x nilelt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x mIn x (let (k, e) := a in (k, f k e) :: mapi f m) -> In x (a :: m)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x mIn x (let (k, e) := a in (k, f k e) :: mapi f m) -> In x (a :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x mIn x ((x', f x' e) :: mapi f m) -> In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x mhyp:In x ((x', f x' e) :: mapi f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x mhyp:In x ((x', f x' e) :: mapi f m)x0:elt'H:MapsTo x x0 ((x', f x' e) :: mapi f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x mx0:elt'H:MapsTo x x0 ((x', f x' e) :: mapi f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x me':elt'H:MapsTo x e' ((x', f x' e) :: mapi f m)H1:eqke (x, e') (x', f x' e)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x me':elt'H:MapsTo x e' ((x', f x' e) :: mapi f m)H1:InA (eqke (elt:=elt')) (x, e') (mapi f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x me':elt'H:MapsTo x e' ((x', f x' e) :: mapi f m)H1:eqke (x, e') (x', f x' e)eqke (x, e) (x', e)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x me':elt'H:MapsTo x e' ((x', f x' e) :: mapi f m)H1:InA (eqke (elt:=elt')) (x, e') (mapi f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'IHm:In x (mapi f m) -> In x me':elt'H:MapsTo x e' ((x', f x' e) :: mapi f m)H1:InA (eqke (elt:=elt')) (x, e') (mapi f m)In x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'e':elt'H:MapsTo x e' ((x', f x' e) :: mapi f m)H1:InA (eqke (elt:=elt')) (x, e') (mapi f m)In x (mapi f m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'e':elt'H:MapsTo x e' ((x', f x' e) :: mapi f m)H1:InA (eqke (elt:=elt')) (x, e') (mapi f m)e'':elthyp:MapsTo x e'' mIn x ((x', e) :: m)elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'e':elt'H:MapsTo x e' ((x', f x' e) :: mapi f m)H1:InA (eqke (elt:=elt')) (x, e') (mapi f m)e'':elthyp:MapsTo x e'' mIn x ((x', e) :: m)constructor 2; auto. Qed.elt, elt':Typex':X.te:eltm:list (X.t * elt)x:keyf:key -> elt -> elt'e':elt'H:MapsTo x e' ((x', f x' e) :: mapi f m)H1:InA (eqke (elt:=elt')) (x, e') (mapi f m)e'':elthyp:MapsTo x e'' mMapsTo x e'' ((x', e) :: m)elt, elt':Typeforall (m : t elt) (x : key) (e : elt) (f : key -> elt -> elt'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, f x e) (mapi f m)elt, elt':Typeforall (m : t elt) (x : key) (e : elt) (f : key -> elt -> elt'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, f x e) (mapi f m)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)IHm:forall (x : key) (e : elt) (f : key -> elt -> elt'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, f x e) (mapi f m)forall (x : key) (e : elt) (f : key -> elt -> elt'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, f x e) (let (k, e0) := a in (k, f k e0) :: mapi f m)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)IHm:forall (x0 : key) (e0 : elt) (f0 : key -> elt -> elt'), HdRel (ltk (elt:=elt)) (x0, e0) m -> HdRel (ltk (elt:=elt')) (x0, f0 x0 e0) (mapi f0 m)x:keye:eltf:key -> elt -> elt'H:HdRel (ltk (elt:=elt)) (x, e) (a :: m)HdRel (ltk (elt:=elt')) (x, f x e) (let (k, e0) := a in (k, f k e0) :: mapi f m)inversion_clear H; auto. Qed. Hint Resolve mapi_lelistA : core.elt, elt':Typex':X.te':eltm:list (X.t * elt)IHm:forall (x0 : key) (e0 : elt) (f0 : key -> elt -> elt'), HdRel (ltk (elt:=elt)) (x0, e0) m -> HdRel (ltk (elt:=elt')) (x0, f0 x0 e0) (mapi f0 m)x:keye:eltf:key -> elt -> elt'H:HdRel (ltk (elt:=elt)) (x, e) ((x', e') :: m)HdRel (ltk (elt:=elt')) (x, f x e) ((x', f x' e') :: mapi f m)elt, elt':Typeforall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall f : key -> elt -> elt', Sorted (ltk (elt:=elt')) (mapi f m)elt, elt':Typeforall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall f : key -> elt -> elt', Sorted (ltk (elt:=elt')) (mapi f m)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall f : key -> elt -> elt', Sorted (ltk (elt:=elt')) (mapi f m)Sorted (ltk (elt:=elt)) (a :: m) -> forall f : key -> elt -> elt', Sorted (ltk (elt:=elt')) (let (k, e) := a in (k, f k e) :: mapi f m)elt, elt':Typea:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall f0 : key -> elt -> elt', Sorted (ltk (elt:=elt')) (mapi f0 m)Hm:Sorted (ltk (elt:=elt)) (a :: m)f:key -> elt -> elt'Sorted (ltk (elt:=elt')) (let (k, e) := a in (k, f k e) :: mapi f m)inversion_clear Hm; auto. Qed. End Elt2. Section Elt3.elt, elt':Typex':X.te':eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall f0 : key -> elt -> elt', Sorted (ltk (elt:=elt')) (mapi f0 m)Hm:Sorted (ltk (elt:=elt)) ((x', e') :: m)f:key -> elt -> elt'Sorted (ltk (elt:=elt')) ((x', f x' e') :: mapi f m)
Variable elt elt' elt'' : Type. Variable f : option elt -> option elt' -> option elt''. Definition option_cons (A:Type)(k:key)(o:option A)(l:list (key*A)) := match o with | Some e => (k,e)::l | None => l end. Fixpoint map2_l (m : t elt) : t elt'' := match m with | nil => nil | (k,e)::l => option_cons k (f (Some e) None) (map2_l l) end. Fixpoint map2_r (m' : t elt') : t elt'' := match m' with | nil => nil | (k,e')::l' => option_cons k (f None (Some e')) (map2_r l') end. Fixpoint map2 (m : t elt) : t elt' -> t elt'' := match m with | nil => map2_r | (k,e) :: l => fix map2_aux (m' : t elt') : t elt'' := match m' with | nil => map2_l m | (k',e') :: l' => match X.compare k k' with | LT _ => option_cons k (f (Some e) None) (map2 l m') | EQ _ => option_cons k (f (Some e) (Some e')) (map2 l l') | GT _ => option_cons k' (f None (Some e')) (map2_aux l') end end end. Notation oee' := (option elt * option elt')%type. Fixpoint combine (m : t elt) : t elt' -> t oee' := match m with | nil => map (fun e' => (None,Some e')) | (k,e) :: l => fix combine_aux (m':t elt') : list (key * oee') := match m' with | nil => map (fun e => (Some e,None)) m | (k',e') :: l' => match X.compare k k' with | LT _ => (k,(Some e, None))::combine l m' | EQ _ => (k,(Some e, Some e'))::combine l l' | GT _ => (k',(None,Some e'))::combine_aux l' end end end. Definition fold_right_pair (A B C:Type)(f: A->B->C->C)(l:list (A*B))(i:C) := List.fold_right (fun p => f (fst p) (snd p)) i l. Definition map2_alt m m' := let m0 : t oee' := combine m m' in let m1 : t (option elt'') := map (fun p => f (fst p) (snd p)) m0 in fold_right_pair (option_cons (A:=elt'')) m1 nil.elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall (m : t elt) (m' : t elt'), map2_alt m m' = map2 m m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall (m : t elt) (m' : t elt'), map2_alt m m' = map2 m m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall (m : t elt) (m' : t elt'), fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil = map2 m m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m' : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine nil m')) nil = map2 nil m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall m' : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil = map2 m m'forall m' : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine (a :: m) m')) nil = map2 (a :: m) m'(* map2_r *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m':t elt'fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (map (fun e' : elt' => (None, Some e')) m')) nil = map2_r m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall m' : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil = map2 m m'forall m' : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine (a :: m) m')) nil = map2 (a :: m) m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:elt'm':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (map (fun e' : elt' => (None, Some e')) m')) nil = map2_r m'option_cons t0 (f None (Some e)) (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (map (fun e' : elt' => (None, Some e')) m')) nil) = option_cons t0 (f None (Some e)) (map2_r m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall m' : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil = map2 m m'forall m' : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine (a :: m) m')) nil = map2 (a :: m) m'(* fin map2_r *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall m' : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil = map2 m m'forall m' : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine (a :: m) m')) nil = map2 (a :: m) m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m' : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil = map2 m m'fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) nil)) nil = map2 ((t0, e) :: m) nilelt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0a0:(X.t * elt')%typem':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) (a0 :: m'))) nil = map2 ((t0, e) :: m) (a0 :: m')(* map2_l *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m' : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil = map2 m m'fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (map (fun e0 : elt => (Some e0, None)) m)) nil = map2_l melt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0a0:(X.t * elt')%typem':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) (a0 :: m'))) nil = map2 ((t0, e) :: m) (a0 :: m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (map (fun e0 : elt => (Some e0, None)) m)) nil = map2_l melt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0a0:(X.t * elt')%typem':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) (a0 :: m'))) nil = map2 ((t0, e) :: m) (a0 :: m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltt1:X.te0:eltm:list (X.t * elt)IHm:fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (map (fun e1 : elt => (Some e1, None)) m)) nil = map2_l moption_cons t1 (f (Some e0) None) (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (map (fun e1 : elt => (Some e1, None)) m)) nil) = option_cons t1 (f (Some e0) None) (map2_l m)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0a0:(X.t * elt')%typem':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) (a0 :: m'))) nil = map2 ((t0, e) :: m) (a0 :: m')(* fin map2_l *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0a0:(X.t * elt')%typem':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) (a0 :: m'))) nil = map2 ((t0, e) :: m) (a0 :: m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0t1:X.te0:elt'm':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) ((t1, e0) :: m'))) nil = map2 ((t0, e) :: m) ((t1, e0) :: m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0t1:X.te0:elt'm':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) match X.compare t0 t1 with | LT _ => (t0, (Some e, None)) :: combine m ((t1, e0) :: m') | EQ _ => (t0, (Some e, Some e0)) :: combine m m' | GT _ => (t1, (None, Some e0)) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (t0, (Some e, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k', e') :: l' => match X.compare t0 k' with | LT _ => (t0, (Some e, None)) :: combine m m'0 | EQ _ => (t0, (Some e, Some e')) :: combine m l' | GT _ => (k', (None, Some e')) :: combine_aux l' end end) m' end) nil = match X.compare t0 t1 with | LT _ => option_cons t0 (f (Some e) None) (map2 m ((t1, e0) :: m')) | EQ _ => option_cons t0 (f (Some e) (Some e0)) (map2 m m') | GT _ => option_cons t1 (f None (Some e0)) ((fix map2_aux (m'0 : t elt') : t elt'' := match m'0 with | nil => option_cons t0 (f (Some e) None) (map2_l m) | (k', e') :: l' => match X.compare t0 k' with | LT _ => option_cons t0 (f (Some e) None) (map2 m m'0) | EQ _ => option_cons t0 (f (Some e) (Some e')) (map2 m l') | GT _ => option_cons k' (f None (Some e')) (map2_aux l') end end) m') endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0t1:X.te0:elt'm':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'l:X.lt t0 t1fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m ((t1, e0) :: m'))) nil = map2 m ((t1, e0) :: m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0t1:X.te0:elt'm':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'e1:X.eq t0 t1fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil = map2 m m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0t1:X.te0:elt'm':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'l:X.lt t1 t0fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (t0, (Some e, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k', e') :: l' => match X.compare t0 k' with | LT _ => (t0, (Some e, None)) :: combine m m'0 | EQ _ => (t0, (Some e, Some e')) :: combine m l' | GT _ => (k', (None, Some e')) :: combine_aux l' end end) m')) nil = (fix map2_aux (m'0 : t elt') : t elt'' := match m'0 with | nil => option_cons t0 (f (Some e) None) (map2_l m) | (k', e') :: l' => match X.compare t0 k' with | LT _ => option_cons t0 (f (Some e) None) (map2 m m'0) | EQ _ => option_cons t0 (f (Some e) (Some e')) (map2 m l') | GT _ => option_cons k' (f None (Some e')) (map2_aux l') end end) m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0t1:X.te0:elt'm':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'e1:X.eq t0 t1fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil = map2 m m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0t1:X.te0:elt'm':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'l:X.lt t1 t0fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (t0, (Some e, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k', e') :: l' => match X.compare t0 k' with | LT _ => (t0, (Some e, None)) :: combine m m'0 | EQ _ => (t0, (Some e, Some e')) :: combine m l' | GT _ => (k', (None, Some e')) :: combine_aux l' end end) m')) nil = (fix map2_aux (m'0 : t elt') : t elt'' := match m'0 with | nil => option_cons t0 (f (Some e) None) (map2_l m) | (k', e') :: l' => match X.compare t0 k' with | LT _ => option_cons t0 (f (Some e) None) (map2 m m'0) | EQ _ => option_cons t0 (f (Some e) (Some e')) (map2 m l') | GT _ => option_cons k' (f None (Some e')) (map2_aux l') end end) m'apply IHm'. Qed.elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:forall m'0 : t elt', fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m'0)) nil = map2 m m'0t1:X.te0:elt'm':list (X.t * elt')IHm':fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine ((t0, e) :: m) m')) nil = map2 ((t0, e) :: m) m'l:X.lt t1 t0fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (t0, (Some e, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k', e') :: l' => match X.compare t0 k' with | LT _ => (t0, (Some e, None)) :: combine m m'0 | EQ _ => (t0, (Some e, Some e')) :: combine m l' | GT _ => (k', (None, Some e')) :: combine_aux l' end end) m')) nil = (fix map2_aux (m'0 : t elt') : t elt'' := match m'0 with | nil => option_cons t0 (f (Some e) None) (map2_l m) | (k', e') :: l' => match X.compare t0 k' with | LT _ => option_cons t0 (f (Some e) None) (map2 m m'0) | EQ _ => option_cons t0 (f (Some e) (Some e')) (map2 m l') | GT _ => option_cons k' (f None (Some e')) (map2_aux l') end end) m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall (m : list (X.t * elt)) (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall (m : list (X.t * elt)) (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) nil -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine nil m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m')forall (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m':list (X.t * elt')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) nilH0:HdRel (ltk (elt:=elt')) (x, e') m'HdRel (ltk (elt:=oee')) (x, e'') (combine nil m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m')forall (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m':list (X.t * elt')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) nilH0:HdRel (ltk (elt:=elt')) (x, e') m'HdRel (ltk (elt:=oee')) (x, e'') (map (fun e'0 : elt' => (None, Some e'0)) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m')forall (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m')forall (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m' : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m')forall (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') nil -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m'0 -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m'0)a0:(X.t * elt')%typem':list (X.t * elt')IHm':forall (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) m')forall (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') (a0 :: m') -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m' : list (X.t * elt')) (x0 : key) (e0 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e0) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) (a :: m)H0:HdRel (ltk (elt:=elt')) (x, e') nilHdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m'0 -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m'0)a0:(X.t * elt')%typem':list (X.t * elt')IHm':forall (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) m')forall (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') (a0 :: m') -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te0:eltm:list (X.t * elt)IHm:forall (m' : list (X.t * elt')) (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) ((t0, e0) :: m)H0:HdRel (ltk (elt:=elt')) (x, e') nilHdRel (ltk (elt:=oee')) (x, e'') (combine ((t0, e0) :: m) nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m'0 -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m'0)a0:(X.t * elt')%typem':list (X.t * elt')IHm':forall (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) m')forall (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') (a0 :: m') -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te0:eltm:list (X.t * elt)IHm:forall (m' : list (X.t * elt')) (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) ((t0, e0) :: m)H0:HdRel (ltk (elt:=elt')) (x, e') nilHdRel (ltk (elt:=oee')) (x, e'') (map (fun e1 : elt => (Some e1, None)) ((t0, e0) :: m))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m'0 -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m'0)a0:(X.t * elt')%typem':list (X.t * elt')IHm':forall (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) m')forall (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') (a0 :: m') -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) m -> HdRel (ltk (elt:=elt')) (x, e') m'0 -> HdRel (ltk (elt:=oee')) (x, e'') (combine m m'0)a0:(X.t * elt')%typem':list (X.t * elt')IHm':forall (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') m' -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) m')forall (x : key) (e : elt) (e' : elt') (e'' : oee'), HdRel (ltk (elt:=elt)) (x, e) (a :: m) -> HdRel (ltk (elt:=elt')) (x, e') (a0 :: m') -> HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x0 : key) (e0 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e0) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m'0 -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m'0)a0:(X.t * elt')%typem':list (X.t * elt')IHm':forall (x0 : key) (e0 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e0) (a :: m) -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine (a :: m) m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) (a :: m)H0:HdRel (ltk (elt:=elt')) (x, e') (a0 :: m')HdRel (ltk (elt:=oee')) (x, e'') (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x0 : key) (e0 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e0) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m'0 -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m'0)a0:(X.t * elt')%typem':list (X.t * elt')IHm':forall (x0 : key) (e0 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e0) (a :: m) -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine (a :: m) m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) (a :: m)H0:HdRel (ltk (elt:=elt')) (x, e') (a0 :: m')HdRel (ltk (elt:=oee')) (x, e'') ((let (k, e0) := a in fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => let (k0, e1) := a in (k0, (Some e1, None)) :: map (fun e2 : elt => (Some e2, None)) m | (k', e'0) :: l' => match X.compare k k' with | LT _ => (k, (Some e0, None)) :: combine m m'0 | EQ _ => (k, (Some e0, Some e'0)) :: combine m l' | GT _ => (k', (None, Some e'0)) :: combine_aux l' end end) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te0:eltm:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m'0 -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m'0)k':X.te0':elt'm':list (X.t * elt')IHm':forall (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) ((k, e0) :: m) -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine ((k, e0) :: m) m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) ((k, e0) :: m)H0:HdRel (ltk (elt:=elt')) (x, e') ((k', e0') :: m')HdRel (ltk (elt:=oee')) (x, e'') match X.compare k k' with | LT _ => (k, (Some e0, None)) :: combine m ((k', e0') :: m') | EQ _ => (k, (Some e0, Some e0')) :: combine m m' | GT _ => (k', (None, Some e0')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e0, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e0, None)) :: combine m m'0 | EQ _ => (k, (Some e0, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m' endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te0:eltm:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m'0 -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m'0)k':X.te0':elt'm':list (X.t * elt')IHm':forall (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) ((k, e0) :: m) -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine ((k, e0) :: m) m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) ((k, e0) :: m)H0:HdRel (ltk (elt:=elt')) (x, e') ((k', e0') :: m')l:X.lt k k'HdRel (ltk (elt:=oee')) (x, e'') ((k, (Some e0, None)) :: combine m ((k', e0') :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te0:eltm:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x0 : key) (e2 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e2) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m'0 -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m'0)k':X.te0':elt'm':list (X.t * elt')IHm':forall (x0 : key) (e2 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e2) ((k, e0) :: m) -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine ((k, e0) :: m) m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) ((k, e0) :: m)H0:HdRel (ltk (elt:=elt')) (x, e') ((k', e0') :: m')e1:X.eq k k'HdRel (ltk (elt:=oee')) (x, e'') ((k, (Some e0, Some e0')) :: combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te0:eltm:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m'0 -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m'0)k':X.te0':elt'm':list (X.t * elt')IHm':forall (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) ((k, e0) :: m) -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine ((k, e0) :: m) m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) ((k, e0) :: m)H0:HdRel (ltk (elt:=elt')) (x, e') ((k', e0') :: m')l:X.lt k' kHdRel (ltk (elt:=oee')) (x, e'') ((k', (None, Some e0')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e0, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e0, None)) :: combine m m'0 | EQ _ => (k, (Some e0, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te0:eltm:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x0 : key) (e2 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e2) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m'0 -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m'0)k':X.te0':elt'm':list (X.t * elt')IHm':forall (x0 : key) (e2 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e2) ((k, e0) :: m) -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine ((k, e0) :: m) m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) ((k, e0) :: m)H0:HdRel (ltk (elt:=elt')) (x, e') ((k', e0') :: m')e1:X.eq k k'HdRel (ltk (elt:=oee')) (x, e'') ((k, (Some e0, Some e0')) :: combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te0:eltm:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m'0 -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m'0)k':X.te0':elt'm':list (X.t * elt')IHm':forall (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) ((k, e0) :: m) -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine ((k, e0) :: m) m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) ((k, e0) :: m)H0:HdRel (ltk (elt:=elt')) (x, e') ((k', e0') :: m')l:X.lt k' kHdRel (ltk (elt:=oee')) (x, e'') ((k', (None, Some e0')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e0, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e0, None)) :: combine m m'0 | EQ _ => (k, (Some e0, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')inversion_clear H0; auto. Qed. Hint Resolve combine_lelistA : core.elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te0:eltm:list (X.t * elt)IHm:forall (m'0 : list (X.t * elt')) (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) m -> HdRel (ltk (elt:=elt')) (x0, e'0) m'0 -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine m m'0)k':X.te0':elt'm':list (X.t * elt')IHm':forall (x0 : key) (e1 : elt) (e'0 : elt') (e''0 : oee'), HdRel (ltk (elt:=elt)) (x0, e1) ((k, e0) :: m) -> HdRel (ltk (elt:=elt')) (x0, e'0) m' -> HdRel (ltk (elt:=oee')) (x0, e''0) (combine ((k, e0) :: m) m')x:keye:elte':elt'e'':oee'H:HdRel (ltk (elt:=elt)) (x, e) ((k, e0) :: m)H0:HdRel (ltk (elt:=elt')) (x, e') ((k', e0') :: m')l:X.lt k' kHdRel (ltk (elt:=oee')) (x, e'') ((k', (None, Some e0')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e0, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e0, None)) :: combine m m'0 | EQ _ => (k, (Some e0, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''Sorted (ltk (elt:=elt)) nil -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine nil m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine m m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine (a :: m) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'Sorted (ltk (elt:=oee')) (combine nil m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine m m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine (a :: m) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'Sorted (ltk (elt:=oee')) (map (fun e' : elt' => (None, Some e')) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine m m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine (a :: m) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine m m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine (a :: m) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine m m')Hm:Sorted (ltk (elt:=elt)) (a :: m)Sorted (ltk (elt:=elt')) nil -> Sorted (ltk (elt:=oee')) (combine (a :: m) nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine (a :: m) m')Sorted (ltk (elt:=elt')) (a0 :: m') -> Sorted (ltk (elt:=oee')) (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine m m')Hm:Sorted (ltk (elt:=elt)) (a :: m)Sorted (ltk (elt:=oee')) (combine (a :: m) nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine (a :: m) m')Sorted (ltk (elt:=elt')) (a0 :: m') -> Sorted (ltk (elt:=oee')) (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine m m')Hm:Sorted (ltk (elt:=elt)) ((t0, e) :: m)Sorted (ltk (elt:=oee')) (combine ((t0, e) :: m) nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine (a :: m) m')Sorted (ltk (elt:=elt')) (a0 :: m') -> Sorted (ltk (elt:=oee')) (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine m m')Hm:Sorted (ltk (elt:=elt)) ((t0, e) :: m)Sorted (ltk (elt:=oee')) (map (fun e0 : elt => (Some e0, None)) ((t0, e) :: m))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine (a :: m) m')Sorted (ltk (elt:=elt')) (a0 :: m') -> Sorted (ltk (elt:=oee')) (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine (a :: m) m')Sorted (ltk (elt:=elt')) (a0 :: m') -> Sorted (ltk (elt:=oee')) (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine (a :: m) m')Hm':Sorted (ltk (elt:=elt')) (a0 :: m')Sorted (ltk (elt:=oee')) (combine (a :: m) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine (a :: m) m')Hm':Sorted (ltk (elt:=elt')) (a0 :: m')Sorted (ltk (elt:=oee')) ((let (k, e) := a in fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => let (k0, e0) := a in (k0, (Some e0, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k', e') :: l' => match X.compare k k' with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e')) :: combine m l' | GT _ => (k', (None, Some e')) :: combine_aux l' end end) (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Sorted (ltk (elt:=oee')) match X.compare k k' with | LT _ => (k, (Some e, None)) :: combine m ((k', e') :: m') | EQ _ => (k, (Some e, Some e')) :: combine m m' | GT _ => (k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m' endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k k'Sorted (ltk (elt:=oee')) ((k, (Some e, None)) :: combine m ((k', e') :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Heq:X.eq k k'Sorted (ltk (elt:=oee')) ((k, (Some e, Some e')) :: combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k' kSorted (ltk (elt:=oee')) ((k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k k'H:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (k, e) mSorted (ltk (elt:=oee')) ((k, (Some e, None)) :: combine m ((k', e') :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Heq:X.eq k k'Sorted (ltk (elt:=oee')) ((k, (Some e, Some e')) :: combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k' kSorted (ltk (elt:=oee')) ((k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k k'H:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (k, e) mHdRel (ltk (elt:=oee')) (k, (Some e, None)) (combine m ((k', e') :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Heq:X.eq k k'Sorted (ltk (elt:=oee')) ((k, (Some e, Some e')) :: combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k' kSorted (ltk (elt:=oee')) ((k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k k'H:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (k, e) mH1:HdRel (ltk (elt:=elt')) (k, e') ((k', e') :: m')HdRel (ltk (elt:=oee')) (k, (Some e, None)) (combine m ((k', e') :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Heq:X.eq k k'Sorted (ltk (elt:=oee')) ((k, (Some e, Some e')) :: combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k' kSorted (ltk (elt:=oee')) ((k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Heq:X.eq k k'Sorted (ltk (elt:=oee')) ((k, (Some e, Some e')) :: combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k' kSorted (ltk (elt:=oee')) ((k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Heq:X.eq k k'H:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (k, e) mH1:Sorted (ltk (elt:=elt')) m'H2:HdRel (ltk (elt:=elt')) (k', e') m'Sorted (ltk (elt:=oee')) ((k, (Some e, Some e')) :: combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k' kSorted (ltk (elt:=oee')) ((k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Heq:X.eq k k'H:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (k, e) mH1:Sorted (ltk (elt:=elt')) m'H2:HdRel (ltk (elt:=elt')) (k', e') m'HdRel (ltk (elt:=oee')) (k, (Some e, Some e')) (combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k' kSorted (ltk (elt:=oee')) ((k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Heq:X.eq k k'H:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (k, e) mH1:Sorted (ltk (elt:=elt')) m'H2:HdRel (ltk (elt:=elt')) (k', e') m'H3:HdRel (ltk (elt:=elt')) (k, e') m'HdRel (ltk (elt:=oee')) (k, (Some e, Some e')) (combine m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k' kSorted (ltk (elt:=oee')) ((k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')Hlt:X.lt k' kSorted (ltk (elt:=oee')) ((k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hlt:X.lt k' kH:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (k, e) mH1:Sorted (ltk (elt:=elt')) m'H2:HdRel (ltk (elt:=elt')) (k', e') m'Sorted (ltk (elt:=oee')) ((k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hlt:X.lt k' kH:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (k, e) mH1:Sorted (ltk (elt:=elt')) m'H2:HdRel (ltk (elt:=elt')) (k', e') m'HdRel (ltk (elt:=oee')) (k', (None, Some e')) ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hlt:X.lt k' kH:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (k, e) mH1:Sorted (ltk (elt:=elt')) m'H2:HdRel (ltk (elt:=elt')) (k', e') m'HdRel (ltk (elt:=oee')) (k', (None, Some e')) (combine ((k, e) :: m) m')exact (combine_lelistA _ H3 H2). Qed.elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> Sorted (ltk (elt:=oee')) (combine m m'0)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=oee')) (combine ((k, e) :: m) m')Hlt:X.lt k' kH:Sorted (ltk (elt:=elt)) mH0:HdRel (ltk (elt:=elt)) (k, e) mH1:Sorted (ltk (elt:=elt')) m'H2:HdRel (ltk (elt:=elt')) (k', e') m'H3:HdRel (ltk (elt:=elt)) (k', e) ((k, e) :: m)HdRel (ltk (elt:=oee')) (k', (None, Some e')) (combine ((k, e) :: m) m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=elt'')) (map2 m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> Sorted (ltk (elt:=elt'')) (map2 m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'Sorted (ltk (elt:=elt'')) (map2 m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'Sorted (ltk (elt:=elt'')) (map2_alt m m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'H0:Sorted (ltk (elt:=oee')) (combine m m')Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'l0:t oee'H0:Sorted (ltk (elt:=oee')) l0Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) l0) nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'l0:t oee'H0:Sorted (ltk (elt:=oee')) l0f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) (map f' l0) nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'l0:t oee'H0:Sorted (ltk (elt:=oee')) l0f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''H1:Sorted (ltk (elt:=option elt'')) (map f' l0)Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) (map f' l0) nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'l0:t oee'H0:Sorted (ltk (elt:=oee')) l0f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''l1:t (option elt'')H1:Sorted (ltk (elt:=option elt'')) l1Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) l1 nil)elt, elt', elt'':Typel1:t (option elt'')H1:Sorted (ltk (elt:=option elt'')) l1Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) l1 nil)elt, elt', elt'':TypeH1:Sorted (ltk (elt:=option elt'')) nilSorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) nil nil)elt, elt', elt'':Typea:(X.t * option elt'')%typel1:list (X.t * option elt'')H1:Sorted (ltk (elt:=option elt'')) (a :: l1)IHl1:Sorted (ltk (elt:=option elt'')) l1 -> Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) l1 nil)Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) (a :: l1) nil)elt, elt', elt'':Typea:(X.t * option elt'')%typel1:list (X.t * option elt'')H1:Sorted (ltk (elt:=option elt'')) (a :: l1)IHl1:Sorted (ltk (elt:=option elt'')) l1 -> Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) l1 nil)Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) (a :: l1) nil)elt, elt', elt'':Typea:(X.t * option elt'')%typel1:list (X.t * option elt'')IHl1:Sorted (ltk (elt:=option elt'')) l1 -> Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) l1 nil)H:Sorted (ltk (elt:=option elt'')) l1H0:HdRel (ltk (elt:=option elt'')) a l1Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) (a :: l1) nil)elt, elt', elt'':Typet0:X.te:elt''l1:list (X.t * option elt'')IHl1:Sorted (ltk (elt:=option elt'')) l1 -> Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) l1 nil)H:Sorted (ltk (elt:=option elt'')) l1H0:HdRel (ltk (elt:=option elt'')) (t0, Some e) l1Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) ((t0, Some e) :: l1) nil)elt, elt', elt'':Typet0:X.te:elt''l1:list (X.t * option elt'')IHl1:Sorted (ltk (elt:=option elt'')) l1 -> Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) l1 nil)H:Sorted (ltk (elt:=option elt'')) l1H0:HdRel (ltk (elt:=option elt'')) (t0, Some e) l1Sorted (ltk (elt:=elt'')) ((t0, e) :: fold_right_pair (option_cons (A:=elt'')) l1 nil)elt, elt', elt'':Typet0:X.te:elt''l1:list (X.t * option elt'')IHl1:Sorted (ltk (elt:=option elt'')) l1 -> Sorted (ltk (elt:=elt'')) (fold_right_pair (option_cons (A:=elt'')) l1 nil)H:Sorted (ltk (elt:=option elt'')) l1H0:HdRel (ltk (elt:=option elt'')) (t0, Some e) l1HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)elt, elt', elt'':Typet0:X.te:elt''l1:list (X.t * option elt'')H:Sorted (ltk (elt:=option elt'')) l1H0:HdRel (ltk (elt:=option elt'')) (t0, Some e) l1HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)elt, elt', elt'':Typet0:X.te:elt''H:Sorted (ltk (elt:=option elt'')) nilH0:HdRel (ltk (elt:=option elt'')) (t0, Some e) nilHdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) nil nil)elt, elt', elt'':Typet0:X.te:elt''a:(X.t * option elt'')%typel1:list (X.t * option elt'')H:Sorted (ltk (elt:=option elt'')) (a :: l1)H0:HdRel (ltk (elt:=option elt'')) (t0, Some e) (a :: l1)IHl1:Sorted (ltk (elt:=option elt'')) l1 -> HdRel (ltk (elt:=option elt'')) (t0, Some e) l1 -> HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) (a :: l1) nil)elt, elt', elt'':Typet0:X.te:elt''a:(X.t * option elt'')%typel1:list (X.t * option elt'')H:Sorted (ltk (elt:=option elt'')) (a :: l1)H0:HdRel (ltk (elt:=option elt'')) (t0, Some e) (a :: l1)IHl1:Sorted (ltk (elt:=option elt'')) l1 -> HdRel (ltk (elt:=option elt'')) (t0, Some e) l1 -> HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) (a :: l1) nil)elt, elt', elt'':Typet0:X.te:elt''t1:X.te0:elt''l1:list (X.t * option elt'')H:Sorted (ltk (elt:=option elt'')) ((t1, Some e0) :: l1)H0:HdRel (ltk (elt:=option elt'')) (t0, Some e) ((t1, Some e0) :: l1)IHl1:Sorted (ltk (elt:=option elt'')) l1 -> HdRel (ltk (elt:=option elt'')) (t0, Some e) l1 -> HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)HdRel (ltk (elt:=elt'')) (t0, e) ((t1, e0) :: fold_right_pair (option_cons (A:=elt'')) l1 nil)elt, elt', elt'':Typet0:X.te:elt''t1:X.tl1:list (X.t * option elt'')H:Sorted (ltk (elt:=option elt'')) ((t1, None) :: l1)H0:HdRel (ltk (elt:=option elt'')) (t0, Some e) ((t1, None) :: l1)IHl1:Sorted (ltk (elt:=option elt'')) l1 -> HdRel (ltk (elt:=option elt'')) (t0, Some e) l1 -> HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)elt, elt', elt'':Typet0:X.te:elt''t1:X.tl1:list (X.t * option elt'')H:Sorted (ltk (elt:=option elt'')) ((t1, None) :: l1)H0:HdRel (ltk (elt:=option elt'')) (t0, Some e) ((t1, None) :: l1)IHl1:Sorted (ltk (elt:=option elt'')) l1 -> HdRel (ltk (elt:=option elt'')) (t0, Some e) l1 -> HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)elt, elt', elt'':Typet0:X.te:elt''t1:X.tl1:list (X.t * option elt'')H:Sorted (ltk (elt:=option elt'')) ((t1, None) :: l1)IHl1:Sorted (ltk (elt:=option elt'')) l1 -> HdRel (ltk (elt:=option elt'')) (t0, Some e) l1 -> HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)H1:ltk (t0, Some e) (t1, None)HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)elt, elt', elt'':Typet0:X.te:elt''t1:X.tl1:list (X.t * option elt'')H:Sorted (ltk (elt:=option elt'')) ((t1, None) :: l1)IHl1:Sorted (ltk (elt:=option elt'')) l1 -> HdRel (ltk (elt:=option elt'')) (t0, Some e) l1 -> HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)H1:X.lt t0 t1HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)elt, elt', elt'':Typet0:X.te:elt''t1:X.tl1:list (X.t * option elt'')IHl1:Sorted (ltk (elt:=option elt'')) l1 -> HdRel (ltk (elt:=option elt'')) (t0, Some e) l1 -> HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)H1:X.lt t0 t1H0:Sorted (ltk (elt:=option elt'')) l1H2:HdRel (ltk (elt:=option elt'')) (t1, None) l1HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)apply Inf_lt with (t1, None (A:=elt'')); auto. Qed. Definition at_least_one (o:option elt)(o':option elt') := match o, o' with | None, None => None | _, _ => Some (o,o') end.elt, elt', elt'':Typet0:X.te:elt''t1:X.tl1:list (X.t * option elt'')IHl1:Sorted (ltk (elt:=option elt'')) l1 -> HdRel (ltk (elt:=option elt'')) (t0, Some e) l1 -> HdRel (ltk (elt:=elt'')) (t0, e) (fold_right_pair (option_cons (A:=elt'')) l1 nil)H1:X.lt t0 t1H0:Sorted (ltk (elt:=option elt'')) l1H2:HdRel (ltk (elt:=option elt'')) (t1, None) l1HdRel (ltk (elt:=option elt'')) (t0, Some e) l1elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''Sorted (ltk (elt:=elt)) nil -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine nil m') = at_least_one (find (elt:=elt) x nil) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''Hm:Sorted (ltk (elt:=elt)) nilm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keyfind (elt:=oee') x (combine nil m') = at_least_one (find (elt:=elt) x nil) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''Hm:Sorted (ltk (elt:=elt)) nilm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keyfind (elt:=oee') x (map (fun e' : elt' => (None, Some e')) m') = match find (elt:=elt') x m' with | Some _ => Some (None, find (elt:=elt') x m') | None => None endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''Hm:Sorted (ltk (elt:=elt)) nilHm':Sorted (ltk (elt:=elt')) nilx:keyfind (elt:=oee') x (map (fun e' : elt' => (None, Some e')) nil) = match find (elt:=elt') x nil with | Some _ => Some (None, find (elt:=elt') x nil) | None => None endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''Hm:Sorted (ltk (elt:=elt)) nila:(X.t * elt')%typem':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) (a :: m')x:keyIHm':Sorted (ltk (elt:=elt')) m' -> find (elt:=oee') x (map (fun e' : elt' => (None, Some e')) m') = match find (elt:=elt') x m' with | Some _ => Some (None, find (elt:=elt') x m') | None => None endfind (elt:=oee') x (map (fun e' : elt' => (None, Some e')) (a :: m')) = match find (elt:=elt') x (a :: m') with | Some _ => Some (None, find (elt:=elt') x (a :: m')) | None => None endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''Hm:Sorted (ltk (elt:=elt)) nila:(X.t * elt')%typem':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) (a :: m')x:keyIHm':Sorted (ltk (elt:=elt')) m' -> find (elt:=oee') x (map (fun e' : elt' => (None, Some e')) m') = match find (elt:=elt') x m' with | Some _ => Some (None, find (elt:=elt') x m') | None => None endfind (elt:=oee') x (map (fun e' : elt' => (None, Some e')) (a :: m')) = match find (elt:=elt') x (a :: m') with | Some _ => Some (None, find (elt:=elt') x (a :: m')) | None => None endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''Hm:Sorted (ltk (elt:=elt)) nilt0:X.te:elt'm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) ((t0, e) :: m')x:keyIHm':Sorted (ltk (elt:=elt')) m' -> find (elt:=oee') x (map (fun e' : elt' => (None, Some e')) m') = match find (elt:=elt') x m' with | Some _ => Some (None, find (elt:=elt') x m') | None => None endfind (elt:=oee') x ((t0, (None, Some e)) :: map (fun e' : elt' => (None, Some e')) m') = match match X.compare x t0 with | LT _ => None | EQ _ => Some e | GT _ => find (elt:=elt') x m' end with | Some _ => Some (None, match X.compare x t0 with | LT _ => None | EQ _ => Some e | GT _ => find (elt:=elt') x m' end) | None => None endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''Hm:Sorted (ltk (elt:=elt)) nilt0:X.te:elt'm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) ((t0, e) :: m')x:keyIHm':Sorted (ltk (elt:=elt')) m' -> find (elt:=oee') x (map (fun e' : elt' => (None, Some e')) m') = match find (elt:=elt') x m' with | Some _ => Some (None, find (elt:=elt') x m') | None => None endl:X.lt t0 xfind (elt:=oee') x (map (fun e' : elt' => (None, Some e')) m') = match find (elt:=elt') x m' with | Some _ => Some (None, find (elt:=elt') x m') | None => None endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')Sorted (ltk (elt:=elt)) (a :: m) -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')(* m' = nil *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')Hm:Sorted (ltk (elt:=elt)) (a :: m)Sorted (ltk (elt:=elt')) nil -> forall x : key, find (elt:=oee') x (combine (a :: m) nil) = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x nil)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x : key, find (elt:=oee') x (combine m m'0) = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')Sorted (ltk (elt:=elt')) (a0 :: m') -> forall x : key, find (elt:=oee') x (combine (a :: m) (a0 :: m')) = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine m m') = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m')Hm:Sorted (ltk (elt:=elt)) ((t0, e) :: m)Hm':Sorted (ltk (elt:=elt')) nilx:keymatch X.compare x t0 with | LT _ => None | EQ _ => Some (Some e, None) | GT _ => find (elt:=oee') x (map (fun e0 : elt => (Some e0, None)) m) end = at_least_one match X.compare x t0 with | LT _ => None | EQ _ => Some e | GT _ => find (elt:=elt) x m end Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x : key, find (elt:=oee') x (combine m m'0) = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')Sorted (ltk (elt:=elt')) (a0 :: m') -> forall x : key, find (elt:=oee') x (combine (a :: m) (a0 :: m')) = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''t0:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine m m') = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m')Hm:Sorted (ltk (elt:=elt)) ((t0, e) :: m)Hm':Sorted (ltk (elt:=elt')) nilx:keyHlt:X.lt t0 xfind (elt:=oee') x (map (fun e0 : elt => (Some e0, None)) m) = at_least_one (find (elt:=elt) x m) Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x : key, find (elt:=oee') x (combine m m'0) = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')Sorted (ltk (elt:=elt')) (a0 :: m') -> forall x : key, find (elt:=oee') x (combine (a :: m) (a0 :: m')) = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''e:eltm:list (X.t * elt)x:keyH:Sorted (ltk (elt:=elt)) mfind (elt:=oee') x (map (fun e0 : elt => (Some e0, None)) m) = at_least_one (find (elt:=elt) x m) Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x : key, find (elt:=oee') x (combine m m'0) = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')Sorted (ltk (elt:=elt')) (a0 :: m') -> forall x : key, find (elt:=oee') x (combine (a :: m) (a0 :: m')) = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''e:elta:(X.t * elt)%typem:list (X.t * elt)x:keyH:Sorted (ltk (elt:=elt)) (a :: m)IHm:Sorted (ltk (elt:=elt)) m -> find (elt:=oee') x (map (fun e0 : elt => (Some e0, None)) m) = at_least_one (find (elt:=elt) x m) Nonefind (elt:=oee') x (let (k, e0) := a in (k, (Some e0, None)) :: map (fun e1 : elt => (Some e1, None)) m) = at_least_one (let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=elt) x m end) Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x : key, find (elt:=oee') x (combine m m'0) = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')Sorted (ltk (elt:=elt')) (a0 :: m') -> forall x : key, find (elt:=oee') x (combine (a :: m) (a0 :: m')) = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''e:elta:(X.t * elt)%typem:list (X.t * elt)x:keyIHm:Sorted (ltk (elt:=elt)) m -> find (elt:=oee') x (map (fun e0 : elt => (Some e0, None)) m) = at_least_one (find (elt:=elt) x m) NoneH0:Sorted (ltk (elt:=elt)) mH1:HdRel (ltk (elt:=elt)) a mfind (elt:=oee') x (let (k, e0) := a in (k, (Some e0, None)) :: map (fun e1 : elt => (Some e1, None)) m) = at_least_one (let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=elt) x m end) Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x : key, find (elt:=oee') x (combine m m'0) = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')Sorted (ltk (elt:=elt')) (a0 :: m') -> forall x : key, find (elt:=oee') x (combine (a :: m) (a0 :: m')) = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''e:eltt0:X.te0:eltm:list (X.t * elt)x:keyIHm:Sorted (ltk (elt:=elt)) m -> find (elt:=oee') x (map (fun e1 : elt => (Some e1, None)) m) = at_least_one (find (elt:=elt) x m) NoneH0:Sorted (ltk (elt:=elt)) mH1:HdRel (ltk (elt:=elt)) (t0, e0) mfind (elt:=oee') x ((t0, (Some e0, None)) :: map (fun e1 : elt => (Some e1, None)) m) = at_least_one match X.compare x t0 with | LT _ => None | EQ _ => Some e0 | GT _ => find (elt:=elt) x m end Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x : key, find (elt:=oee') x (combine m m'0) = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')Sorted (ltk (elt:=elt')) (a0 :: m') -> forall x : key, find (elt:=oee') x (combine (a :: m) (a0 :: m')) = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x (a0 :: m'))(* m' <> nil *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x : key, find (elt:=oee') x (combine m m'0) = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=oee') x (combine (a :: m) m') = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x m')Sorted (ltk (elt:=elt')) (a0 :: m') -> forall x : key, find (elt:=oee') x (combine (a :: m) (a0 :: m')) = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''a:(X.t * elt)%typem:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) (a :: m)a0:(X.t * elt')%typem':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine (a :: m) m') = at_least_one (find (elt:=elt) x0 (a :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) (a0 :: m')x:keyfind (elt:=oee') x (combine (a :: m) (a0 :: m')) = at_least_one (find (elt:=elt) x (a :: m)) (find (elt:=elt') x (a0 :: m'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyfind (elt:=oee') x match X.compare k k' with | LT _ => (k, (Some e, None)) :: combine m ((k', e') :: m') | EQ _ => (k, (Some e, Some e')) :: combine m m' | GT _ => (k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m' end = at_least_one match X.compare x k with | LT _ => None | EQ _ => Some e | GT _ => find (elt:=elt) x m end match X.compare x k' with | LT _ => None | EQ _ => Some e' | GT _ => find (elt:=elt') x m' endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'find (elt:=oee') x match X.compare k k' with | LT _ => (k, (Some e, None)) :: combine m ((k', e') :: m') | EQ _ => (k, (Some e, Some e')) :: combine m m' | GT _ => (k', (None, Some e')) :: (fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m' end = at_least_one match X.compare x k with | LT _ => None | EQ _ => Some e | GT _ => find (elt:=elt) x m end match X.compare x k' with | LT _ => None | EQ _ => Some e' | GT _ => find (elt:=elt') x m' endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k k'l0:X.lt k xH:X.lt x k'find (elt:=oee') x (combine m ((k', e') :: m')) = at_least_one (find (elt:=elt) x m) Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k k'l0:X.lt k xH:X.eq x k'find (elt:=oee') x (combine m ((k', e') :: m')) = at_least_one (find (elt:=elt) x m) (Some e')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k k'l0:X.lt k xH:X.lt k' xfind (elt:=oee') x (combine m ((k', e') :: m')) = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt x kH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = match find (elt:=elt') x m' with | Some _ => Some (None, find (elt:=elt') x m') | None => None endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' ke0:X.eq x kH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = Some (Some e, find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k k'l0:X.lt k xH:X.eq x k'find (elt:=oee') x (combine m ((k', e') :: m')) = at_least_one (find (elt:=elt) x m) (Some e')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k k'l0:X.lt k xH:X.lt k' xfind (elt:=oee') x (combine m ((k', e') :: m')) = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt x kH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = match find (elt:=elt') x m' with | Some _ => Some (None, find (elt:=elt') x m') | None => None endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' ke0:X.eq x kH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = Some (Some e, find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k k'l0:X.lt k xH:X.lt k' xfind (elt:=oee') x (combine m ((k', e') :: m')) = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt x kH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = match find (elt:=elt') x m' with | Some _ => Some (None, find (elt:=elt') x m') | None => None endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' ke0:X.eq x kH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = Some (Some e, find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt x kH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = match find (elt:=elt') x m' with | Some _ => Some (None, find (elt:=elt') x m') | None => None endelt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' ke0:X.eq x kH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = Some (Some e, find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt x kH:X.lt k' xfind (elt:=oee') x (combine ((k, e) :: m) m') = at_least_one None (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' ke0:X.eq x kH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = Some (Some e, find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt x kH:X.lt k' xat_least_one (find (elt:=elt) x ((k, e) :: m)) (find (elt:=elt') x m') = at_least_one None (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' ke0:X.eq x kH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = Some (Some e, find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' ke0:X.eq x kH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e1 : elt => (Some e1, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = Some (Some e, find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' ke0:X.eq x kH:X.lt k' xfind (elt:=oee') x (combine ((k, e) :: m) m') = Some (Some e, find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' ke0:X.eq x kH:X.lt k' xat_least_one (find (elt:=elt) x ((k, e) :: m)) (find (elt:=elt') x m') = Some (Some e, find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xfind (elt:=oee') x ((fix combine_aux (m'0 : t elt') : list (key * oee') := match m'0 with | nil => (k, (Some e, None)) :: map (fun e0 : elt => (Some e0, None)) m | (k'0, e'0) :: l' => match X.compare k k'0 with | LT _ => (k, (Some e, None)) :: combine m m'0 | EQ _ => (k, (Some e, Some e'0)) :: combine m l' | GT _ => (k'0, (None, Some e'0)) :: combine_aux l' end end) m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xfind (elt:=oee') x (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')simpl find; elim_comp; auto. Qed. Definition at_least_one_then_f (o:option elt)(o':option elt') := match o, o' with | None, None => None | _, _ => f o o' end.elt, elt', elt'':Typef:option elt -> option elt' -> option elt''k:X.te:eltm:list (X.t * elt)IHm:Sorted (ltk (elt:=elt)) m -> forall m'0 : list (X.t * elt'), Sorted (ltk (elt:=elt')) m'0 -> forall x0 : key, find (elt:=oee') x0 (combine m m'0) = at_least_one (find (elt:=elt) x0 m) (find (elt:=elt') x0 m'0)Hm:Sorted (ltk (elt:=elt)) ((k, e) :: m)k':X.te':elt'm':list (X.t * elt')IHm':Sorted (ltk (elt:=elt')) m' -> forall x0 : key, find (elt:=oee') x0 (combine ((k, e) :: m) m') = at_least_one (find (elt:=elt) x0 ((k, e) :: m)) (find (elt:=elt') x0 m')Hm':Sorted (ltk (elt:=elt')) ((k', e') :: m')x:keyH1:Sorted (ltk (elt:=elt)) mH2:HdRel (ltk (elt:=elt)) (k, e) mH5:Sorted (ltk (elt:=elt')) m'H6:HdRel (ltk (elt:=elt')) (k', e') m'l:X.lt k' kl0:X.lt k xH:X.lt k' xat_least_one (find (elt:=elt) x ((k, e) :: m)) (find (elt:=elt') x m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=elt'') x (map2 m m') = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, find (elt:=elt'') x (map2 m m') = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keyfind (elt:=elt'') x (map2 m m') = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keyfind (elt:=elt'') x (map2_alt m m') = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keyfind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil) = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keyH:find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil) = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keyH:find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')H2:Sorted (ltk (elt:=oee')) (combine m m')find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) (combine m m')) nil) = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keyH:find (elt:=oee') x (combine m m') = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')H2:Sorted (ltk (elt:=oee')) (combine m m')f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' (combine m m')) nil) = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keym0:t oee'H:find (elt:=oee') x m0 = at_least_one (find (elt:=elt) x m) (find (elt:=elt') x m')H2:Sorted (ltk (elt:=oee')) m0f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keym0:t oee'o:option eltH:find (elt:=oee') x m0 = at_least_one o (find (elt:=elt') x m')H2:Sorted (ltk (elt:=oee')) m0f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keym0:t oee'o:option elto':option elt'H:find (elt:=oee') x m0 = at_least_one o o'H2:Sorted (ltk (elt:=oee')) m0f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keym0:t oee'o:option elto':option elt'H:find (elt:=oee') x m0 = at_least_one o o'H2:Sorted (ltk (elt:=oee')) m0f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keym0:t oee'o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) m0f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keym0:t oee'o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) m0f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''(find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyo:option elto':option elt'H2:Sorted (ltk (elt:=oee')) nilX:option elt''H:None = at_least_one o o'None = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = at_least_one o o'find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = at_least_one o o'find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) ((k, (oo, oo')) :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:match X.compare x k with | LT _ => None | EQ _ => Some (oo, oo') | GT _ => find (elt:=oee') x m0 end = at_least_one o o'find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:match X.compare x k with | LT _ => None | EQ _ => Some (oo, oo') | GT _ => find (elt:=oee') x m0 end = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = None(* x < k *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0e:elt''match X.compare x k with | LT _ => None | EQ _ => Some e | GT _ => find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) end = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0e:elt''H2:X.lt x kNone = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'find (elt:=oee') x m0 = at_least_one o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'find (elt:=oee') x m0 = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'p:oee'H3:find (elt:=oee') x m0 = Some pSome p = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'p:oee'H3:find (elt:=oee') x m0 = Some pltk (x, (oo, oo')) (k, (oo, oo'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'p:oee'H3:find (elt:=oee') x m0 = Some pH4:ltk (x, (oo, oo')) (k, (oo, oo'))Some p = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'p:oee'H3:find (elt:=oee') x m0 = Some pH4:ltk (x, (oo, oo')) (k, (oo, oo'))Some p = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o'p:oee'H3:find (elt:=oee') x m0 = Some pH4:ltk (x, (oo, oo')) (k, (oo, oo'))In x m0elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = None(* x = k *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0at_least_one_then_f o o' = f oo oo'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = f oo oo'find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = f oo oo'find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = f oo oo'find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = f oo oo'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = f oo oo'find (elt:=elt'') x (option_cons k (f oo oo') (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil)) = f oo oo'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0e:elt''H2:at_least_one_then_f o o' = Some ematch X.compare x k with | LT _ => None | EQ _ => Some e | GT _ => find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) end = Some eelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = NoneH4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonefind (elt:=oee') x m0 = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = NoneH4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonep:oee'H3:find (elt:=oee') x m0 = Some pSome p = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = NoneH4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonep:oee'H3:find (elt:=oee') x m0 = Some peqk (k, (oo, oo')) (x, (oo, oo'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = NoneH4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonep:oee'H3:find (elt:=oee') x m0 = Some pH5:eqk (k, (oo, oo')) (x, (oo, oo'))Some p = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = NoneH4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonep:oee'H3:find (elt:=oee') x m0 = Some pH5:eqk (k, (oo, oo')) (x, (oo, oo'))Some p = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H2:at_least_one_then_f o o' = NoneH4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonep:oee'H3:find (elt:=oee') x m0 = Some pH5:eqk (k, (oo, oo')) (x, (oo, oo'))In x m0elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = None(* k < x *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f oo oo') (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil)) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0e:elt''match X.compare x k with | LT _ => None | EQ _ => Some e | GT _ => find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) end = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0e:elt''H2:X.lt k xfind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = at_least_one o o'H0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = at_least_one_then_f o o'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = None(* None -> None *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keya:(X.t * oee')%typem0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) (a :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:(let (k', x0) := a in match X.compare x k' with | LT _ => None | EQ _ => Some x0 | GT _ => find (elt:=oee') x m0 end) = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (let (k, e) := a in (k, f' e) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) ((k, (oo, oo')) :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:match X.compare x k with | LT _ => None | EQ _ => Some (oo, oo') | GT _ => find (elt:=oee') x m0 end = Nonefind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) ((k, f' (oo, oo')) :: map f' m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'H2:Sorted (ltk (elt:=oee')) ((k, (oo, oo')) :: m0)f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:match X.compare x k with | LT _ => None | EQ _ => Some (oo, oo') | GT _ => find (elt:=oee') x m0 end = Nonefind (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)H:match X.compare x k with | LT _ => None | EQ _ => Some (oo, oo') | GT _ => find (elt:=oee') x m0 end = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = None(* x < k *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f oo oo') (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0e:elt''match X.compare x k with | LT _ => None | EQ _ => Some e | GT _ => find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) end = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonefind (elt:=oee') x m0 = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonep:oee'H2:find (elt:=oee') x m0 = Some pSome p = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonep:oee'H2:find (elt:=oee') x m0 = Some pltk (x, (oo, oo')) (k, (oo, oo'))elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonep:oee'H2:find (elt:=oee') x m0 = Some pH3:ltk (x, (oo, oo')) (k, (oo, oo'))Some p = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonep:oee'H2:find (elt:=oee') x m0 = Some pH3:ltk (x, (oo, oo')) (k, (oo, oo'))Some p = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p0 : oee' => f (fst p0) (snd p0):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt x kH:None = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0H4:find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = Nonep:oee'H2:find (elt:=oee') x m0 = Some pH3:ltk (x, (oo, oo')) (k, (oo, oo'))In x m0elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = None(* x = k *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Heq:X.eq x kH:Some (oo, oo') = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = None(* k < x *)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f' (oo, oo')) (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (option_cons k (f oo oo') (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil)) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0e:elt''match X.compare x k with | LT _ => None | EQ _ => Some e | GT _ => find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) end = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0e:elt''H2:X.lt k xfind (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = Noneelt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = Nonedestruct (IHm0 H0) as (_,H4); apply H4; auto. Qed.elt, elt', elt'':Typef:option elt -> option elt' -> option elt''x:keyk:X.too:option eltoo':option elt'm0:list (X.t * oee')o:option elto':option elt'f':=fun p : oee' => f (fst p) (snd p):oee' -> option elt''IHm0:Sorted (ltk (elt:=oee')) m0 -> (find (elt:=oee') x m0 = at_least_one o o' -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = at_least_one_then_f o o') /\ (find (elt:=oee') x m0 = None -> find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map f' m0) nil) = None)Hlt:X.lt k xH:find (elt:=oee') x m0 = NoneH0:Sorted (ltk (elt:=oee')) m0H1:HdRel (ltk (elt:=oee')) (k, (oo, oo')) m0find (elt:=elt'') x (fold_right_pair (option_cons (A:=elt'')) (map (fun p : oee' => f (fst p) (snd p)) m0) nil) = None
Specification of map2
elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, In x m \/ In x m' -> find (elt:=elt'') x (map2 m m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, In x m \/ In x m' -> find (elt:=elt'') x (map2 m m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keyH:In x m \/ In x m'find (elt:=elt'') x (map2 m m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keyH:In x m \/ In x m'at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:eltH:MapsTo x e mat_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt'H:MapsTo x e m'at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:eltH:MapsTo x e mat_least_one_then_f (Some e) (find (elt:=elt') x m') = f (Some e) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt'H:MapsTo x e m'at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt'H:MapsTo x e m'at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')destruct (find x m); simpl; auto. Qed.elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt'H:MapsTo x e m'at_least_one_then_f (find (elt:=elt) x m) (Some e) = f (find (elt:=elt) x m) (Some e)elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, In x (map2 m m') -> In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''forall m : list (X.t * elt), Sorted (ltk (elt:=elt)) m -> forall m' : list (X.t * elt'), Sorted (ltk (elt:=elt')) m' -> forall x : key, In x (map2 m m') -> In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keyH:In x (map2 m m')In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')find (elt:=elt'') x (map2 m m') = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m') -> In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')Some e = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m') -> In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')(forall e0 : elt, find (elt:=elt) x m = Some e0 -> MapsTo x e0 m) -> Some e = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m') -> In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')(forall e0 : elt', find (elt:=elt') x m' = Some e0 -> MapsTo x e0 m') -> (forall e0 : elt, find (elt:=elt) x m = Some e0 -> MapsTo x e0 m) -> Some e = at_least_one_then_f (find (elt:=elt) x m) (find (elt:=elt') x m') -> In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')e0:elte1:elt'H0:forall e2 : elt', Some e1 = Some e2 -> MapsTo x e2 m'H1:forall e2 : elt, Some e0 = Some e2 -> MapsTo x e2 mH2:Some e = f (Some e0) (Some e1)In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')e0:eltH0:forall e1 : elt', None = Some e1 -> MapsTo x e1 m'H1:forall e1 : elt, Some e0 = Some e1 -> MapsTo x e1 mH2:Some e = f (Some e0) NoneIn x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')e0:elt'H0:forall e1 : elt', Some e0 = Some e1 -> MapsTo x e1 m'H1:forall e1 : elt, None = Some e1 -> MapsTo x e1 mH2:Some e = f None (Some e0)In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')H0:forall e0 : elt', None = Some e0 -> MapsTo x e0 m'H1:forall e0 : elt, None = Some e0 -> MapsTo x e0 mH2:Some e = NoneIn x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')e0:eltH0:forall e1 : elt', None = Some e1 -> MapsTo x e1 m'H1:forall e1 : elt, Some e0 = Some e1 -> MapsTo x e1 mH2:Some e = f (Some e0) NoneIn x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')e0:elt'H0:forall e1 : elt', Some e0 = Some e1 -> MapsTo x e1 m'H1:forall e1 : elt, None = Some e1 -> MapsTo x e1 mH2:Some e = f None (Some e0)In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')H0:forall e0 : elt', None = Some e0 -> MapsTo x e0 m'H1:forall e0 : elt, None = Some e0 -> MapsTo x e0 mH2:Some e = NoneIn x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')e0:elt'H0:forall e1 : elt', Some e0 = Some e1 -> MapsTo x e1 m'H1:forall e1 : elt, None = Some e1 -> MapsTo x e1 mH2:Some e = f None (Some e0)In x m \/ In x m'elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')H0:forall e0 : elt', None = Some e0 -> MapsTo x e0 m'H1:forall e0 : elt, None = Some e0 -> MapsTo x e0 mH2:Some e = NoneIn x m \/ In x m'discriminate. Qed. End Elt3. End Raw. Module Make (X: OrderedType) <: S with Module E := X. Module Raw := Raw X. Module E := X. Definition key := E.t. Record slist (elt:Type) := {this :> Raw.t elt; sorted : sort (@Raw.PX.ltk elt) this}. Definition t (elt:Type) : Type := slist elt. Section Elt. Variable elt elt' elt'':Type. Implicit Types m : t elt. Implicit Types x y : key. Implicit Types e : elt. Definition empty : t elt := Build_slist (Raw.empty_sorted elt). Definition is_empty m : bool := Raw.is_empty (this m). Definition add x e m : t elt := Build_slist (Raw.add_sorted (sorted m) x e). Definition find x m : option elt := Raw.find x (this m). Definition remove x m : t elt := Build_slist (Raw.remove_sorted (sorted m) x). Definition mem x m : bool := Raw.mem x (this m). Definition map f m : t elt' := Build_slist (Raw.map_sorted (sorted m) f). Definition mapi (f:key->elt->elt') m : t elt' := Build_slist (Raw.mapi_sorted (sorted m) f). Definition map2 f m (m':t elt') : t elt'' := Build_slist (Raw.map2_sorted f (sorted m) (sorted m')). Definition elements m : list (key*elt) := @Raw.elements elt (this m). Definition cardinal m := length (this m). Definition fold (A:Type)(f:key->elt->A->A) m (i:A) : A := @Raw.fold elt A f (this m) i. Definition equal cmp m m' : bool := @Raw.equal elt cmp (this m) (this m'). Definition MapsTo x e m : Prop := Raw.PX.MapsTo x e (this m). Definition In x m : Prop := Raw.PX.In x (this m). Definition Empty m : Prop := Raw.Empty (this m). Definition Equal m m' := forall y, find y m = find y m'. Definition Equiv (eq_elt:elt->elt->Prop) m m' := (forall k, In k m <-> In k m') /\ (forall k e e', MapsTo k e m -> MapsTo k e' m' -> eq_elt e e'). Definition Equivb cmp m m' : Prop := @Raw.Equivb elt cmp (this m) (this m'). Definition eq_key : (key*elt) -> (key*elt) -> Prop := @Raw.PX.eqk elt. Definition eq_key_elt : (key*elt) -> (key*elt) -> Prop:= @Raw.PX.eqke elt. Definition lt_key : (key*elt) -> (key*elt) -> Prop := @Raw.PX.ltk elt.elt, elt', elt'':Typef:option elt -> option elt' -> option elt''m:list (X.t * elt)Hm:Sorted (ltk (elt:=elt)) mm':list (X.t * elt')Hm':Sorted (ltk (elt:=elt')) m'x:keye:elt''H:MapsTo x e (map2 m m')H0:forall e0 : elt', None = Some e0 -> MapsTo x e0 m'H1:forall e0 : elt, None = Some e0 -> MapsTo x e0 mH2:Some e = NoneIn x m \/ In x m'elt, elt', elt'':Typeforall (m : t elt) (x y : key) (e : elt), E.eq x y -> MapsTo x e m -> MapsTo y e mintros m; exact (@Raw.PX.MapsTo_eq elt (this m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x y : key) (e : elt), E.eq x y -> MapsTo x e m -> MapsTo y e melt, elt', elt'':Typeforall (m : t elt) (x : key), In x m -> mem x m = trueintros m; exact (@Raw.mem_1 elt (this m) (sorted m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x : key), In x m -> mem x m = trueelt, elt', elt'':Typeforall (m : t elt) (x : key), mem x m = true -> In x mintros m; exact (@Raw.mem_2 elt (this m) (sorted m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x : key), mem x m = true -> In x melt, elt', elt'':TypeEmpty emptyexact (@Raw.empty_1 elt). Qed.elt, elt', elt'':TypeEmpty emptyelt, elt', elt'':Typeforall m : t elt, Empty m -> is_empty m = trueintros m; exact (@Raw.is_empty_1 elt (this m)). Qed.elt, elt', elt'':Typeforall m : t elt, Empty m -> is_empty m = trueelt, elt', elt'':Typeforall m : t elt, is_empty m = true -> Empty mintros m; exact (@Raw.is_empty_2 elt (this m)). Qed.elt, elt', elt'':Typeforall m : t elt, is_empty m = true -> Empty melt, elt', elt'':Typeforall (m : t elt) (x y : key) (e : elt), E.eq x y -> MapsTo y e (add x e m)intros m; exact (@Raw.add_1 elt (this m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x y : key) (e : elt), E.eq x y -> MapsTo y e (add x e m)elt, elt', elt'':Typeforall (m : t elt) (x y : key) (e e' : elt), ~ E.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m)intros m; exact (@Raw.add_2 elt (this m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x y : key) (e e' : elt), ~ E.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m)elt, elt', elt'':Typeforall (m : t elt) (x y : key) (e e' : elt), ~ E.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e mintros m; exact (@Raw.add_3 elt (this m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x y : key) (e e' : elt), ~ E.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e melt, elt', elt'':Typeforall (m : t elt) (x y : key), E.eq x y -> ~ In y (remove x m)intros m; exact (@Raw.remove_1 elt (this m) (sorted m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x y : key), E.eq x y -> ~ In y (remove x m)elt, elt', elt'':Typeforall (m : t elt) (x y : key) (e : elt), ~ E.eq x y -> MapsTo y e m -> MapsTo y e (remove x m)intros m; exact (@Raw.remove_2 elt (this m) (sorted m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x y : key) (e : elt), ~ E.eq x y -> MapsTo y e m -> MapsTo y e (remove x m)elt, elt', elt'':Typeforall (m : t elt) (x y : key) (e : elt), MapsTo y e (remove x m) -> MapsTo y e mintros m; exact (@Raw.remove_3 elt (this m) (sorted m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x y : key) (e : elt), MapsTo y e (remove x m) -> MapsTo y e melt, elt', elt'':Typeforall (m : t elt) (x : key) (e : elt), MapsTo x e m -> find x m = Some eintros m; exact (@Raw.find_1 elt (this m) (sorted m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x : key) (e : elt), MapsTo x e m -> find x m = Some eelt, elt', elt'':Typeforall (m : t elt) (x : key) (e : elt), find x m = Some e -> MapsTo x e mintros m; exact (@Raw.find_2 elt (this m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x : key) (e : elt), find x m = Some e -> MapsTo x e melt, elt', elt'':Typeforall (m : t elt) (x : key) (e : elt), MapsTo x e m -> InA eq_key_elt (x, e) (elements m)intros m; exact (@Raw.elements_1 elt (this m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x : key) (e : elt), MapsTo x e m -> InA eq_key_elt (x, e) (elements m)elt, elt', elt'':Typeforall (m : t elt) (x : key) (e : elt), InA eq_key_elt (x, e) (elements m) -> MapsTo x e mintros m; exact (@Raw.elements_2 elt (this m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (x : key) (e : elt), InA eq_key_elt (x, e) (elements m) -> MapsTo x e melt, elt', elt'':Typeforall m : t elt, Sorted lt_key (elements m)intros m; exact (@Raw.elements_3 elt (this m) (sorted m)). Qed.elt, elt', elt'':Typeforall m : t elt, Sorted lt_key (elements m)elt, elt', elt'':Typeforall m : t elt, NoDupA eq_key (elements m)intros m; exact (@Raw.elements_3w elt (this m) (sorted m)). Qed.elt, elt', elt'':Typeforall m : t elt, NoDupA eq_key (elements m)elt, elt', elt'':Typeforall m : t elt, cardinal m = length (elements m)intros; reflexivity. Qed.elt, elt', elt'':Typeforall m : t elt, cardinal m = length (elements m)elt, elt', elt'':Typeforall (m : t elt) (A : Type) (i : A) (f : key -> elt -> A -> A), fold f m i = fold_left (fun (a : A) (p : key * elt) => f (fst p) (snd p) a) (elements m) iintros m; exact (@Raw.fold_1 elt (this m)). Qed.elt, elt', elt'':Typeforall (m : t elt) (A : Type) (i : A) (f : key -> elt -> A -> A), fold f m i = fold_left (fun (a : A) (p : key * elt) => f (fst p) (snd p) a) (elements m) ielt, elt', elt'':Typeforall (m m' : t elt) (cmp : elt -> elt -> bool), Equivb cmp m m' -> equal cmp m m' = trueintros m m'; exact (@Raw.equal_1 elt (this m) (sorted m) (this m') (sorted m')). Qed.elt, elt', elt'':Typeforall (m m' : t elt) (cmp : elt -> elt -> bool), Equivb cmp m m' -> equal cmp m m' = trueelt, elt', elt'':Typeforall (m m' : t elt) (cmp : elt -> elt -> bool), equal cmp m m' = true -> Equivb cmp m m'intros m m'; exact (@Raw.equal_2 elt (this m) (sorted m) (this m') (sorted m')). Qed. End Elt.elt, elt', elt'':Typeforall (m m' : t elt) (cmp : elt -> elt -> bool), equal cmp m m' = true -> Equivb cmp m m'forall (elt elt' : Type) (m : t elt) (x : key) (e : elt) (f : elt -> elt'), MapsTo x e m -> MapsTo x (f e) (map f m)intros elt elt' m; exact (@Raw.map_1 elt elt' (this m)). Qed.forall (elt elt' : Type) (m : t elt) (x : key) (e : elt) (f : elt -> elt'), MapsTo x e m -> MapsTo x (f e) (map f m)forall (elt elt' : Type) (m : t elt) (x : key) (f : elt -> elt'), In (elt:=elt') x (map f m) -> In (elt:=elt) x mintros elt elt' m; exact (@Raw.map_2 elt elt' (this m)). Qed.forall (elt elt' : Type) (m : t elt) (x : key) (f : elt -> elt'), In (elt:=elt') x (map f m) -> In (elt:=elt) x mforall (elt elt' : Type) (m : t elt) (x : key) (e : elt) (f : key -> elt -> elt'), MapsTo x e m -> exists y : E.t, E.eq y x /\ MapsTo x (f y e) (mapi f m)intros elt elt' m; exact (@Raw.mapi_1 elt elt' (this m)). Qed.forall (elt elt' : Type) (m : t elt) (x : key) (e : elt) (f : key -> elt -> elt'), MapsTo x e m -> exists y : E.t, E.eq y x /\ MapsTo x (f y e) (mapi f m)forall (elt elt' : Type) (m : t elt) (x : key) (f : key -> elt -> elt'), In (elt:=elt') x (mapi f m) -> In (elt:=elt) x mintros elt elt' m; exact (@Raw.mapi_2 elt elt' (this m)). Qed.forall (elt elt' : Type) (m : t elt) (x : key) (f : key -> elt -> elt'), In (elt:=elt') x (mapi f m) -> In (elt:=elt) x mforall (elt elt' elt'' : Type) (m : t elt) (m' : t elt') (x : key) (f : option elt -> option elt' -> option elt''), In (elt:=elt) x m \/ In (elt:=elt') x m' -> find (elt:=elt'') x (map2 f m m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')intros elt elt' elt'' m m' x f; exact (@Raw.map2_1 elt elt' elt'' f (this m) (sorted m) (this m') (sorted m') x). Qed.forall (elt elt' elt'' : Type) (m : t elt) (m' : t elt') (x : key) (f : option elt -> option elt' -> option elt''), In (elt:=elt) x m \/ In (elt:=elt') x m' -> find (elt:=elt'') x (map2 f m m') = f (find (elt:=elt) x m) (find (elt:=elt') x m')forall (elt elt' elt'' : Type) (m : t elt) (m' : t elt') (x : key) (f : option elt -> option elt' -> option elt''), In (elt:=elt'') x (map2 f m m') -> In (elt:=elt) x m \/ In (elt:=elt') x m'intros elt elt' elt'' m m' x f; exact (@Raw.map2_2 elt elt' elt'' f (this m) (sorted m) (this m') (sorted m') x). Qed. End Make. Module Make_ord (X: OrderedType)(D : OrderedType) <: Sord with Module Data := D with Module MapS.E := X. Module Data := D. Module MapS := Make(X). Import MapS. Module MD := OrderedTypeFacts(D). Import MD. Definition t := MapS.t D.t. Definition cmp e e' := match D.compare e e' with EQ _ => true | _ => false end. Fixpoint eq_list (m m' : list (X.t * D.t)) : Prop := match m, m' with | nil, nil => True | (x,e)::l, (x',e')::l' => match X.compare x x' with | EQ _ => D.eq e e' /\ eq_list l l' | _ => False end | _, _ => False end. Definition eq m m' := eq_list (this m) (this m'). Fixpoint lt_list (m m' : list (X.t * D.t)) : Prop := match m, m' with | nil, nil => False | nil, _ => True | _, nil => False | (x,e)::l, (x',e')::l' => match X.compare x x' with | LT _ => True | GT _ => False | EQ _ => D.lt e e' \/ (D.eq e e' /\ lt_list l l') end end. Definition lt m m' := lt_list (this m) (this m').forall (elt elt' elt'' : Type) (m : t elt) (m' : t elt') (x : key) (f : option elt -> option elt' -> option elt''), In (elt:=elt'') x (map2 f m m') -> In (elt:=elt) x m \/ In (elt:=elt') x m'forall m m' : slist D.t, eq m m' <-> equal cmp m m' = trueforall m m' : slist D.t, eq m m' <-> equal cmp m m' = trueHl:Sorted (Raw.PX.ltk (elt:=D.t)) nilforall m' : slist D.t, eq {| this := nil; sorted := Hl |} m' <-> equal cmp {| this := nil; sorted := Hl |} m' = truea:(X.t * D.t)%typel:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) (a :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = trueforall m' : slist D.t, eq {| this := a :: l; sorted := Hl |} m' <-> equal cmp {| this := a :: l; sorted := Hl |} m' = trueHl:Sorted (Raw.PX.ltk (elt:=D.t)) nill':Raw.t D.tHl':Sorted (Raw.PX.ltk (elt:=D.t)) l'match l' with | nil => True | _ :: _ => False end <-> equal cmp {| this := nil; sorted := Hl |} {| this := l'; sorted := Hl' |} = truea:(X.t * D.t)%typel:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) (a :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = trueforall m' : slist D.t, eq {| this := a :: l; sorted := Hl |} m' <-> equal cmp {| this := a :: l; sorted := Hl |} m' = truea:(X.t * D.t)%typel:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) (a :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = trueforall m' : slist D.t, eq {| this := a :: l; sorted := Hl |} m' <-> equal cmp {| this := a :: l; sorted := Hl |} m' = truea:(X.t * D.t)%typel:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) (a :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truel':Raw.t D.tHl':Sorted (Raw.PX.ltk (elt:=D.t)) l'eq_list {| this := a :: l; sorted := Hl |} {| this := l'; sorted := Hl' |} <-> equal cmp {| this := a :: l; sorted := Hl |} {| this := l'; sorted := Hl' |} = truea:(X.t * D.t)%typel:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) (a :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = trueHl':Sorted (Raw.PX.ltk (elt:=D.t)) nileq_list {| this := a :: l; sorted := Hl |} {| this := nil; sorted := Hl' |} <-> equal cmp {| this := a :: l; sorted := Hl |} {| this := nil; sorted := Hl' |} = truea:(X.t * D.t)%typel:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) (a :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truep:(X.t * D.t)%typel':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) (p :: l')eq_list {| this := a :: l; sorted := Hl |} {| this := p :: l'; sorted := Hl' |} <-> equal cmp {| this := a :: l; sorted := Hl |} {| this := p :: l'; sorted := Hl' |} = truea:(X.t * D.t)%typel:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) (a :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truep:(X.t * D.t)%typel':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) (p :: l')eq_list {| this := a :: l; sorted := Hl |} {| this := p :: l'; sorted := Hl' |} <-> equal cmp {| this := a :: l; sorted := Hl |} {| this := p :: l'; sorted := Hl' |} = truex:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truep:(X.t * D.t)%typel':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) (p :: l')eq_list {| this := (x, e) :: l; sorted := Hl |} {| this := p :: l'; sorted := Hl' |} <-> equal cmp {| this := (x, e) :: l; sorted := Hl |} {| this := p :: l'; sorted := Hl' |} = truex:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')eq_list {| this := (x, e) :: l; sorted := Hl |} {| this := (x', e') :: l'; sorted := Hl' |} <-> equal cmp {| this := (x, e) :: l; sorted := Hl |} {| this := (x', e') :: l'; sorted := Hl' |} = truex:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')match X.compare x x' with | EQ _ => D.eq e e' /\ eq_list l l' | _ => False end <-> match X.compare x x' with | EQ _ => cmp e e' && Raw.equal cmp l l' | _ => false end = truex:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H0:D.eq e e'H1:eq_list l l'cmp e e' && Raw.equal cmp l l' = truex:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueD.eq e e'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueeq_list l l'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H0:D.eq e e'H1:eq_list l l'match D.compare e e' with | EQ _ => true | _ => false end && Raw.equal cmp l l' = truex:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueD.eq e e'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueeq_list l l'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H0:D.eq e e'H1:eq_list l l'Raw.equal cmp l l' = truex:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueD.eq e e'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueeq_list l l'x:X.te:D.tl:list (X.t * D.t)IHl:forall (Hl : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl |} m' <-> equal cmp {| this := l; sorted := Hl |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H0:D.eq e e'H1:eq_list l l'H:Sorted (Raw.PX.ltk (elt:=D.t)) lH2:HdRel (Raw.PX.ltk (elt:=D.t)) (x, e) lRaw.equal cmp l l' = truex:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueD.eq e e'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueeq_list l l'x:X.te:D.tl:list (X.t * D.t)IHl:forall (Hl : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl |} m' <-> equal cmp {| this := l; sorted := Hl |} m' = truex':X.te':D.tl':list (X.t * D.t)Heq:X.eq x x'H0:D.eq e e'H1:eq_list l l'H:Sorted (Raw.PX.ltk (elt:=D.t)) lH2:HdRel (Raw.PX.ltk (elt:=D.t)) (x, e) lH3:Sorted (Raw.PX.ltk (elt:=D.t)) l'H4:HdRel (Raw.PX.ltk (elt:=D.t)) (x', e') l'Raw.equal cmp l l' = truex:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueD.eq e e'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueeq_list l l'x:X.te:D.tl:list (X.t * D.t)IHl:forall (Hl : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl |} m' <-> equal cmp {| this := l; sorted := Hl |} m' = truex':X.te':D.tl':list (X.t * D.t)Heq:X.eq x x'H0:D.eq e e'H1:eq_list l l'H:Sorted (Raw.PX.ltk (elt:=D.t)) lH2:HdRel (Raw.PX.ltk (elt:=D.t)) (x, e) lH3:Sorted (Raw.PX.ltk (elt:=D.t)) l'H4:HdRel (Raw.PX.ltk (elt:=D.t)) (x', e') l'H5:eq {| this := l; sorted := H |} {| this := l'; sorted := H3 |} -> equal cmp {| this := l; sorted := H |} {| this := l'; sorted := H3 |} = trueH6:equal cmp {| this := l; sorted := H |} {| this := l'; sorted := H3 |} = true -> eq {| this := l; sorted := H |} {| this := l'; sorted := H3 |}Raw.equal cmp l l' = truex:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueD.eq e e'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueeq_list l l'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueD.eq e e'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueeq_list l l'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H0:cmp e e' = trueH1:Raw.equal cmp l l' = trueD.eq e e'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueeq_list l l'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H0:cmp e e' = trueH1:Raw.equal cmp l l' = truematch D.compare e e' with | EQ _ => true | _ => false end = true -> D.eq e e'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueeq_list l l'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H:cmp e e' && Raw.equal cmp l l' = trueeq_list l l'x:X.te:D.tl:list (X.t * D.t)Hl:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: l)IHl:forall (Hl0 : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl0 |} m' <-> equal cmp {| this := l; sorted := Hl0 |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H0:cmp e e' = trueH1:Raw.equal cmp l l' = trueeq_list l l'x:X.te:D.tl:list (X.t * D.t)IHl:forall (Hl : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl |} m' <-> equal cmp {| this := l; sorted := Hl |} m' = truex':X.te':D.tl':list (X.t * D.t)Hl':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: l')Heq:X.eq x x'H0:cmp e e' = trueH1:Raw.equal cmp l l' = trueH:Sorted (Raw.PX.ltk (elt:=D.t)) lH2:HdRel (Raw.PX.ltk (elt:=D.t)) (x, e) leq_list l l'x:X.te:D.tl:list (X.t * D.t)IHl:forall (Hl : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl |} m' <-> equal cmp {| this := l; sorted := Hl |} m' = truex':X.te':D.tl':list (X.t * D.t)Heq:X.eq x x'H0:cmp e e' = trueH1:Raw.equal cmp l l' = trueH:Sorted (Raw.PX.ltk (elt:=D.t)) lH2:HdRel (Raw.PX.ltk (elt:=D.t)) (x, e) lH3:Sorted (Raw.PX.ltk (elt:=D.t)) l'H4:HdRel (Raw.PX.ltk (elt:=D.t)) (x', e') l'eq_list l l'unfold equal, eq in H6; simpl in H6; auto. Qed.x:X.te:D.tl:list (X.t * D.t)IHl:forall (Hl : Sorted (Raw.PX.ltk (elt:=D.t)) l) (m' : slist D.t), eq {| this := l; sorted := Hl |} m' <-> equal cmp {| this := l; sorted := Hl |} m' = truex':X.te':D.tl':list (X.t * D.t)Heq:X.eq x x'H0:cmp e e' = trueH1:Raw.equal cmp l l' = trueH:Sorted (Raw.PX.ltk (elt:=D.t)) lH2:HdRel (Raw.PX.ltk (elt:=D.t)) (x, e) lH3:Sorted (Raw.PX.ltk (elt:=D.t)) l'H4:HdRel (Raw.PX.ltk (elt:=D.t)) (x', e') l'H5:eq {| this := l; sorted := H |} {| this := l'; sorted := H3 |} -> equal cmp {| this := l; sorted := H |} {| this := l'; sorted := H3 |} = trueH6:equal cmp {| this := l; sorted := H |} {| this := l'; sorted := H3 |} = true -> eq {| this := l; sorted := H |} {| this := l'; sorted := H3 |}eq_list l l'forall m m' : MapS.t D.t, Equivb cmp m m' -> eq m m'forall m m' : MapS.t D.t, Equivb cmp m m' -> eq m m'm, m':MapS.t D.tH:Equivb cmp m m'eq m m'm, m':MapS.t D.tH:Equivb cmp m m'(Equivb cmp m m' -> equal cmp m m' = true) -> eq m m'intuition. Qed.m, m':MapS.t D.tH:Equivb cmp m m'eq m m' <-> equal cmp m m' = true -> (Equivb cmp m m' -> equal cmp m m' = true) -> eq m m'forall m m' : slist D.t, eq m m' -> Equivb cmp m m'forall m m' : slist D.t, eq m m' -> Equivb cmp m m'm, m':slist D.tH:eq m m'Equivb cmp m m'm, m':slist D.tH:eq m m'(equal cmp m m' = true -> Equivb cmp m m') -> Equivb cmp m m'intuition. Qed.m, m':slist D.tH:eq m m'eq m m' <-> equal cmp m m' = true -> (equal cmp m m' = true -> Equivb cmp m m') -> Equivb cmp m m'forall m : t, eq m mforall m : t, eq m ma:(X.t * D.t)%typem:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) (a :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}let (x, e) := a in let (x', e') := a in match X.compare x x' with | EQ _ => D.eq e e' /\ eq_list m m | _ => False endt0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}match X.compare t0 t0 with | EQ _ => D.eq t1 t1 /\ eq_list m m | _ => False endt0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Hlt:X.lt t0 t0Falset0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Heq:X.eq t0 t0D.eq t1 t1 /\ eq_list m mt0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Hlt:X.lt t0 t0Falset0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Heq:X.eq t0 t0D.eq t1 t1 /\ eq_list m mt0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Hlt:X.lt t0 t0Falset0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Heq:X.eq t0 t0D.eq t1 t1t0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Heq:X.eq t0 t0eq_list m mt0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Hlt:X.lt t0 t0Falset0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Heq:X.eq t0 t0eq_list m mt0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Hlt:X.lt t0 t0Falset0:X.tt1:D.tm:list (X.t * D.t)IHm:forall Hm : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm |} {| this := m; sorted := Hm |}Heq:X.eq t0 t0H:Sorted (Raw.PX.ltk (elt:=D.t)) mH0:HdRel (Raw.PX.ltk (elt:=D.t)) (t0, t1) meq_list m mt0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Hlt:X.lt t0 t0Falseapply (MapS.Raw.MX.lt_antirefl Hlt); auto. Qed.t0:X.tt1:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((t0, t1) :: m)IHm:forall Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m, eq {| this := m; sorted := Hm0 |} {| this := m; sorted := Hm0 |}Hlt:X.lt t0 t0Falseforall m1 m2 : t, eq m1 m2 -> eq m2 m1forall m1 m2 : t, eq m1 m2 -> eq m2 m1x:X.te:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m)IHm:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m) (m2 : t), eq {| this := m; sorted := Hm0 |} m2 -> eq m2 {| this := m; sorted := Hm0 |}x':X.te':D.tm':list (X.t * D.t)Hm':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m')match X.compare x x' with | EQ _ => D.eq e e' /\ eq_list m m' | _ => False end -> match X.compare x' x with | EQ _ => D.eq e' e /\ eq_list m' m | _ => False endx:X.te:D.tm:list (X.t * D.t)Hm:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m)IHm:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m) (m2 : t), eq {| this := m; sorted := Hm0 |} m2 -> eq m2 {| this := m; sorted := Hm0 |}x':X.te':D.tm':list (X.t * D.t)Hm':Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m')Heq:X.eq x x'H:X.eq x' xH1:D.eq e e'H2:eq_list m m'eq_list m' mapply (IHm H0 (Build_slist H4)); auto. Qed.x:X.te:D.tm:list (X.t * D.t)IHm:forall (Hm : Sorted (Raw.PX.ltk (elt:=D.t)) m) (m2 : t), eq {| this := m; sorted := Hm |} m2 -> eq m2 {| this := m; sorted := Hm |}x':X.te':D.tm':list (X.t * D.t)Heq:X.eq x x'H:X.eq x' xH1:D.eq e e'H2:eq_list m m'H0:Sorted (Raw.PX.ltk (elt:=D.t)) mH3:HdRel (Raw.PX.ltk (elt:=D.t)) (x, e) mH4:Sorted (Raw.PX.ltk (elt:=D.t)) m'H5:HdRel (Raw.PX.ltk (elt:=D.t)) (x', e') m'eq_list m' mforall m1 m2 m3 : t, eq m1 m2 -> eq m2 m3 -> eq m1 m3forall m1 m2 m3 : t, eq m1 m2 -> eq m2 m3 -> eq m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), eq {| this := m1; sorted := Hm0 |} m0 -> eq m0 m4 -> eq {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)match X.compare x x' with | EQ _ => D.eq e e' /\ eq_list m1 m2 | _ => False end -> match X.compare x' x'' with | EQ _ => D.eq e' e'' /\ eq_list m2 m3 | _ => False end -> match X.compare x x'' with | EQ _ => D.eq e e'' /\ eq_list m1 m3 | _ => False endx:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), eq {| this := m1; sorted := Hm0 |} m0 -> eq m0 m4 -> eq {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H2:D.eq e e'H3:eq_list m1 m2H0:D.eq e' e''H4:eq_list m2 m3D.eq e e''x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), eq {| this := m1; sorted := Hm0 |} m0 -> eq m0 m4 -> eq {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H2:D.eq e e'H3:eq_list m1 m2H0:D.eq e' e''H4:eq_list m2 m3eq_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), eq {| this := m1; sorted := Hm0 |} m0 -> eq m0 m4 -> eq {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H2:D.eq e e'H3:eq_list m1 m2H0:D.eq e' e''H4:eq_list m2 m3eq_list m1 m3apply (IHm1 H1 (Build_slist H6) (Build_slist H8)); intuition. Qed.x:X.te:D.tm1:list (X.t * D.t)IHm1:forall (Hm1 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), eq {| this := m1; sorted := Hm1 |} m0 -> eq m0 m4 -> eq {| this := m1; sorted := Hm1 |} m4x':X.te':D.tm2:list (X.t * D.t)x'':X.te'':D.tm3:list (X.t * D.t)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H2:D.eq e e'H3:eq_list m1 m2H0:D.eq e' e''H4:eq_list m2 m3H1:Sorted (Raw.PX.ltk (elt:=D.t)) m1H5:HdRel (Raw.PX.ltk (elt:=D.t)) (x, e) m1H6:Sorted (Raw.PX.ltk (elt:=D.t)) m2H7:HdRel (Raw.PX.ltk (elt:=D.t)) (x', e') m2H8:Sorted (Raw.PX.ltk (elt:=D.t)) m3H9:HdRel (Raw.PX.ltk (elt:=D.t)) (x'', e'') m3eq_list m1 m3forall m1 m2 m3 : t, lt m1 m2 -> lt m2 m3 -> lt m1 m3forall m1 m2 m3 : t, lt m1 m2 -> lt m2 m3 -> lt m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)match X.compare x x' with | LT _ => True | EQ _ => D.lt e e' \/ D.eq e e' /\ lt_list m1 m2 | GT _ => False end -> match X.compare x' x'' with | LT _ => True | EQ _ => D.lt e' e'' \/ D.eq e' e'' /\ lt_list m2 m3 | GT _ => False end -> match X.compare x x'' with | LT _ => True | EQ _ => D.lt e e'' \/ D.eq e e'' /\ lt_list m1 m3 | GT _ => False endx:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H2:D.lt e e'H0:D.lt e' e''D.lt e e'' \/ D.eq e e'' /\ lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H2:D.lt e e'H1:D.eq e' e''H3:lt_list m2 m3D.lt e e'' \/ D.eq e e'' /\ lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H2:D.lt e' e''D.lt e e'' \/ D.eq e e'' /\ lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H1:D.eq e' e''H4:lt_list m2 m3D.lt e e'' \/ D.eq e e'' /\ lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H2:D.lt e e'H1:D.eq e' e''H3:lt_list m2 m3D.lt e e'' \/ D.eq e e'' /\ lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H2:D.lt e' e''D.lt e e'' \/ D.eq e e'' /\ lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H1:D.eq e' e''H4:lt_list m2 m3D.lt e e'' \/ D.eq e e'' /\ lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H2:D.lt e' e''D.lt e e'' \/ D.eq e e'' /\ lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H1:D.eq e' e''H4:lt_list m2 m3D.lt e e'' \/ D.eq e e'' /\ lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H1:D.eq e' e''H4:lt_list m2 m3D.lt e e'' \/ D.eq e e'' /\ lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H1:D.eq e' e''H4:lt_list m2 m3D.eq e e'' /\ lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H1:D.eq e' e''H4:lt_list m2 m3D.eq e e''x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H1:D.eq e' e''H4:lt_list m2 m3lt_list m1 m3x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm0 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm0 |} m4x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)x'':X.te'':D.tm3:list (X.t * D.t)Hm3:Sorted (Raw.PX.ltk (elt:=D.t)) ((x'', e'') :: m3)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H1:D.eq e' e''H4:lt_list m2 m3lt_list m1 m3apply (IHm1 H2 (Build_slist H6) (Build_slist H8)); intuition. Qed.x:X.te:D.tm1:list (X.t * D.t)IHm1:forall (Hm1 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 m4 : t), lt {| this := m1; sorted := Hm1 |} m0 -> lt m0 m4 -> lt {| this := m1; sorted := Hm1 |} m4x':X.te':D.tm2:list (X.t * D.t)x'':X.te'':D.tm3:list (X.t * D.t)Heq:X.eq x x'Heq':X.eq x' x''H:X.eq x x''H0:D.eq e e'H3:lt_list m1 m2H1:D.eq e' e''H4:lt_list m2 m3H2:Sorted (Raw.PX.ltk (elt:=D.t)) m1H5:HdRel (Raw.PX.ltk (elt:=D.t)) (x, e) m1H6:Sorted (Raw.PX.ltk (elt:=D.t)) m2H7:HdRel (Raw.PX.ltk (elt:=D.t)) (x', e') m2H8:Sorted (Raw.PX.ltk (elt:=D.t)) m3H9:HdRel (Raw.PX.ltk (elt:=D.t)) (x'', e'') m3lt_list m1 m3forall m1 m2 : t, lt m1 m2 -> ~ eq m1 m2forall m1 m2 : t, lt m1 m2 -> ~ eq m1 m2x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : t), lt {| this := m1; sorted := Hm0 |} m0 -> ~ eq {| this := m1; sorted := Hm0 |} m0x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)match X.compare x x' with | LT _ => True | EQ _ => D.lt e e' \/ D.eq e e' /\ lt_list m1 m2 | GT _ => False end -> ~ match X.compare x x' with | EQ _ => D.eq e e' /\ eq_list m1 m2 | _ => False endx:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : t), lt {| this := m1; sorted := Hm0 |} m0 -> ~ eq {| this := m1; sorted := Hm0 |} m0x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)Heq:X.eq x x'D.lt e e' \/ D.eq e e' /\ lt_list m1 m2 -> ~ (D.eq e e' /\ eq_list m1 m2)x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : t), lt {| this := m1; sorted := Hm0 |} m0 -> eq {| this := m1; sorted := Hm0 |} m0 -> Falsex':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)Heq:X.eq x x'H1:D.eq e e'H2:eq_list m1 m2H0:D.lt e e'Falsex:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : t), lt {| this := m1; sorted := Hm0 |} m0 -> eq {| this := m1; sorted := Hm0 |} m0 -> Falsex':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)Heq:X.eq x x'H1:D.eq e e'H2:eq_list m1 m2H:D.eq e e'H3:lt_list m1 m2Falsex:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : t), lt {| this := m1; sorted := Hm0 |} m0 -> eq {| this := m1; sorted := Hm0 |} m0 -> Falsex':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)Heq:X.eq x x'H1:D.eq e e'H2:eq_list m1 m2H:D.eq e e'H3:lt_list m1 m2Falseapply (IHm1 H0 (Build_slist H5)); intuition. Qed. Ltac cmp_solve := unfold eq, lt; simpl; try Raw.MX.elim_comp; auto.x:X.te:D.tm1:list (X.t * D.t)IHm1:forall (Hm1 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : t), lt {| this := m1; sorted := Hm1 |} m0 -> eq {| this := m1; sorted := Hm1 |} m0 -> Falsex':X.te':D.tm2:list (X.t * D.t)Heq:X.eq x x'H1:D.eq e e'H2:eq_list m1 m2H:D.eq e e'H3:lt_list m1 m2H0:Sorted (Raw.PX.ltk (elt:=D.t)) m1H4:HdRel (Raw.PX.ltk (elt:=D.t)) (x, e) m1H5:Sorted (Raw.PX.ltk (elt:=D.t)) m2H6:HdRel (Raw.PX.ltk (elt:=D.t)) (x', e') m2Falseforall m1 m2 : slist D.t, Compare lt eq m1 m2forall m1 m2 : slist D.t, Compare lt eq m1 m2a:(X.t * D.t)%typem1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) (a :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : slist D.t), Compare lt eq {| this := m1; sorted := Hm0 |} m0p:(X.t * D.t)%typem2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) (p :: m2)Compare (fun m m' : slist D.t => lt_list m m') (fun m m' : slist D.t => eq_list m m') {| this := a :: m1; sorted := Hm1 |} {| this := p :: m2; sorted := Hm2 |}x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : slist D.t), Compare lt eq {| this := m1; sorted := Hm0 |} m0x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)Compare (fun m m' : slist D.t => lt_list m m') (fun m m' : slist D.t => eq_list m m') {| this := (x, e) :: m1; sorted := Hm1 |} {| this := (x', e') :: m2; sorted := Hm2 |}x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : slist D.t), Compare lt eq {| this := m1; sorted := Hm0 |} m0x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)e0:X.eq x x'Compare (fun m m' : slist D.t => lt_list m m') (fun m m' : slist D.t => eq_list m m') {| this := (x, e) :: m1; sorted := Hm1 |} {| this := (x', e') :: m2; sorted := Hm2 |}x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : slist D.t), Compare lt eq {| this := m1; sorted := Hm0 |} m0x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)e0:X.eq x x'e1:D.eq e e'Compare (fun m m' : slist D.t => lt_list m m') (fun m m' : slist D.t => eq_list m m') {| this := (x, e) :: m1; sorted := Hm1 |} {| this := (x', e') :: m2; sorted := Hm2 |}x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : slist D.t), Compare lt eq {| this := m1; sorted := Hm0 |} m0x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)e0:X.eq x x'e1:D.eq e e'Sorted (Raw.PX.ltk (elt:=D.t)) m1x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : slist D.t), Compare lt eq {| this := m1; sorted := Hm0 |} m0x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)e0:X.eq x x'e1:D.eq e e'Hm11:Sorted (Raw.PX.ltk (elt:=D.t)) m1Compare (fun m m' : slist D.t => lt_list m m') (fun m m' : slist D.t => eq_list m m') {| this := (x, e) :: m1; sorted := Hm1 |} {| this := (x', e') :: m2; sorted := Hm2 |}x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : slist D.t), Compare lt eq {| this := m1; sorted := Hm0 |} m0x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)e0:X.eq x x'e1:D.eq e e'Hm11:Sorted (Raw.PX.ltk (elt:=D.t)) m1Compare (fun m m' : slist D.t => lt_list m m') (fun m m' : slist D.t => eq_list m m') {| this := (x, e) :: m1; sorted := Hm1 |} {| this := (x', e') :: m2; sorted := Hm2 |}x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : slist D.t), Compare lt eq {| this := m1; sorted := Hm0 |} m0x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)e0:X.eq x x'e1:D.eq e e'Hm11:Sorted (Raw.PX.ltk (elt:=D.t)) m1Sorted (Raw.PX.ltk (elt:=D.t)) m2x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : slist D.t), Compare lt eq {| this := m1; sorted := Hm0 |} m0x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)e0:X.eq x x'e1:D.eq e e'Hm11:Sorted (Raw.PX.ltk (elt:=D.t)) m1Hm22:Sorted (Raw.PX.ltk (elt:=D.t)) m2Compare (fun m m' : slist D.t => lt_list m m') (fun m m' : slist D.t => eq_list m m') {| this := (x, e) :: m1; sorted := Hm1 |} {| this := (x', e') :: m2; sorted := Hm2 |}destruct (IHm1 Hm11 (Build_slist Hm22)); [ apply LT | apply EQ | apply GT ]; cmp_solve. Qed. End Make_ord.x:X.te:D.tm1:list (X.t * D.t)Hm1:Sorted (Raw.PX.ltk (elt:=D.t)) ((x, e) :: m1)IHm1:forall (Hm0 : Sorted (Raw.PX.ltk (elt:=D.t)) m1) (m0 : slist D.t), Compare lt eq {| this := m1; sorted := Hm0 |} m0x':X.te':D.tm2:list (X.t * D.t)Hm2:Sorted (Raw.PX.ltk (elt:=D.t)) ((x', e') :: m2)e0:X.eq x x'e1:D.eq e e'Hm11:Sorted (Raw.PX.ltk (elt:=D.t)) m1Hm22:Sorted (Raw.PX.ltk (elt:=D.t)) m2Compare (fun m m' : slist D.t => lt_list m m') (fun m m' : slist D.t => eq_list m m') {| this := (x, e) :: m1; sorted := Hm1 |} {| this := (x', e') :: m2; sorted := Hm2 |}