- Node {α : Type u} {β : Type v} : Option β → (RBNode α fun (x : α) => PrefixTreeNode α β) → PrefixTreeNode α β
Instances For
instance
Lean.instInhabitedPrefixTreeNode
{α : Type u_1}
{β : Type u_2}
:
Inhabited (PrefixTreeNode α β)
Equations
Instances For
@[specialize #[]]
def
Lean.PrefixTreeNode.insert
{α : Type u_1}
{β : Type u_2}
(t : PrefixTreeNode α β)
(cmp : α → α → Ordering)
(k : List α)
(val : β)
:
PrefixTreeNode α β
Equations
- t.insert cmp k val = Lean.PrefixTreeNode.insert.loop cmp val t k
Instances For
partial def
Lean.PrefixTreeNode.insert.insertEmpty
{α : Type u_1}
{β : Type u_2}
(val : β)
(k : List α)
:
PrefixTreeNode α β
partial def
Lean.PrefixTreeNode.insert.loop
{α : Type u_1}
{β : Type u_2}
(cmp : α → α → Ordering)
(val : β)
:
PrefixTreeNode α β → List α → PrefixTreeNode α β
@[specialize #[]]
def
Lean.PrefixTreeNode.find?
{α : Type u_1}
{β : Type u_2}
(t : PrefixTreeNode α β)
(cmp : α → α → Ordering)
(k : List α)
:
Option β
Equations
- t.find? cmp k = Lean.PrefixTreeNode.find?.loop cmp t k
Instances For
partial def
Lean.PrefixTreeNode.find?.loop
{α : Type u_1}
{β : Type u_2}
(cmp : α → α → Ordering)
:
PrefixTreeNode α β → List α → Option β
@[specialize #[]]
def
Lean.PrefixTreeNode.findLongestPrefix?
{α : Type u_1}
{β : Type u_2}
(t : PrefixTreeNode α β)
(cmp : α → α → Ordering)
(k : List α)
:
Option β
Returns the the value of the longest key in t
that is a prefix of k
, if any.
Equations
- t.findLongestPrefix? cmp k = Lean.PrefixTreeNode.findLongestPrefix?.loop cmp none t k
Instances For
partial def
Lean.PrefixTreeNode.findLongestPrefix?.loop
{α : Type u_1}
{β : Type u_2}
(cmp : α → α → Ordering)
(acc? : Option β)
:
PrefixTreeNode α β → List α → Option β
@[specialize #[]]
def
Lean.PrefixTreeNode.foldMatchingM
{m : Type u_1 → Type u_2}
{α : Type u_3}
{β : Type u_4}
{σ : Type u_1}
[Monad m]
(t : PrefixTreeNode α β)
(cmp : α → α → Ordering)
(k : List α)
(init : σ)
(f : β → σ → m σ)
:
m σ
Equations
- t.foldMatchingM cmp k init f = Lean.PrefixTreeNode.foldMatchingM.find cmp init f k t init
Instances For
partial def
Lean.PrefixTreeNode.foldMatchingM.fold
{m : Type u_1 → Type u_2}
{α : Type u_3}
{β : Type u_4}
{σ : Type u_1}
[Monad m]
(f : β → σ → m σ)
:
PrefixTreeNode α β → σ → m σ
inductive
Lean.PrefixTreeNode.WellFormed
{α : Type u_1}
(cmp : α → α → Ordering)
{β : Type u_2}
:
PrefixTreeNode α β → Prop
- emptyWff {α : Type u_1} {cmp : α → α → Ordering} {x✝ : Type u_2} : WellFormed cmp empty
- insertWff {α : Type u_1} {cmp : α → α → Ordering} {x✝ : Type u_2} {t : PrefixTreeNode α x✝} {k : List α} {val : x✝} : WellFormed cmp t → WellFormed cmp (t.insert cmp k val)
Instances For
Equations
- Lean.PrefixTree α β cmp = { t : Lean.PrefixTreeNode α β // Lean.PrefixTreeNode.WellFormed cmp t }
Instances For
Equations
Instances For
instance
Lean.instInhabitedPrefixTree
{α : Type u_1}
{β : Type u_2}
{p : α → α → Ordering}
:
Inhabited (PrefixTree α β p)
Equations
- Lean.instInhabitedPrefixTree = { default := Lean.PrefixTree.empty }
instance
Lean.instEmptyCollectionPrefixTree
{α : Type u_1}
{β : Type u_2}
{p : α → α → Ordering}
:
EmptyCollection (PrefixTree α β p)
Equations
- Lean.instEmptyCollectionPrefixTree = { emptyCollection := Lean.PrefixTree.empty }
@[inline]
def
Lean.PrefixTree.insert
{α : Type u_1}
{β : Type u_2}
{p : α → α → Ordering}
(t : PrefixTree α β p)
(k : List α)
(v : β)
:
PrefixTree α β p
Instances For
@[inline]
def
Lean.PrefixTree.findLongestPrefix?
{α : Type u_1}
{β : Type u_2}
{p : α → α → Ordering}
(t : PrefixTree α β p)
(k : List α)
:
Option β
Returns the the value of the longest key in t
that is a prefix of k
, if any.
Equations
- t.findLongestPrefix? k = t.val.findLongestPrefix? p k
Instances For
@[inline]
def
Lean.PrefixTree.foldMatchingM
{m : Type u_1 → Type u_2}
{α : Type u_3}
{β : Type u_4}
{p : α → α → Ordering}
{σ : Type u_1}
[Monad m]
(t : PrefixTree α β p)
(k : List α)
(init : σ)
(f : β → σ → m σ)
:
m σ
Equations
- t.foldMatchingM k init f = t.val.foldMatchingM p k init f
Instances For
@[inline]
def
Lean.PrefixTree.foldM
{m : Type u_1 → Type u_2}
{α : Type u_3}
{β : Type u_4}
{p : α → α → Ordering}
{σ : Type u_1}
[Monad m]
(t : PrefixTree α β p)
(init : σ)
(f : β → σ → m σ)
:
m σ
Equations
- t.foldM init f = t.foldMatchingM [] init f
Instances For
@[inline]
def
Lean.PrefixTree.forMatchingM
{m : Type → Type u_1}
{α : Type u_2}
{β : Type u_3}
{p : α → α → Ordering}
[Monad m]
(t : PrefixTree α β p)
(k : List α)
(f : β → m Unit)
:
m Unit
Equations
- t.forMatchingM k f = t.val.foldMatchingM p k () fun (b : β) (x : Unit) => f b