Documentation

Mathlib.Data.Ordmap.Ordnode

Ordered sets #

This file defines a data structure for ordered sets, supporting a variety of useful operations including insertion and deletion, logarithmic time lookup, set operations, folds, and conversion from lists.

The Ordnode α operations all assume that α has the structure of a total preorder, meaning a operation that is

For example, in order to use this data structure as a map type, one can store pairs (k, v) where (k, v) ≤ (k', v') is defined to mean k ≤ k' (assuming that the key values are linearly ordered).

Two values x,y are equivalent if x ≤ y and y ≤ x. An Ordnode α maintains the invariant that it never stores two equivalent nodes; the insertion operation comes with two variants depending on whether you want to keep the old value or the new value in case you insert a value that is equivalent to one in the set.

The operations in this file are not verified, in the sense that they provide "raw operations" that work for programming purposes but the invariants are not explicitly in the structure. See Ordset for a verified version of this data structure.

Main definitions #

Implementation notes #

Based on weight balanced trees:

Ported from Haskell's Data.Set.

Tags #

ordered map, ordered set, data structure

inductive Ordnode (α : Type u) :

An Ordnode α is a finite set of values, represented as a tree. The operations on this type maintain that the tree is balanced and correctly stores subtree sizes at each level.

Instances For
    Equations
    • Ordnode.instEmptyCollection = { emptyCollection := Ordnode.nil }
    instance Ordnode.instInhabited {α : Type u_1} :
    Equations
    • Ordnode.instInhabited = { default := Ordnode.nil }
    @[inline]

    Internal use only

    The maximal relative difference between the sizes of two trees, it corresponds with the w in Adams' paper.

    According to the Haskell comment, only (delta, ratio) settings of (3, 2) and (4, 2) will work, and the proofs in Ordset.lean assume delta := 3 and ratio := 2.

    Equations
    Instances For
      @[inline]

      Internal use only

      The ratio between an outer and inner sibling of the heavier subtree in an unbalanced setting. It determines whether a double or single rotation should be performed to restore balance. It is corresponds with the inverse of α in Adam's article.

      Equations
      Instances For
        @[inline]
        def Ordnode.singleton {α : Type u_1} (a : α) :

        O(1). Construct a singleton set containing value a.

        singleton 3 = {3}

        Equations
        Instances For
          instance Ordnode.instSingleton {α : Type u_1} :
          Equations
          • Ordnode.instSingleton = { singleton := Ordnode.singleton }
          @[inline]
          def Ordnode.size {α : Type u_1} :
          Ordnode α

          O(1). Get the size of the set.

          size {2, 1, 1, 4} = 3

          Equations
          • x.size = match x with | Ordnode.nil => 0 | Ordnode.node sz l x r => sz
          Instances For
            @[simp]
            theorem Ordnode.size_nil {α : Type u_1} :
            Ordnode.nil.size = 0
            @[simp]
            theorem Ordnode.size_node {α : Type u_1} (sz : ) (l : Ordnode α) (x : α) (r : Ordnode α) :
            (Ordnode.node sz l x r).size = sz
            @[inline]
            def Ordnode.empty {α : Type u_1} :
            Ordnode αBool

            O(1). Is the set empty?

            empty ∅ = tt empty {1, 2, 3} = ff

            Equations
            Instances For
              def Ordnode.dual {α : Type u_1} :
              Ordnode αOrdnode α

              Internal use only, because it violates the BST property on the original order.

              O(n). The dual of a tree is a tree with its left and right sides reversed throughout. The dual of a valid BST is valid under the dual order. This is convenient for exploiting symmetries in the algorithms.

              Equations
              Instances For
                @[reducible, inline]
                def Ordnode.node' {α : Type u_1} (l : Ordnode α) (x : α) (r : Ordnode α) :

                Internal use only

                O(1). Construct a node with the correct size information, without rebalancing.

                Equations
                Instances For
                  def Ordnode.repr {α : Type u_2} [Repr α] (o : Ordnode α) (n : ) :

                  Basic pretty printing for Ordnode α that shows the structure of the tree.

                  repr {3, 1, 2, 4} = ((∅ 1 ∅) 2 ((∅ 3 ∅) 4 ∅))

                  Equations
                  Instances For
                    instance Ordnode.instRepr {α : Type u_2} [Repr α] :
                    Equations
                    • Ordnode.instRepr = { reprPrec := Ordnode.repr }
                    def Ordnode.balanceL {α : Type u_1} (l : Ordnode α) (x : α) (r : Ordnode α) :

                    Internal use only

                    O(1). Rebalance a tree which was previously balanced but has had its left side grow by 1, or its right side shrink by 1.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      def Ordnode.balanceR {α : Type u_1} (l : Ordnode α) (x : α) (r : Ordnode α) :

                      Internal use only

                      O(1). Rebalance a tree which was previously balanced but has had its right side grow by 1, or its left side shrink by 1.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        def Ordnode.balance {α : Type u_1} (l : Ordnode α) (x : α) (r : Ordnode α) :

                        Internal use only

                        O(1). Rebalance a tree which was previously balanced but has had one side change by at most 1.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          def Ordnode.All {α : Type u_1} (P : αProp) :
                          Ordnode αProp

                          O(n). Does every element of the map satisfy property P?

                          All (fun x ↦ x < 5) {1, 2, 3} = True All (fun x ↦ x < 5) {1, 2, 3, 5} = False

                          Equations
                          Instances For
                            instance Ordnode.All.decidable {α : Type u_1} {P : αProp} (t : Ordnode α) [DecidablePred P] :
                            Equations
                            def Ordnode.Any {α : Type u_1} (P : αProp) :
                            Ordnode αProp

                            O(n). Does any element of the map satisfy property P?

                            Any (fun x ↦ x < 2) {1, 2, 3} = True Any (fun x ↦ x < 2) {2, 3, 5} = False

                            Equations
                            Instances For
                              instance Ordnode.Any.decidable {α : Type u_1} {P : αProp} (t : Ordnode α) [DecidablePred P] :
                              Equations
                              def Ordnode.Emem {α : Type u_1} (x : α) :
                              Ordnode αProp

                              O(n). Exact membership in the set. This is useful primarily for stating correctness properties; use for a version that actually uses the BST property of the tree.

                              Emem 2 {1, 2, 3} = true
                              Emem 4 {1, 2, 3} = false 
                              
                              Equations
                              Instances For
                                instance Ordnode.Emem.decidable {α : Type u_1} (x : α) [DecidableEq α] (t : Ordnode α) :
                                Equations
                                def Ordnode.Amem {α : Type u_1} [LE α] (x : α) :
                                Ordnode αProp

                                O(n). Approximate membership in the set, that is, whether some element in the set is equivalent to this one in the preorder. This is useful primarily for stating correctness properties; use for a version that actually uses the BST property of the tree.

                                Amem 2 {1, 2, 3} = true
                                Amem 4 {1, 2, 3} = false
                                

                                To see the difference with Emem, we need a preorder that is not a partial order. For example, suppose we compare pairs of numbers using only their first coordinate. Then: -- Porting note: Verify below example emem (0, 1) {(0, 0), (1, 2)} = false amem (0, 1) {(0, 0), (1, 2)} = true (0, 1) ∈ {(0, 0), (1, 2)} = true

                                The relation is equivalent to Amem as long as the Ordnode is well formed, and should always be used instead of Amem.

                                Equations
                                Instances For
                                  instance Ordnode.Amem.decidable {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) (t : Ordnode α) :
                                  Equations
                                  def Ordnode.findMin' {α : Type u_1} :
                                  Ordnode ααα

                                  O(log n). Return the minimum element of the tree, or the provided default value.

                                  findMin' 37 {1, 2, 3} = 1 findMin' 37 ∅ = 37

                                  Equations
                                  • Ordnode.nil.findMin' x = x
                                  • (Ordnode.node size l x_2 r).findMin' x = l.findMin' x_2
                                  Instances For
                                    def Ordnode.findMin {α : Type u_1} :
                                    Ordnode αOption α

                                    O(log n). Return the minimum element of the tree, if it exists.

                                    findMin {1, 2, 3} = some 1 findMin ∅ = none

                                    Equations
                                    • x.findMin = match x with | Ordnode.nil => none | Ordnode.node size l x r => some (l.findMin' x)
                                    Instances For
                                      def Ordnode.findMax' {α : Type u_1} :
                                      αOrdnode αα

                                      O(log n). Return the maximum element of the tree, or the provided default value.

                                      findMax' 37 {1, 2, 3} = 3 findMax' 37 ∅ = 37

                                      Equations
                                      Instances For
                                        def Ordnode.findMax {α : Type u_1} :
                                        Ordnode αOption α

                                        O(log n). Return the maximum element of the tree, if it exists.

                                        findMax {1, 2, 3} = some 3 findMax ∅ = none

                                        Equations
                                        Instances For
                                          def Ordnode.eraseMin {α : Type u_1} :
                                          Ordnode αOrdnode α

                                          O(log n). Remove the minimum element from the tree, or do nothing if it is already empty.

                                          eraseMin {1, 2, 3} = {2, 3} eraseMin ∅ = ∅

                                          Equations
                                          Instances For
                                            def Ordnode.eraseMax {α : Type u_1} :
                                            Ordnode αOrdnode α

                                            O(log n). Remove the maximum element from the tree, or do nothing if it is already empty.

                                            eraseMax {1, 2, 3} = {1, 2} eraseMax ∅ = ∅

                                            Equations
                                            Instances For
                                              def Ordnode.splitMin' {α : Type u_1} :
                                              Ordnode ααOrdnode αα × Ordnode α

                                              Internal use only, because it requires a balancing constraint on the inputs.

                                              O(log n). Extract and remove the minimum element from a nonempty tree.

                                              Equations
                                              • Ordnode.nil.splitMin' x✝ x = (x✝, x)
                                              • (Ordnode.node size ll lx lr).splitMin' x✝ x = match ll.splitMin' lx lr with | (xm, l') => (xm, l'.balanceR x✝ x)
                                              Instances For
                                                def Ordnode.splitMin {α : Type u_1} :
                                                Ordnode αOption (α × Ordnode α)

                                                O(log n). Extract and remove the minimum element from the tree, if it exists.

                                                split_min {1, 2, 3} = some (1, {2, 3}) split_min ∅ = none

                                                Equations
                                                • x.splitMin = match x with | Ordnode.nil => none | Ordnode.node size l x r => some (l.splitMin' x r)
                                                Instances For
                                                  def Ordnode.splitMax' {α : Type u_1} :
                                                  Ordnode ααOrdnode αOrdnode α × α

                                                  Internal use only, because it requires a balancing constraint on the inputs.

                                                  O(log n). Extract and remove the maximum element from a nonempty tree.

                                                  Equations
                                                  • x✝.splitMax' x Ordnode.nil = (x✝, x)
                                                  • x✝.splitMax' x (Ordnode.node size rl rx rr) = match rl.splitMax' rx rr with | (r', xm) => (x✝.balanceL x r', xm)
                                                  Instances For
                                                    def Ordnode.splitMax {α : Type u_1} :
                                                    Ordnode αOption (Ordnode α × α)

                                                    O(log n). Extract and remove the maximum element from the tree, if it exists.

                                                    split_max {1, 2, 3} = some ({1, 2}, 3) split_max ∅ = none

                                                    Equations
                                                    • x.splitMax = match x with | Ordnode.nil => none | Ordnode.node size x l r => some (x.splitMax' l r)
                                                    Instances For
                                                      def Ordnode.glue {α : Type u_1} :
                                                      Ordnode αOrdnode αOrdnode α

                                                      Internal use only

                                                      O(log(m + n)). Concatenate two trees that are balanced and ordered with respect to each other.

                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Instances For
                                                        def Ordnode.merge {α : Type u_1} (l : Ordnode α) :
                                                        Ordnode αOrdnode α

                                                        O(log(m + n)). Concatenate two trees that are ordered with respect to each other.

                                                        merge {1, 2} {3, 4} = {1, 2, 3, 4} merge {3, 4} {1, 2} = precondition violation

                                                        Equations
                                                        • One or more equations did not get rendered due to their size.
                                                        Instances For
                                                          def Ordnode.insertMax {α : Type u_1} :
                                                          Ordnode ααOrdnode α

                                                          O(log n). Insert an element above all the others, without any comparisons. (Assumes that the element is in fact above all the others).

                                                          insertMax {1, 2} 4 = {1, 2, 4}
                                                          insertMax {1, 2} 0 = precondition violation 
                                                          
                                                          Equations
                                                          Instances For
                                                            def Ordnode.insertMin {α : Type u_1} (x : α) :
                                                            Ordnode αOrdnode α

                                                            O(log n). Insert an element below all the others, without any comparisons. (Assumes that the element is in fact below all the others).

                                                            insertMin {1, 2} 0 = {0, 1, 2}
                                                            insertMin {1, 2} 4 = precondition violation 
                                                            
                                                            Equations
                                                            Instances For
                                                              def Ordnode.filter {α : Type u_1} (p : αProp) [DecidablePred p] :
                                                              Ordnode αOrdnode α

                                                              O(n). Filter the elements of a tree satisfying a predicate.

                                                              filter (fun x ↦ x < 3) {1, 2, 4} = {1, 2} filter (fun x ↦ x > 5) {1, 2, 4} = ∅

                                                              Equations
                                                              Instances For
                                                                def Ordnode.partition {α : Type u_1} (p : αProp) [DecidablePred p] :
                                                                Ordnode αOrdnode α × Ordnode α

                                                                O(n). Split the elements of a tree into those satisfying, and not satisfying, a predicate.

                                                                partition (fun x ↦ x < 3) {1, 2, 4} = ({1, 2}, {3})

                                                                Equations
                                                                • One or more equations did not get rendered due to their size.
                                                                • Ordnode.partition p Ordnode.nil = (Ordnode.nil, Ordnode.nil)
                                                                Instances For
                                                                  def Ordnode.map {α : Type u_1} {β : Type u_2} (f : αβ) :
                                                                  Ordnode αOrdnode β

                                                                  O(n). Map a function across a tree, without changing the structure. Only valid when the function is strictly monotone, i.e. x < y → f x < f y.

                                                                   partition (fun x ↦ x + 2) {1, 2, 4} = {2, 3, 6}
                                                                   partition (fun x : ℕ ↦ x - 2) {1, 2, 4} = precondition violation 
                                                                  
                                                                  Equations
                                                                  Instances For
                                                                    def Ordnode.fold {α : Type u_1} {β : Sort u_2} (z : β) (f : βαββ) :
                                                                    Ordnode αβ

                                                                    O(n). Fold a function across the structure of a tree.

                                                                     fold z f {1, 2, 4} = f (f z 1 z) 2 (f z 4 z)
                                                                    

                                                                    The exact structure of function applications depends on the tree and so is unspecified.

                                                                    Equations
                                                                    Instances For
                                                                      def Ordnode.foldl {α : Type u_1} {β : Sort u_2} (f : βαβ) :
                                                                      βOrdnode αβ

                                                                      O(n). Fold a function from left to right (in increasing order) across the tree.

                                                                      foldl f z {1, 2, 4} = f (f (f z 1) 2) 4

                                                                      Equations
                                                                      Instances For
                                                                        def Ordnode.foldr {α : Type u_1} {β : Sort u_2} (f : αββ) :
                                                                        Ordnode αββ

                                                                        O(n). Fold a function from right to left (in decreasing order) across the tree.

                                                                        foldr f {1, 2, 4} z = f 1 (f 2 (f 4 z))

                                                                        Equations
                                                                        Instances For
                                                                          def Ordnode.toList {α : Type u_1} (t : Ordnode α) :
                                                                          List α

                                                                          O(n). Build a list of elements in ascending order from the tree.

                                                                          toList {1, 2, 4} = [1, 2, 4] toList {2, 1, 1, 4} = [1, 2, 4]

                                                                          Equations
                                                                          Instances For
                                                                            def Ordnode.toRevList {α : Type u_1} (t : Ordnode α) :
                                                                            List α

                                                                            O(n). Build a list of elements in descending order from the tree.

                                                                            toRevList {1, 2, 4} = [4, 2, 1] toRevList {2, 1, 1, 4} = [4, 2, 1]

                                                                            Equations
                                                                            Instances For
                                                                              instance Ordnode.instToString {α : Type u_1} [ToString α] :
                                                                              Equations
                                                                              • Ordnode.instToString = { toString := fun (t : Ordnode α) => "{" ++ ", ".intercalate (List.map toString t.toList) ++ "}" }
                                                                              Equations
                                                                              def Ordnode.Equiv {α : Type u_1} (t₁ : Ordnode α) (t₂ : Ordnode α) :

                                                                              O(n). True if the trees have the same elements, ignoring structural differences.

                                                                              Equiv {1, 2, 4} {2, 1, 1, 4} = true Equiv {1, 2, 4} {1, 2, 3} = false

                                                                              Equations
                                                                              • t₁.Equiv t₂ = (t₁.size = t₂.size t₁.toList = t₂.toList)
                                                                              Instances For
                                                                                Equations
                                                                                • x✝.instDecidableRelEquivOfDecidableEq x = And.decidable
                                                                                def Ordnode.powerset {α : Type u_1} (t : Ordnode α) :

                                                                                O(2^n). Constructs the powerset of a given set, that is, the set of all subsets.

                                                                                powerset {1, 2, 3} = {∅, {1}, {2}, {3}, {1,2}, {1,3}, {2,3}, {1,2,3}}

                                                                                Equations
                                                                                • One or more equations did not get rendered due to their size.
                                                                                Instances For
                                                                                  def Ordnode.prod {α : Type u_1} {β : Type u_2} (t₁ : Ordnode α) (t₂ : Ordnode β) :
                                                                                  Ordnode (α × β)

                                                                                  O(m * n). The cartesian product of two sets: (a, b) ∈ s.prod t iff a ∈ s and b ∈ t.

                                                                                  prod {1, 2} {2, 3} = {(1, 2), (1, 3), (2, 2), (2, 3)}

                                                                                  Equations
                                                                                  Instances For
                                                                                    def Ordnode.copair {α : Type u_1} {β : Type u_2} (t₁ : Ordnode α) (t₂ : Ordnode β) :
                                                                                    Ordnode (α β)

                                                                                    O(m + n). Build a set on the disjoint union by combining sets on the factors. Or.inl a ∈ s.copair t iff a ∈ s, and Or.inr b ∈ s.copair t iff b ∈ t.

                                                                                    copair {1, 2} {2, 3} = {inl 1, inl 2, inr 2, inr 3} 
                                                                                    
                                                                                    Equations
                                                                                    Instances For
                                                                                      def Ordnode.pmap {α : Type u_1} {P : αProp} {β : Type u_2} (f : (a : α) → P aβ) (t : Ordnode α) :
                                                                                      Ordnode.All P tOrdnode β

                                                                                      O(n). Map a partial function across a set. The result depends on a proof that the function is defined on all members of the set.

                                                                                      pmap (fin.mk : ∀ n, n < 4 → fin 4) {1, 2} H = {(1 : fin 4), (2 : fin 4)} 
                                                                                      
                                                                                      Equations
                                                                                      Instances For
                                                                                        def Ordnode.attach' {α : Type u_1} {P : αProp} (t : Ordnode α) :
                                                                                        Ordnode.All P tOrdnode { a : α // P a }

                                                                                        O(n). "Attach" the information that every element of t satisfies property P to these elements inside the set, producing a set in the subtype.

                                                                                        attach' (fun x ↦ x < 4) {1, 2} H = ({1, 2} : Ordnode {x // x<4}) 
                                                                                        
                                                                                        Equations
                                                                                        Instances For
                                                                                          def Ordnode.nth {α : Type u_1} :
                                                                                          Ordnode αOption α

                                                                                          O(log n). Get the ith element of the set, by its index from left to right.

                                                                                          nth {a, b, c, d} 2 = some c nth {a, b, c, d} 5 = none

                                                                                          Equations
                                                                                          • Ordnode.nil.nth x = none
                                                                                          • (Ordnode.node size l x_2 r).nth x = match x.psub' l.size with | none => l.nth x | some 0 => some x_2 | some j.succ => r.nth j
                                                                                          Instances For
                                                                                            def Ordnode.removeNth {α : Type u_1} :
                                                                                            Ordnode αOrdnode α

                                                                                            O(log n). Remove the ith element of the set, by its index from left to right.

                                                                                            remove_nth {a, b, c, d} 2 = {a, b, d} remove_nth {a, b, c, d} 5 = {a, b, c, d}

                                                                                            Equations
                                                                                            • Ordnode.nil.removeNth x = Ordnode.nil
                                                                                            • (Ordnode.node size l x_2 r).removeNth x = match x.psub' l.size with | none => (l.removeNth x).balanceR x_2 r | some 0 => l.glue r | some j.succ => l.balanceL x_2 (r.removeNth j)
                                                                                            Instances For
                                                                                              def Ordnode.takeAux {α : Type u_1} :
                                                                                              Ordnode αOrdnode α

                                                                                              Auxiliary definition for take. (Can also be used in lieu of take if you know the index is within the range of the data structure.)

                                                                                              takeAux {a, b, c, d} 2 = {a, b}
                                                                                              takeAux {a, b, c, d} 5 = {a, b, c, d} 
                                                                                              
                                                                                              Equations
                                                                                              • One or more equations did not get rendered due to their size.
                                                                                              • Ordnode.nil.takeAux x = Ordnode.nil
                                                                                              Instances For
                                                                                                def Ordnode.take {α : Type u_1} (i : ) (t : Ordnode α) :

                                                                                                O(log n). Get the first i elements of the set, counted from the left.

                                                                                                take 2 {a, b, c, d} = {a, b} take 5 {a, b, c, d} = {a, b, c, d}

                                                                                                Equations
                                                                                                Instances For
                                                                                                  def Ordnode.dropAux {α : Type u_1} :
                                                                                                  Ordnode αOrdnode α

                                                                                                  Auxiliary definition for drop. (Can also be used in lieu of drop if you know the index is within the range of the data structure.)

                                                                                                  drop_aux {a, b, c, d} 2 = {c, d}
                                                                                                  drop_aux {a, b, c, d} 5 = ∅ 
                                                                                                  
                                                                                                  Equations
                                                                                                  • One or more equations did not get rendered due to their size.
                                                                                                  • Ordnode.nil.dropAux x = Ordnode.nil
                                                                                                  Instances For
                                                                                                    def Ordnode.drop {α : Type u_1} (i : ) (t : Ordnode α) :

                                                                                                    O(log n). Remove the first i elements of the set, counted from the left.

                                                                                                    drop 2 {a, b, c, d} = {c, d} drop 5 {a, b, c, d} = ∅

                                                                                                    Equations
                                                                                                    Instances For
                                                                                                      def Ordnode.splitAtAux {α : Type u_1} :
                                                                                                      Ordnode αOrdnode α × Ordnode α

                                                                                                      Auxiliary definition for splitAt. (Can also be used in lieu of splitAt if you know the index is within the range of the data structure.)

                                                                                                      splitAtAux {a, b, c, d} 2 = ({a, b}, {c, d})
                                                                                                      splitAtAux {a, b, c, d} 5 = ({a, b, c, d}, ∅) 
                                                                                                      
                                                                                                      Equations
                                                                                                      • One or more equations did not get rendered due to their size.
                                                                                                      • Ordnode.nil.splitAtAux x = (Ordnode.nil, Ordnode.nil)
                                                                                                      Instances For
                                                                                                        def Ordnode.splitAt {α : Type u_1} (i : ) (t : Ordnode α) :

                                                                                                        O(log n). Split a set at the ith element, getting the first i and everything else.

                                                                                                        splitAt 2 {a, b, c, d} = ({a, b}, {c, d}) splitAt 5 {a, b, c, d} = ({a, b, c, d}, ∅)

                                                                                                        Equations
                                                                                                        Instances For
                                                                                                          def Ordnode.takeWhile {α : Type u_1} (p : αProp) [DecidablePred p] :
                                                                                                          Ordnode αOrdnode α

                                                                                                          O(log n). Get an initial segment of the set that satisfies the predicate p. p is required to be antitone, that is, x < y → p y → p x.

                                                                                                          takeWhile (fun x ↦ x < 4) {1, 2, 3, 4, 5} = {1, 2, 3}
                                                                                                          takeWhile (fun x ↦ x > 4) {1, 2, 3, 4, 5} = precondition violation 
                                                                                                          
                                                                                                          Equations
                                                                                                          Instances For
                                                                                                            def Ordnode.dropWhile {α : Type u_1} (p : αProp) [DecidablePred p] :
                                                                                                            Ordnode αOrdnode α

                                                                                                            O(log n). Remove an initial segment of the set that satisfies the predicate p. p is required to be antitone, that is, x < y → p y → p x.

                                                                                                            dropWhile (fun x ↦ x < 4) {1, 2, 3, 4, 5} = {4, 5}
                                                                                                            dropWhile (fun x ↦ x > 4) {1, 2, 3, 4, 5} = precondition violation 
                                                                                                            
                                                                                                            Equations
                                                                                                            Instances For
                                                                                                              def Ordnode.span {α : Type u_1} (p : αProp) [DecidablePred p] :
                                                                                                              Ordnode αOrdnode α × Ordnode α

                                                                                                              O(log n). Split the set into those satisfying and not satisfying the predicate p. p is required to be antitone, that is, x < y → p y → p x.

                                                                                                              span (fun x ↦ x < 4) {1, 2, 3, 4, 5} = ({1, 2, 3}, {4, 5})
                                                                                                              span (fun x ↦ x > 4) {1, 2, 3, 4, 5} = precondition violation 
                                                                                                              
                                                                                                              Equations
                                                                                                              • One or more equations did not get rendered due to their size.
                                                                                                              • Ordnode.span p Ordnode.nil = (Ordnode.nil, Ordnode.nil)
                                                                                                              Instances For
                                                                                                                def Ordnode.ofAscListAux₁ {α : Type u_1} (l : List α) :
                                                                                                                Ordnode α × { l' : List α // l'.length l.length }

                                                                                                                Auxiliary definition for ofAscList.

                                                                                                                Note: This function is defined by well founded recursion, so it will probably not compute in the kernel, meaning that you probably can't prove things like ofAscList [1, 2, 3] = {1, 2, 3} by rfl. This implementation is optimized for VM evaluation.

                                                                                                                Equations
                                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                                • Ordnode.ofAscListAux₁ [] = fun (x : ) => (Ordnode.nil, [], )
                                                                                                                Instances For
                                                                                                                  def Ordnode.ofAscListAux₂ {α : Type u_1} :
                                                                                                                  List αOrdnode αOrdnode α

                                                                                                                  Auxiliary definition for ofAscList.

                                                                                                                  Equations
                                                                                                                  Instances For
                                                                                                                    def Ordnode.ofAscList {α : Type u_1} :
                                                                                                                    List αOrdnode α

                                                                                                                    O(n). Build a set from a list which is already sorted. Performs no comparisons.

                                                                                                                    ofAscList [1, 2, 3] = {1, 2, 3} ofAscList [3, 2, 1] = precondition violation

                                                                                                                    Equations
                                                                                                                    Instances For
                                                                                                                      def Ordnode.mem {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                      Ordnode αBool

                                                                                                                      O(log n). Does the set (approximately) contain the element x? That is, is there an element that is equivalent to x in the order?

                                                                                                                      1 ∈ {1, 2, 3} = true
                                                                                                                      4 ∈ {1, 2, 3} = false
                                                                                                                      

                                                                                                                      Using a preorder on ℕ × ℕ that only compares the first coordinate:

                                                                                                                      (1, 1) ∈ {(0, 1), (1, 2)} = true
                                                                                                                      (3, 1) ∈ {(0, 1), (1, 2)} = false 
                                                                                                                      
                                                                                                                      Equations
                                                                                                                      Instances For
                                                                                                                        def Ordnode.find {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                        Ordnode αOption α

                                                                                                                        O(log n). Retrieve an element in the set that is equivalent to x in the order, if it exists.

                                                                                                                        find 1 {1, 2, 3} = some 1
                                                                                                                        find 4 {1, 2, 3} = none
                                                                                                                        

                                                                                                                        Using a preorder on ℕ × ℕ that only compares the first coordinate:

                                                                                                                        find (1, 1) {(0, 1), (1, 2)} = some (1, 2)
                                                                                                                        find (3, 1) {(0, 1), (1, 2)} = none 
                                                                                                                        
                                                                                                                        Equations
                                                                                                                        Instances For
                                                                                                                          instance Ordnode.instMembership {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] :
                                                                                                                          Equations
                                                                                                                          instance Ordnode.mem.decidable {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) (t : Ordnode α) :
                                                                                                                          Equations
                                                                                                                          def Ordnode.insertWith {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (f : αα) (x : α) :
                                                                                                                          Ordnode αOrdnode α

                                                                                                                          O(log n). Insert an element into the set, preserving balance and the BST property. If an equivalent element is already in the set, the function f is used to generate the element to insert (being passed the current value in the set).

                                                                                                                          insertWith f 0 {1, 2, 3} = {0, 1, 2, 3}
                                                                                                                          insertWith f 1 {1, 2, 3} = {f 1, 2, 3}
                                                                                                                          

                                                                                                                          Using a preorder on ℕ × ℕ that only compares the first coordinate:

                                                                                                                          insertWith f (1, 1) {(0, 1), (1, 2)} = {(0, 1), f (1, 2)}
                                                                                                                          insertWith f (3, 1) {(0, 1), (1, 2)} = {(0, 1), (1, 2), (3, 1)} 
                                                                                                                          
                                                                                                                          Equations
                                                                                                                          Instances For
                                                                                                                            def Ordnode.adjustWith {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (f : αα) (x : α) :
                                                                                                                            Ordnode αOrdnode α

                                                                                                                            O(log n). Modify an element in the set with the given function, doing nothing if the key is not found. Note that the element returned by f must be equivalent to x.

                                                                                                                            adjustWith f 0 {1, 2, 3} = {1, 2, 3}
                                                                                                                            adjustWith f 1 {1, 2, 3} = {f 1, 2, 3}
                                                                                                                            

                                                                                                                            Using a preorder on ℕ × ℕ that only compares the first coordinate:

                                                                                                                            adjustWith f (1, 1) {(0, 1), (1, 2)} = {(0, 1), f (1, 2)}
                                                                                                                            adjustWith f (3, 1) {(0, 1), (1, 2)} = {(0, 1), (1, 2)} 
                                                                                                                            
                                                                                                                            Equations
                                                                                                                            • One or more equations did not get rendered due to their size.
                                                                                                                            • Ordnode.adjustWith f x Ordnode.nil = Ordnode.nil
                                                                                                                            Instances For
                                                                                                                              def Ordnode.updateWith {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (f : αOption α) (x : α) :
                                                                                                                              Ordnode αOrdnode α

                                                                                                                              O(log n). Modify an element in the set with the given function, doing nothing if the key is not found. Note that the element returned by f must be equivalent to x.

                                                                                                                              updateWith f 0 {1, 2, 3} = {1, 2, 3}
                                                                                                                              updateWith f 1 {1, 2, 3} = {2, 3}     if f 1 = none
                                                                                                                                                        = {a, 2, 3}  if f 1 = some a 
                                                                                                                              
                                                                                                                              Equations
                                                                                                                              • One or more equations did not get rendered due to their size.
                                                                                                                              • Ordnode.updateWith f x Ordnode.nil = Ordnode.nil
                                                                                                                              Instances For
                                                                                                                                def Ordnode.alter {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (f : Option αOption α) (x : α) :
                                                                                                                                Ordnode αOrdnode α

                                                                                                                                O(log n). Modify an element in the set with the given function, doing nothing if the key is not found. Note that the element returned by f must be equivalent to x.

                                                                                                                                alter f 0 {1, 2, 3} = {1, 2, 3}     if f none = none
                                                                                                                                                    = {a, 1, 2, 3}  if f none = some a
                                                                                                                                alter f 1 {1, 2, 3} = {2, 3}     if f 1 = none
                                                                                                                                                    = {a, 2, 3}  if f 1 = some a 
                                                                                                                                
                                                                                                                                Equations
                                                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                                                • Ordnode.alter f x Ordnode.nil = Option.recOn (f none) Ordnode.nil Ordnode.singleton
                                                                                                                                Instances For
                                                                                                                                  def Ordnode.insert {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                  Ordnode αOrdnode α

                                                                                                                                  O(log n). Insert an element into the set, preserving balance and the BST property. If an equivalent element is already in the set, this replaces it.

                                                                                                                                  insert 1 {1, 2, 3} = {1, 2, 3}
                                                                                                                                  insert 4 {1, 2, 3} = {1, 2, 3, 4}
                                                                                                                                  

                                                                                                                                  Using a preorder on ℕ × ℕ that only compares the first coordinate:

                                                                                                                                  insert (1, 1) {(0, 1), (1, 2)} = {(0, 1), (1, 1)}
                                                                                                                                  insert (3, 1) {(0, 1), (1, 2)} = {(0, 1), (1, 2), (3, 1)} 
                                                                                                                                  
                                                                                                                                  Equations
                                                                                                                                  Instances For
                                                                                                                                    instance Ordnode.instInsert {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] :
                                                                                                                                    Insert α (Ordnode α)
                                                                                                                                    Equations
                                                                                                                                    • Ordnode.instInsert = { insert := Ordnode.insert }
                                                                                                                                    def Ordnode.insert' {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                    Ordnode αOrdnode α

                                                                                                                                    O(log n). Insert an element into the set, preserving balance and the BST property. If an equivalent element is already in the set, the set is returned as is.

                                                                                                                                    insert' 1 {1, 2, 3} = {1, 2, 3}
                                                                                                                                    insert' 4 {1, 2, 3} = {1, 2, 3, 4}
                                                                                                                                    

                                                                                                                                    Using a preorder on ℕ × ℕ that only compares the first coordinate:

                                                                                                                                    insert' (1, 1) {(0, 1), (1, 2)} = {(0, 1), (1, 2)}
                                                                                                                                    insert' (3, 1) {(0, 1), (1, 2)} = {(0, 1), (1, 2), (3, 1)} 
                                                                                                                                    
                                                                                                                                    Equations
                                                                                                                                    Instances For
                                                                                                                                      def Ordnode.split {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                      Ordnode αOrdnode α × Ordnode α

                                                                                                                                      O(log n). Split the tree into those smaller than x and those greater than it. If an element equivalent to x is in the set, it is discarded.

                                                                                                                                      split 2 {1, 2, 4} = ({1}, {4})
                                                                                                                                      split 3 {1, 2, 4} = ({1, 2}, {4})
                                                                                                                                      split 4 {1, 2, 4} = ({1, 2}, ∅)
                                                                                                                                      

                                                                                                                                      Using a preorder on ℕ × ℕ that only compares the first coordinate:

                                                                                                                                      split (1, 1) {(0, 1), (1, 2)} = ({(0, 1)}, ∅)
                                                                                                                                      split (3, 1) {(0, 1), (1, 2)} = ({(0, 1), (1, 2)}, ∅) 
                                                                                                                                      
                                                                                                                                      Equations
                                                                                                                                      • One or more equations did not get rendered due to their size.
                                                                                                                                      • Ordnode.split x Ordnode.nil = (Ordnode.nil, Ordnode.nil)
                                                                                                                                      Instances For
                                                                                                                                        def Ordnode.split3 {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                        Ordnode αOrdnode α × Option α × Ordnode α

                                                                                                                                        O(log n). Split the tree into those smaller than x and those greater than it, plus an element equivalent to x, if it exists.

                                                                                                                                        split3 2 {1, 2, 4} = ({1}, some 2, {4})
                                                                                                                                        split3 3 {1, 2, 4} = ({1, 2}, none, {4})
                                                                                                                                        split3 4 {1, 2, 4} = ({1, 2}, some 4, ∅)
                                                                                                                                        

                                                                                                                                        Using a preorder on ℕ × ℕ that only compares the first coordinate:

                                                                                                                                        split3 (1, 1) {(0, 1), (1, 2)} = ({(0, 1)}, some (1, 2), ∅)
                                                                                                                                        split3 (3, 1) {(0, 1), (1, 2)} = ({(0, 1), (1, 2)}, none, ∅) 
                                                                                                                                        
                                                                                                                                        Equations
                                                                                                                                        • One or more equations did not get rendered due to their size.
                                                                                                                                        • Ordnode.split3 x Ordnode.nil = (Ordnode.nil, none, Ordnode.nil)
                                                                                                                                        Instances For
                                                                                                                                          def Ordnode.erase {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                          Ordnode αOrdnode α

                                                                                                                                          O(log n). Remove an element from the set equivalent to x. Does nothing if there is no such element.

                                                                                                                                          erase 1 {1, 2, 3} = {2, 3}
                                                                                                                                          erase 4 {1, 2, 3} = {1, 2, 3}
                                                                                                                                          

                                                                                                                                          Using a preorder on ℕ × ℕ that only compares the first coordinate:

                                                                                                                                          erase (1, 1) {(0, 1), (1, 2)} = {(0, 1)}
                                                                                                                                          erase (3, 1) {(0, 1), (1, 2)} = {(0, 1), (1, 2)} 
                                                                                                                                          
                                                                                                                                          Equations
                                                                                                                                          • One or more equations did not get rendered due to their size.
                                                                                                                                          • Ordnode.erase x Ordnode.nil = Ordnode.nil
                                                                                                                                          Instances For
                                                                                                                                            def Ordnode.findLtAux {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                            Ordnode ααα

                                                                                                                                            Auxiliary definition for findLt.

                                                                                                                                            Equations
                                                                                                                                            Instances For
                                                                                                                                              def Ordnode.findLt {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                              Ordnode αOption α

                                                                                                                                              O(log n). Get the largest element in the tree that is < x.

                                                                                                                                              findLt 2 {1, 2, 4} = some 1 findLt 3 {1, 2, 4} = some 2 findLt 0 {1, 2, 4} = none

                                                                                                                                              Equations
                                                                                                                                              Instances For
                                                                                                                                                def Ordnode.findGtAux {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                                Ordnode ααα

                                                                                                                                                Auxiliary definition for findGt.

                                                                                                                                                Equations
                                                                                                                                                Instances For
                                                                                                                                                  def Ordnode.findGt {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                                  Ordnode αOption α

                                                                                                                                                  O(log n). Get the smallest element in the tree that is > x.

                                                                                                                                                  findGt 2 {1, 2, 4} = some 4 findGt 3 {1, 2, 4} = some 4 findGt 4 {1, 2, 4} = none

                                                                                                                                                  Equations
                                                                                                                                                  Instances For
                                                                                                                                                    def Ordnode.findLeAux {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                                    Ordnode ααα

                                                                                                                                                    Auxiliary definition for findLe.

                                                                                                                                                    Equations
                                                                                                                                                    • One or more equations did not get rendered due to their size.
                                                                                                                                                    • Ordnode.findLeAux x✝ Ordnode.nil x = x
                                                                                                                                                    Instances For
                                                                                                                                                      def Ordnode.findLe {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                                      Ordnode αOption α

                                                                                                                                                      O(log n). Get the largest element in the tree that is ≤ x.

                                                                                                                                                      findLe 2 {1, 2, 4} = some 2 findLe 3 {1, 2, 4} = some 2 findLe 0 {1, 2, 4} = none

                                                                                                                                                      Equations
                                                                                                                                                      Instances For
                                                                                                                                                        def Ordnode.findGeAux {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                                        Ordnode ααα

                                                                                                                                                        Auxiliary definition for findGe.

                                                                                                                                                        Equations
                                                                                                                                                        • One or more equations did not get rendered due to their size.
                                                                                                                                                        • Ordnode.findGeAux x✝ Ordnode.nil x = x
                                                                                                                                                        Instances For
                                                                                                                                                          def Ordnode.findGe {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                                          Ordnode αOption α

                                                                                                                                                          O(log n). Get the smallest element in the tree that is ≥ x.

                                                                                                                                                          findGe 2 {1, 2, 4} = some 2 findGe 3 {1, 2, 4} = some 4 findGe 5 {1, 2, 4} = none

                                                                                                                                                          Equations
                                                                                                                                                          Instances For
                                                                                                                                                            def Ordnode.findIndexAux {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) :
                                                                                                                                                            Ordnode αOption

                                                                                                                                                            Auxiliary definition for findIndex.

                                                                                                                                                            Equations
                                                                                                                                                            • One or more equations did not get rendered due to their size.
                                                                                                                                                            • Ordnode.findIndexAux x✝ Ordnode.nil x = none
                                                                                                                                                            Instances For
                                                                                                                                                              def Ordnode.findIndex {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (x : α) (t : Ordnode α) :

                                                                                                                                                              O(log n). Get the index, counting from the left, of an element equivalent to x if it exists.

                                                                                                                                                              findIndex 2 {1, 2, 4} = some 1
                                                                                                                                                              findIndex 4 {1, 2, 4} = some 2
                                                                                                                                                              findIndex 5 {1, 2, 4} = none 
                                                                                                                                                              
                                                                                                                                                              Equations
                                                                                                                                                              Instances For
                                                                                                                                                                def Ordnode.isSubsetAux {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] :
                                                                                                                                                                Ordnode αOrdnode αBool

                                                                                                                                                                Auxiliary definition for isSubset.

                                                                                                                                                                Equations
                                                                                                                                                                • Ordnode.nil.isSubsetAux x = true
                                                                                                                                                                • x.isSubsetAux Ordnode.nil = false
                                                                                                                                                                • (Ordnode.node size l x_2 r).isSubsetAux x = match Ordnode.split3 x_2 x with | (lt, found, gt) => found.isSome && l.isSubsetAux lt && r.isSubsetAux gt
                                                                                                                                                                Instances For
                                                                                                                                                                  def Ordnode.isSubset {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (t₁ : Ordnode α) (t₂ : Ordnode α) :

                                                                                                                                                                  O(m + n). Is every element of t₁ equivalent to some element of t₂?

                                                                                                                                                                  is_subset {1, 4} {1, 2, 4} = tt is_subset {1, 3} {1, 2, 4} = ff

                                                                                                                                                                  Equations
                                                                                                                                                                  • t₁.isSubset t₂ = (decide (t₁.size t₂.size) && t₁.isSubsetAux t₂)
                                                                                                                                                                  Instances For
                                                                                                                                                                    def Ordnode.disjoint {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] :
                                                                                                                                                                    Ordnode αOrdnode αBool

                                                                                                                                                                    O(m + n). Is every element of t₁ not equivalent to any element of t₂?

                                                                                                                                                                    disjoint {1, 3} {2, 4} = tt disjoint {1, 2} {2, 4} = ff

                                                                                                                                                                    Equations
                                                                                                                                                                    Instances For
                                                                                                                                                                      def Ordnode.union {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] :
                                                                                                                                                                      Ordnode αOrdnode αOrdnode α

                                                                                                                                                                      O(m * log(|m ∪ n| + 1)), m ≤ n. The union of two sets, preferring members of t₁ over those of t₂ when equivalent elements are encountered.

                                                                                                                                                                      union {1, 2} {2, 3} = {1, 2, 3} union {1, 3} {2} = {1, 2, 3}

                                                                                                                                                                      Using a preorder on ℕ × ℕ that only compares the first coordinate:

                                                                                                                                                                      union {(1, 1)} {(0, 1), (1, 2)} = {(0, 1), (1, 1)}

                                                                                                                                                                      Equations
                                                                                                                                                                      • One or more equations did not get rendered due to their size.
                                                                                                                                                                      • x.union Ordnode.nil = x
                                                                                                                                                                      • Ordnode.nil.union x = x
                                                                                                                                                                      Instances For
                                                                                                                                                                        def Ordnode.diff {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] :
                                                                                                                                                                        Ordnode αOrdnode αOrdnode α

                                                                                                                                                                        O(m * log(|m ∪ n| + 1)), m ≤ n. Difference of two sets.

                                                                                                                                                                        diff {1, 2} {2, 3} = {1} diff {1, 2, 3} {2} = {1, 3}

                                                                                                                                                                        Equations
                                                                                                                                                                        • One or more equations did not get rendered due to their size.
                                                                                                                                                                        • x.diff Ordnode.nil = x
                                                                                                                                                                        Instances For
                                                                                                                                                                          def Ordnode.inter {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] :
                                                                                                                                                                          Ordnode αOrdnode αOrdnode α

                                                                                                                                                                          O(m * log(|m ∪ n| + 1)), m ≤ n. Intersection of two sets, preferring members of t₁ over those of t₂ when equivalent elements are encountered.

                                                                                                                                                                          inter {1, 2} {2, 3} = {2}
                                                                                                                                                                          inter {1, 3} {2} = ∅ 
                                                                                                                                                                          
                                                                                                                                                                          Equations
                                                                                                                                                                          • One or more equations did not get rendered due to their size.
                                                                                                                                                                          • Ordnode.nil.inter x = Ordnode.nil
                                                                                                                                                                          Instances For
                                                                                                                                                                            def Ordnode.ofList {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] (l : List α) :

                                                                                                                                                                            O(n * log n). Build a set from a list, preferring elements that appear earlier in the list in the case of equivalent elements.

                                                                                                                                                                            ofList [1, 2, 3] = {1, 2, 3}
                                                                                                                                                                            ofList [2, 1, 1, 3] = {1, 2, 3}
                                                                                                                                                                            

                                                                                                                                                                            Using a preorder on ℕ × ℕ that only compares the first coordinate:

                                                                                                                                                                            ofList [(1, 1), (0, 1), (1, 2)] = {(0, 1), (1, 1)} 
                                                                                                                                                                            
                                                                                                                                                                            Equations
                                                                                                                                                                            Instances For
                                                                                                                                                                              def Ordnode.ofList' {α : Type u_1} [LE α] [DecidableRel fun (x x_1 : α) => x x_1] :
                                                                                                                                                                              List αOrdnode α

                                                                                                                                                                              O(n * log n). Adaptively chooses between the linear and log-linear algorithm depending on whether the input list is already sorted.

                                                                                                                                                                              ofList' [1, 2, 3] = {1, 2, 3} ofList' [2, 1, 1, 3] = {1, 2, 3}

                                                                                                                                                                              Equations
                                                                                                                                                                              Instances For
                                                                                                                                                                                def Ordnode.image {α : Type u_2} {β : Type u_3} [LE β] [DecidableRel fun (x x_1 : β) => x x_1] (f : αβ) (t : Ordnode α) :

                                                                                                                                                                                O(n * log n). Map a function on a set. Unlike map this has no requirements on f, and the resulting set may be smaller than the input if f is noninjective. Equivalent elements are selected with a preference for smaller source elements.

                                                                                                                                                                                image (fun x ↦ x + 2) {1, 2, 4} = {3, 4, 6}
                                                                                                                                                                                image (fun x : ℕ ↦ x - 2) {1, 2, 4} = {0, 2} 
                                                                                                                                                                                
                                                                                                                                                                                Equations
                                                                                                                                                                                Instances For