(* -*- coding: utf-8 -*- *) (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* t | Node : t -> A -> t -> t . Fixpoint find (vm : t) (p:positive) {struct vm} : A := match vm with | Empty => default | Leaf i => i | Node l e r => match p with | xH => e | xO p => find l p | xI p => find r p end end. Fixpoint singleton (x:positive) (v : A) : t := match x with | xH => Leaf v | xO p => Node (singleton p v) default Empty | xI p => Node Empty default (singleton p v) end. Fixpoint vm_add (x: positive) (v : A) (m : t) {struct m} : t := match m with | Empty => singleton x v | Leaf vl => match x with | xH => Leaf v | xO p => Node (singleton p v) vl Empty | xI p => Node Empty vl (singleton p v) end | Node l o r => match x with | xH => Node l v r | xI p => Node l o (vm_add p v r) | xO p => Node (vm_add p v l) o r end end. End MakeVarMap.