The following functions can't be defined at Init.Data.List.Basic
, because they depend on Init.Util
,
and Init.Util
depends on Init.Data.List.Basic
.
Alternative getters #
get! #
Returns the i
-th element in the list (zero-based).
If the index is out of bounds (i ≥ as.length
), this function panics when executed, and returns
default
. See get?
and getD
for safer alternatives.
Equations
Instances For
getLast! #
Returns the last element in the list.
If the list is empty, this function panics when executed, and returns default
.
See getLast
and getLastD
for safer alternatives.
Equations
- [].getLast! = panicWithPosWithDecl "Init.Data.List.BasicAux" "List.getLast!" 44 13 "empty list"
- (a :: as).getLast! = (a :: as).getLast ⋯
Instances For
Head and tail #
head! #
Returns the first element in the list.
If the list is empty, this function panics when executed, and returns default
.
See head
and headD
for safer alternatives.
Equations
- [].head! = panicWithPosWithDecl "Init.Data.List.BasicAux" "List.head!" 58 12 "empty list"
- (a :: as).head! = a
Instances For
tail! #
Drops the first element of the list.
If the list is empty, this function panics when executed, and returns the empty list.
See tail
and tailD
for safer alternatives.
Instances For
partitionM #
Monadic generalization of List.partition
.
This uses Array.toList
and which isn't imported by Init.Data.List.Basic
or Init.Data.List.Control
.
def posOrNeg (x : Int) : Except String Bool :=
if x > 0 then pure true
else if x < 0 then pure false
else throw "Zero is not positive or negative"
partitionM posOrNeg [-1, 2, 3] = Except.ok ([2, 3], [-1])
partitionM posOrNeg [0, 2, 3] = Except.error "Zero is not positive or negative"
Instances For
Auxiliary for partitionM
:
partitionM.go p l acc₁ acc₂
returns (acc₁.toList ++ left, acc₂.toList ++ right)
if partitionM p l
returns (left, right)
.
Instances For
partitionMap #
Given a function f : α → β ⊕ γ
, partitionMap f l
maps the list by f
whilst partitioning the result into a pair of lists, List β × List γ
,
partitioning the .inl _
into the left list, and the .inr _
into the right List.
partitionMap (id : Nat ⊕ Nat → Nat ⊕ Nat) [inl 0, inr 1, inl 2] = ([0, 2], [1])
Equations
- List.partitionMap f l = List.partitionMap.go f l #[] #[]
Instances For
Auxiliary for partitionMap
:
partitionMap.go f l acc₁ acc₂ = (acc₁.toList ++ left, acc₂.toList ++ right)
if partitionMap f l = (left, right)
.
Equations
- List.partitionMap.go f [] x✝ x = (x✝.toList, x.toList)
- List.partitionMap.go f (x_3 :: xs) x✝ x = match f x_3 with | Sum.inl a => List.partitionMap.go f xs (x✝.push a) x | Sum.inr b => List.partitionMap.go f xs x✝ (x.push b)
Instances For
mapMono #
This is a performance optimization for List.mapM
that avoids allocating a new list when the result of each f a
is a pointer equal value a
.
For verification purposes, List.mapMono = List.map
.
Monomorphic List.mapM
. The internal implementation uses pointer equality, and does not allocate a new list
if the result of each f a
is a pointer equal value a
.
Equations
Instances For
This tactic, added to the decreasing_trivial
toolbox, proves that
sizeOf a < sizeOf as
when a ∈ as
, which is useful for well founded recursions
over a nested inductive like inductive T | mk : List T → T
.
Equations
- List.tacticSizeOf_list_dec = Lean.ParserDescr.node `List.tacticSizeOf_list_dec 1024 (Lean.ParserDescr.nonReservedSymbol "sizeOf_list_dec" false)