Module docstring
{}
{}
List.asString
definition
(s : List Char) : String
/--
Creates a string that contains the characters in a list, in order.
Examples:
* `['L', '∃', '∀', 'N'].asString = "L∃∀N"`
* `[].asString = ""`
* `['a', 'a', 'a'].asString = "aaa"`
-/
def List.asString (s : List Char) : String :=
⟨s⟩
String.instOfNatPos
instance
: OfNat String.Pos (nat_lit 0)
instance : OfNat String.Pos (nat_lit 0) where
ofNat := {}
String.instLT
instance
: LT String
instance : LT String :=
⟨fun s₁ s₂ => s₁.data < s₂.data⟩
String.decidableLT
instance
(s₁ s₂ : @& String) : Decidable (s₁ < s₂)
@[extern "lean_string_dec_lt"]
instance decidableLT (s₁ s₂ : @& String) : Decidable (s₁ < s₂) :=
List.decidableLT s₁.data s₂.data
String.decLt
abbrev
@[deprecated decidableLT (since := "2024-12-13")] abbrev decLt := @decidableLT
String.le
definition
(a b : String) : Prop
String.instLE
instance
: LE String
instance : LE String :=
⟨String.le⟩
String.decLE
instance
(s₁ s₂ : String) : Decidable (s₁ ≤ s₂)
String.length
definition
: (@& String) → Nat
String.push
definition
: String → Char → String
/--
Adds a character to the end of a string.
The internal implementation uses dynamic arrays and will perform destructive updates
if the string is not shared.
Examples:
* `"abc".push 'd' = "abcd"`
* `"".push 'a' = "a"`
-/
@[extern "lean_string_push"]
def push : String → Char → String
| ⟨s⟩, c => ⟨s ++ [c]⟩
String.append
definition
: String → (@& String) → String
/--
Appends two strings. Usually accessed via the `++` operator.
The internal implementation will perform destructive updates if the string is not shared.
Examples:
* `"abc".append "def" = "abcdef"`
* `"abc" ++ "def" = "abcdef"`
* `"" ++ "" = ""`
-/
@[extern "lean_string_append"]
def append : String → (@& String) → String
| ⟨a⟩, ⟨b⟩ => ⟨a ++ b⟩
String.toList
definition
(s : String) : List Char
/--
Converts a string to a list of characters.
Even though the logical model of strings is as a structure that wraps a list of characters, this
operation takes time and space linear in the length of the string. At runtime, strings are
represented as dynamic arrays of bytes.
Examples:
* `"abc".toList = ['a', 'b', 'c']`
* `"".toList = []`
* `"\n".toList = ['\n']`
-/
def toList (s : String) : List Char :=
s.data
String.Pos.isValid
definition
(s : @& String) (p : @& Pos) : Bool
/--
Returns `true` if `p` is a valid UTF-8 position in the string `s`.
This means that `p ≤ s.endPos` and `p` lies on a UTF-8 character boundary. At runtime, this
operation takes constant time.
Examples:
* `String.Pos.isValid "abc" ⟨0⟩ = true`
* `String.Pos.isValid "abc" ⟨1⟩ = true`
* `String.Pos.isValid "abc" ⟨3⟩ = true`
* `String.Pos.isValid "abc" ⟨4⟩ = false`
* `String.Pos.isValid "𝒫(A)" ⟨0⟩ = true`
* `String.Pos.isValid "𝒫(A)" ⟨1⟩ = false`
* `String.Pos.isValid "𝒫(A)" ⟨2⟩ = false`
* `String.Pos.isValid "𝒫(A)" ⟨3⟩ = false`
* `String.Pos.isValid "𝒫(A)" ⟨4⟩ = true`
-/
@[extern "lean_string_is_valid_pos"]
def Pos.isValid (s : @&String) (p : @& Pos) : Bool :=
go s.data 0
where
go : List Char → Pos → Bool
| [], i => i = p
| c::cs, i => if i = p then true else go cs (i + c)
String.utf8GetAux
definition
: List Char → Pos → Pos → Char
String.get
definition
(s : @& String) (p : @& Pos) : Char
/--
Returns the character at position `p` of a string. If `p` is not a valid position, returns the
fallback value `(default : Char)`, which is `'A'`, but does not panic.
This function is overridden with an efficient implementation in runtime code. See
`String.utf8GetAux` for the reference implementation.
Examples:
* `"abc".get ⟨1⟩ = 'b'`
* `"abc".get ⟨3⟩ = (default : Char)` because byte `3` is at the end of the string.
* `"L∃∀N".get ⟨2⟩ = (default : Char)` because byte `2` is in the middle of `'∃'`.
-/
@[extern "lean_string_utf8_get"]
def get (s : @& String) (p : @& Pos) : Char :=
match s with
| ⟨s⟩ => utf8GetAux s 0 p
String.utf8GetAux?
definition
: List Char → Pos → Pos → Option Char
String.get?
definition
: (@& String) → (@& Pos) → Option Char
/--
Returns the character at position `p` of a string. If `p` is not a valid position, returns `none`.
This function is overridden with an efficient implementation in runtime code. See
`String.utf8GetAux?` for the reference implementation.
Examples:
* `"abc".get? ⟨1⟩ = some 'b'`
* `"abc".get? ⟨3⟩ = none`
* `"L∃∀N".get? ⟨1⟩ = some '∃'`
* `"L∃∀N".get? ⟨2⟩ = none`
-/
@[extern "lean_string_utf8_get_opt"]
def get? : (@& String) → (@& Pos) → Option Char
| ⟨s⟩, p => utf8GetAux? s 0 p
String.get!
definition
(s : @& String) (p : @& Pos) : Char
/--
Returns the character at position `p` of a string. Panics if `p` is not a valid position.
See `String.get?` for a safer alternative.
This function is overridden with an efficient implementation in runtime code. See
`String.utf8GetAux` for the reference implementation.
Examples
* `"abc".get! ⟨1⟩ = 'b'`
-/
@[extern "lean_string_utf8_get_bang"]
def get! (s : @& String) (p : @& Pos) : Char :=
match s with
| ⟨s⟩ => utf8GetAux s 0 p
String.utf8SetAux
definition
(c' : Char) : List Char → Pos → Pos → List Char
String.set
definition
: String → (@& Pos) → Char → String
/--
Replaces the character at a specified position in a string with a new character. If the position is
invalid, the string is returned unchanged.
If both the replacement character and the replaced character are 7-bit ASCII characters and the
string is not shared, then it is updated in-place and not copied.
Examples:
* `"abc".set ⟨1⟩ 'B' = "aBc"`
* `"abc".set ⟨3⟩ 'D' = "abc"`
* `"L∃∀N".set ⟨4⟩ 'X' = "L∃XN"`
* `"L∃∀N".set ⟨2⟩ 'X' = "L∃∀N"` because `'∃'` is a multi-byte character, so the byte index `2` is an
invalid position.
-/
@[extern "lean_string_utf8_set"]
def set : String → (@& Pos) → Char → String
| ⟨s⟩, i, c => ⟨utf8SetAux c s 0 i⟩
String.modify
definition
(s : String) (i : Pos) (f : Char → Char) : String
/--
Replaces the character at position `p` in the string `s` with the result of applying `f` to that
character. If `p` is an invalid position, the string is returned unchanged.
If both the replacement character and the replaced character are 7-bit ASCII characters and the
string is not shared, then it is updated in-place and not copied.
Examples:
* `abc.modify ⟨1⟩ Char.toUpper = "aBc"`
* `abc.modify ⟨3⟩ Char.toUpper = "abc"`
-/
def modify (s : String) (i : Pos) (f : Char → Char) : String :=
s.set i <| f <| s.get i
String.next
definition
(s : @& String) (p : @& Pos) : Pos
/--
Returns the next position in a string after position `p`. The result is unspecified if `p` is not a
valid position or if `p = s.endPos`.
A run-time bounds check is performed to determine whether `p` is at the end of the string. If a
bounds check has already been performed, use `String.next'` to avoid a repeated check.
Some examples where the result is unspecified:
* `"abc".next ⟨3⟩`, since `3 = "abc".endPos`
* `"L∃∀N".next ⟨2⟩`, since `2` points into the middle of a multi-byte UTF-8 character
Examples:
* `"abc".get ("abc".next 0) = 'b'`
* `"L∃∀N".get (0 |> "L∃∀N".next |> "L∃∀N".next) = '∀'`
-/
@[extern "lean_string_utf8_next"]
def next (s : @& String) (p : @& Pos) : Pos :=
let c := get s p
p + c
String.utf8PrevAux
definition
: List Char → Pos → Pos → Pos
String.prev
definition
: (@& String) → (@& Pos) → Pos
/--
Returns the position in a string before a specified position, `p`. If `p = ⟨0⟩`, returns `0`. If `p`
is not a valid position, the result is unspecified.
For example, `"L∃∀N".prev ⟨3⟩` is unspecified, since byte 3 occurs in the middle of the multi-byte
character `'∃'`.
Examples:
* `"abc".get ("abc".endPos |> "abc".prev) = 'c'`
* `"L∃∀N".get ("L∃∀N".endPos |> "L∃∀N".prev |> "L∃∀N".prev |> "L∃∀N".prev) = '∃'`
-/
@[extern "lean_string_utf8_prev"]
def prev : (@& String) → (@& Pos) → Pos
| ⟨s⟩, p => if p = 0 then 0 else utf8PrevAux s 0 p
String.front
definition
(s : String) : Char
String.back
definition
(s : String) : Char
String.atEnd
definition
: (@& String) → (@& Pos) → Bool
/--
Returns `true` if a specified byte position is greater than or equal to the position which points to
the end of a string. Otherwise, returns `false`.
Examples:
* `(0 |> "abc".next |> "abc".next |> "abc".atEnd) = false`
* `(0 |> "abc".next |> "abc".next |> "abc".next |> "abc".next |> "abc".atEnd) = true`
* `(0 |> "L∃∀N".next |> "L∃∀N".next |> "L∃∀N".next |> "L∃∀N".atEnd) = false`
* `(0 |> "L∃∀N".next |> "L∃∀N".next |> "L∃∀N".next |> "L∃∀N".next |> "L∃∀N".atEnd) = true`
* `"abc".atEnd ⟨4⟩ = true`
* `"L∃∀N".atEnd ⟨7⟩ = false`
* `"L∃∀N".atEnd ⟨8⟩ = true`
-/
@[extern "lean_string_utf8_at_end"]
def atEnd : (@& String) → (@& Pos) → Bool
| s, p => p.byteIdx ≥ utf8ByteSize s
String.get'
definition
(s : @& String) (p : @& Pos) (h : ¬s.atEnd p) : Char
/--
Returns the character at position `p` of a string. Returns `(default : Char)`, which is `'A'`, if
`p` is not a valid position.
Requires evidence, `h`, that `p` is within bounds instead of performing a run-time bounds check as
in `String.get`.
A typical pattern combines `get'` with a dependent `if`-expression to avoid the overhead of an
additional bounds check. For example:
```
def getInBounds? (s : String) (p : String.Pos) : Option Char :=
if h : s.atEnd p then none else some (s.get' p h)
```
Even with evidence of `¬ s.atEnd p`, `p` may be invalid if a byte index points into the middle of a
multi-byte UTF-8 character. For example, `"L∃∀N".get' ⟨2⟩ (by decide) = (default : Char)`.
Examples:
* `"abc".get' 0 (by decide) = 'a'`
* `let lean := "L∃∀N"; lean.get' (0 |> lean.next |> lean.next) (by decide) = '∀'`
-/
@[extern "lean_string_utf8_get_fast"]
def get' (s : @& String) (p : @& Pos) (h : ¬ s.atEnd p) : Char :=
match s with
| ⟨s⟩ => utf8GetAux s 0 p
String.next'
definition
(s : @& String) (p : @& Pos) (h : ¬s.atEnd p) : Pos
/--
Returns the next position in a string after position `p`. The result is unspecified if `p` is not a
valid position.
Requires evidence, `h`, that `p` is within bounds. No run-time bounds check is performed, as in
`String.next`.
A typical pattern combines `String.next'` with a dependent `if`-expression to avoid the overhead of
an additional bounds check. For example:
```
def next? (s: String) (p : String.Pos) : Option Char :=
if h : s.atEnd p then none else s.get (s.next' p h)
```
Example:
* `let abc := "abc"; abc.get (abc.next' 0 (by decide)) = 'b'`
-/
@[extern "lean_string_utf8_next_fast"]
def next' (s : @& String) (p : @& Pos) (h : ¬ s.atEnd p) : Pos :=
let c := get s p
p + c
Char.utf8Size_pos
theorem
(c : Char) : 0 < c.utf8Size
theorem _root_.Char.utf8Size_pos (c : Char) : 0 < c.utf8Size := by
repeat first | apply iteInduction (motive := (0 < ·)) <;> intros | decide
Char.utf8Size_le_four
theorem
(c : Char) : c.utf8Size ≤ 4
theorem _root_.Char.utf8Size_le_four (c : Char) : c.utf8Size ≤ 4 := by
repeat first | apply iteInduction (motive := (· ≤ 4)) <;> intros | decide
String.one_le_csize
abbrev
@[deprecated Char.utf8Size_pos (since := "2026-06-04")] abbrev one_le_csize := Char.utf8Size_pos
String.pos_lt_eq
theorem
(p₁ p₂ : Pos) : (p₁ < p₂) = (p₁.1 < p₂.1)
String.pos_add_char
theorem
(p : Pos) (c : Char) : (p + c).byteIdx = p.byteIdx + c.utf8Size
String.Pos.ne_zero_of_lt
theorem
: {a b : Pos} → a < b → b ≠ 0
protected theorem Pos.ne_zero_of_lt : {a b : Pos} → a < b → b ≠ 0
| _, _, hlt, rfl => Nat.not_lt_zero _ hlt
String.lt_next
theorem
(s : String) (i : Pos) : i.1 < (s.next i).1
theorem lt_next (s : String) (i : Pos) : i.1 < (s.next i).1 :=
Nat.add_lt_add_left (Char.utf8Size_pos _) _
String.utf8PrevAux_lt_of_pos
theorem
: ∀ (cs : List Char) (i p : Pos), p ≠ 0 → (utf8PrevAux cs i p).1 < p.1
theorem utf8PrevAux_lt_of_pos : ∀ (cs : List Char) (i p : Pos), p ≠ 0 →
(utf8PrevAux cs i p).1 < p.1
| [], _, _, h =>
Nat.lt_of_le_of_lt (Nat.zero_le _)
(Nat.zero_lt_of_ne_zero (mt (congrArg Pos.mk) h))
| c::cs, i, p, h => by
simp [utf8PrevAux]
apply iteInduction (motive := (Pos.byteIdx · < _)) <;> intro h'
next => exact h' ▸ Nat.add_lt_add_left (Char.utf8Size_pos _) _
next => exact utf8PrevAux_lt_of_pos _ _ _ h
String.prev_lt_of_pos
theorem
(s : String) (i : Pos) (h : i ≠ 0) : (s.prev i).1 < i.1
theorem prev_lt_of_pos (s : String) (i : Pos) (h : i ≠ 0) : (s.prev i).1 < i.1 := by
simp [prev, h]
exact utf8PrevAux_lt_of_pos _ _ _ h
String.posOfAux
definition
(s : String) (c : Char) (stopPos : Pos) (pos : Pos) : Pos
String.posOf
definition
(s : String) (c : Char) : Pos
/--
Returns the position of the first occurrence of a character, `c`, in a string `s`. If `s` does not
contain `c`, returns `s.endPos`.
Examples:
* `"abcba".posOf 'a' = ⟨0⟩`
* `"abcba".posOf 'z' = ⟨5⟩`
* `"L∃∀N".posOf '∀' = ⟨4⟩`
-/
@[inline] def posOf (s : String) (c : Char) : Pos :=
posOfAux s c s.endPos 0
String.revPosOfAux
definition
(s : String) (c : Char) (pos : Pos) : Option Pos
def revPosOfAux (s : String) (c : Char) (pos : Pos) : Option Pos :=
if h : pos = 0 then none
else
have := prev_lt_of_pos s pos h
let pos := s.prev pos
if s.get pos == c then some pos
else revPosOfAux s c pos
termination_by pos.1
String.revPosOf
definition
(s : String) (c : Char) : Option Pos
/--
Returns the position of the last occurrence of a character, `c`, in a string `s`. If `s` does not
contain `c`, returns `none`.
Examples:
* `"abcabc".refPosOf 'a' = some ⟨3⟩`
* `"abcabc".revPosOf 'z' = none`
* `"L∃∀N".revPosOf '∀' = some ⟨4⟩`
-/
@[inline] def revPosOf (s : String) (c : Char) : Option Pos :=
revPosOfAux s c s.endPos
String.findAux
definition
(s : String) (p : Char → Bool) (stopPos : Pos) (pos : Pos) : Pos
String.find
definition
(s : String) (p : Char → Bool) : Pos
/--
Finds the position of the first character in a string for which the Boolean predicate `p` returns
`true`. If there is no such character in the string, then the end position of the string is
returned.
Examples:
* `"coffee tea water".find (·.isWhitespace) = ⟨6⟩`
* `"tea".find (· == 'X') = ⟨3⟩`
* `"".find (· == 'X') = ⟨0⟩`
-/
@[inline] def find (s : String) (p : Char → Bool) : Pos :=
findAux s p s.endPos 0
String.revFindAux
definition
(s : String) (p : Char → Bool) (pos : Pos) : Option Pos
def revFindAux (s : String) (p : Char → Bool) (pos : Pos) : Option Pos :=
if h : pos = 0 then none
else
have := prev_lt_of_pos s pos h
let pos := s.prev pos
if p (s.get pos) then some pos
else revFindAux s p pos
termination_by pos.1
String.revFind
definition
(s : String) (p : Char → Bool) : Option Pos
/--
Finds the position of the last character in a string for which the Boolean predicate `p` returns
`true`. If there is no such character in the string, then `none` is returned.
Examples:
* `"coffee tea water".revFind (·.isWhitespace) = some ⟨10⟩`
* `"tea".revFind (· == 'X') = none`
* `"".revFind (· == 'X') = none`
-/
@[inline] def revFind (s : String) (p : Char → Bool) : Option Pos :=
revFindAux s p s.endPos
String.Pos.min
abbrev
(p₁ p₂ : Pos) : Pos
String.firstDiffPos
definition
(a b : String) : Pos
/--
Returns the first position where the two strings differ.
If one string is a prefix of the other, then the returned position is the end position of the
shorter string. If the strings are identical, then their end position is returned.
Examples:
* `"tea".firstDiffPos "ten" = ⟨2⟩`
* `"tea".firstDiffPos "tea" = ⟨3⟩`
* `"tea".firstDiffPos "teas" = ⟨3⟩`
* `"teas".firstDiffPos "tea" = ⟨3⟩`
-/
def firstDiffPos (a b : String) : Pos :=
let stopPos := a.endPos.min b.endPos
let rec loop (i : Pos) : Pos :=
if h : i < stopPos then
if a.get i != b.get i then i
else
have := Nat.sub_lt_sub_left h (lt_next a i)
loop (a.next i)
else i
termination_by stopPos.1 - i.1
loop 0
String.extract
definition
: (@& String) → (@& Pos) → (@& Pos) → String
/--
Creates a new string that consists of the region of the input string delimited by the two positions.
The result is `""` if the start position is greater than or equal to the end position or if the
start position is at the end of the string. If either position is invalid (that is, if either points
at the middle of a multi-byte UTF-8 character) then the result is unspecified.
Examples:
* `"red green blue".extract ⟨0⟩ ⟨3⟩ = "red"`
* `"red green blue".extract ⟨3⟩ ⟨0⟩ = ""`
* `"red green blue".extract ⟨0⟩ ⟨100⟩ = "red green blue"`
* `"red green blue".extract ⟨4⟩ ⟨100⟩ = "green blue"`
* `"L∃∀N".extract ⟨2⟩ ⟨100⟩ = "green blue"`
-/
@[extern "lean_string_utf8_extract"]
def extract : (@& String) → (@& Pos) → (@& Pos) → String
| ⟨s⟩, b, e => if b.byteIdx ≥ e.byteIdx then "" else ⟨go₁ s 0 b e⟩
where
go₁ : List Char → Pos → Pos → Pos → List Char
| [], _, _, _ => []
| s@(c::cs), i, b, e => if i = b then go₂ s i e else go₁ cs (i + c) b e
go₂ : List Char → Pos → Pos → List Char
| [], _, _ => []
| c::cs, i, e => if i = e then [] else c :: go₂ cs (i + c) e
String.splitAux
definition
(s : String) (p : Char → Bool) (b : Pos) (i : Pos) (r : List String) : List String
@[specialize] def splitAux (s : String) (p : Char → Bool) (b : Pos) (i : Pos) (r : List String) : List String :=
if h : s.atEnd i then
let r := (s.extract b i)::r
r.reverse
else
have := Nat.sub_lt_sub_left (Nat.gt_of_not_le (mt decide_eq_true h)) (lt_next s _)
if p (s.get i) then
let i' := s.next i
splitAux s p i' i' (s.extract b i :: r)
else
splitAux s p b (s.next i) r
termination_by s.endPos.1 - i.1
String.split
definition
(s : String) (p : Char → Bool) : List String
/--
Splits a string at each character for which `p` returns `true`.
The characters that satisfy `p` are not included in any of the resulting strings. If multiple
characters in a row satisfy `p`, then the resulting list will contain empty strings.
Examples:
* `"coffee tea water".split (·.isWhitespace) = ["coffee", "tea", "water"]`
* `"coffee tea water".split (·.isWhitespace) = ["coffee", "", "tea", "", "water"]`
* `"fun x =>\n x + 1\n".split (· == '\n') = ["fun x =>", " x + 1", ""]`
-/
@[specialize] def split (s : String) (p : Char → Bool) : List String :=
splitAux s p 0 0 []
String.splitOnAux
definition
(s sep : String) (b : Pos) (i : Pos) (j : Pos) (r : List String) : List String
/--
Auxiliary for `splitOn`. Preconditions:
* `sep` is not empty
* `b <= i` are indexes into `s`
* `j` is an index into `sep`, and not at the end
It represents the state where we have currently parsed some split parts into `r` (in reverse order),
`b` is the beginning of the string / the end of the previous match of `sep`, and the first `j` bytes
of `sep` match the bytes `i-j .. i` of `s`.
-/
def splitOnAux (s sep : String) (b : Pos) (i : Pos) (j : Pos) (r : List String) : List String :=
if s.atEnd i then
let r := (s.extract b i)::r
r.reverse
else
if s.get i == sep.get j then
let i := s.next i
let j := sep.next j
if sep.atEnd j then
splitOnAux s sep i i 0 (s.extract b (i - j)::r)
else
splitOnAux s sep b i j r
else
splitOnAux s sep b (s.next (i - j)) 0 r
termination_by (s.endPos.1 - (i - j).1, sep.endPos.1 - j.1)
decreasing_by
all_goals simp_wf
focus
rename_i h _ _
left; exact Nat.sub_lt_sub_left
(Nat.lt_of_le_of_lt (Nat.sub_le ..) (Nat.gt_of_not_le (mt decide_eq_true h)))
(Nat.lt_of_le_of_lt (Nat.sub_le ..) (lt_next s _))
focus
rename_i i₀ j₀ _ eq h'
rw [show (s.next i₀ - sep.next j₀).1 = (i₀ - j₀).1 by
show (_ + Char.utf8Size _) - (_ + Char.utf8Size _) = _
rw [(beq_iff_eq ..).1 eq, Nat.add_sub_add_right]; rfl]
right; exact Nat.sub_lt_sub_left
(Nat.lt_of_le_of_lt (Nat.le_add_right ..) (Nat.gt_of_not_le (mt decide_eq_true h')))
(lt_next sep _)
focus
rename_i h _
left; exact Nat.sub_lt_sub_left
(Nat.lt_of_le_of_lt (Nat.sub_le ..) (Nat.gt_of_not_le (mt decide_eq_true h)))
(lt_next s _)
String.instInhabited
instance
: Inhabited String
String.instAppend
instance
: Append String
instance : Append String := ⟨String.append⟩
String.pushn
definition
(s : String) (c : Char) (n : Nat) : String
/--
Adds multiple repetitions of a character to the end of a string.
Returns `s`, with `n` repetitions of `c` at the end. Internally, the implementation repeatedly calls
`String.push`, so the string is modified in-place if there is a unique reference to it.
Examples:
* `"indeed".pushn '!' 2 = "indeed!!"`
* `"indeed".pushn '!' 0 = "indeed"`
* `"".pushn ' ' 4 = " "`
-/
@[inline] def pushn (s : String) (c : Char) (n : Nat) : String :=
n.repeat (fun s => s.push c) s
String.isEmpty
definition
(s : String) : Bool
String.join
definition
(l : List String) : String
/--
Appends all the strings in a list of strings, in order.
Use `String.intercalate` to place a separator string between the strings in a list.
Examples:
* `String.join ["gr", "ee", "n"] = "green"`
* `String.join ["b", "", "l", "", "ue"] = "red"`
* `String.join [] = ""`
-/
@[inline] def join (l : List String) : String :=
l.foldl (fun r s => r ++ s) ""
String.singleton
definition
(c : Char) : String
/--
Returns a new string that contains only the character `c`.
Because strings are encoded in UTF-8, the resulting string may take multiple bytes.
Examples:
* `String.singleton 'L' = "L"`
* `String.singleton ' ' = " "`
* `String.singleton '"' = "\""`
* `String.singleton '𝒫' = "𝒫"`
-/
@[inline] def singleton (c : Char) : String :=
"".push c
String.intercalate
definition
(s : String) : List String → String
/--
Appends the strings in a list of strings, placing the separator `s` between each pair.
Examples:
* `", ".intercalate ["red", "green", "blue"] = "red, green, blue"`
* `" and ".intercalate ["tea", "coffee"] = "tea and coffee"`
* `" | ".intercalate ["M", "", "N"] = "M | | N"`
-/
def intercalate (s : String) : List String → String
| [] => ""
| a :: as => go a s as
where go (acc : String) (s : String) : List String → String
| a :: as => go (acc ++ s ++ a) s as
| [] => acc
String.Iterator
structure
/--
An iterator over the characters (Unicode code points) in a `String`. Typically created by
`String.iter`.
String iterators pair a string with a valid byte index. This allows efficient character-by-character
processing of strings while avoiding the need to manually ensure that byte indices are used with the
correct strings.
An iterator is *valid* if the position `i` is *valid* for the string `s`, meaning `0 ≤ i ≤ s.endPos`
and `i` lies on a UTF8 byte boundary. If `i = s.endPos`, the iterator is at the end of the string.
Most operations on iterators return unspecified values if the iterator is not valid. The functions
in the `String.Iterator` API rule out the creation of invalid iterators, with two exceptions:
- `Iterator.next iter` is invalid if `iter` is already at the end of the string (`iter.atEnd` is
`true`), and
- `Iterator.forward iter n`/`Iterator.nextn iter n` is invalid if `n` is strictly greater than the
number of remaining characters.
-/
structure Iterator where
/-- The string being iterated over. -/
s : String
/-- The current UTF-8 byte position in the string `s`.
This position is not guaranteed to be valid for the string. If the position is not valid, then the
current character is `(default : Char)`, similar to `String.get` on an invalid position.
-/
i : Pos
deriving DecidableEq, Inhabited
String.instDecidableEqIterator
instance
: DecidableEq✝ (@String.Iterator✝)
DecidableEq, Inhabited
String.instInhabitedIterator
instance
: Inhabited✝ (@String.Iterator)
Inhabited
String.mkIterator
definition
(s : String) : Iterator
/-- Creates an iterator at the beginning of the string. -/
@[inline] def mkIterator (s : String) : Iterator :=
⟨s, 0⟩
String.iter
abbrev
@[inherit_doc mkIterator]
abbrev iter := mkIterator
String.instSizeOfIterator
instance
: SizeOf String.Iterator
/--
The size of a string iterator is the number of bytes remaining.
Recursive functions that iterate towards the end of a string will typically decrease this measure.
-/
instance : SizeOf String.Iterator where
sizeOf i := i.1.utf8ByteSize - i.2.byteIdx
String.Iterator.sizeOf_eq
theorem
(i : String.Iterator) : sizeOf i = i.1.utf8ByteSize - i.2.byteIdx
theorem Iterator.sizeOf_eq (i : String.Iterator) : sizeOf i = i.1.utf8ByteSize - i.2.byteIdx :=
rfl
String.Iterator.toString
definition
@[inline, inherit_doc Iterator.s]
def toString := Iterator.s
String.Iterator.remainingBytes
definition
: Iterator → Nat
/--
The number of UTF-8 bytes remaining in the iterator.
-/
@[inline] def remainingBytes : Iterator → Nat
| ⟨s, i⟩ => s.endPos.byteIdx - i.byteIdx
String.Iterator.pos
definition
@[inline, inherit_doc Iterator.i]
def pos := Iterator.i
String.Iterator.curr
definition
: Iterator → Char
String.Iterator.next
definition
: Iterator → Iterator
/--
Moves the iterator's position forward by one character, unconditionally.
It is only valid to call this function if the iterator is not at the end of the string (i.e.
if `Iterator.atEnd` is `false`); otherwise, the resulting iterator will be invalid.
-/
@[inline] def next : Iterator → Iterator
| ⟨s, i⟩ => ⟨s, s.next i⟩
String.Iterator.prev
definition
: Iterator → Iterator
String.Iterator.atEnd
definition
: Iterator → Bool
String.Iterator.hasNext
definition
: Iterator → Bool
String.Iterator.hasPrev
definition
: Iterator → Bool
String.Iterator.curr'
definition
(it : Iterator) (h : it.hasNext) : Char
/--
Gets the character at the iterator's current position.
The proof of `it.hasNext` ensures that there is, in fact, a character at the current position. This
function is faster that `String.Iterator.curr` due to avoiding a run-time bounds check.
-/
@[inline] def curr' (it : Iterator) (h : it.hasNext) : Char :=
match it with
| ⟨s, i⟩ => get' s i (by simpa only [hasNext, endPos, decide_eq_true_eq, String.atEnd, ge_iff_le, Nat.not_le] using h)
String.Iterator.next'
definition
(it : Iterator) (h : it.hasNext) : Iterator
/--
Moves the iterator's position forward by one character, unconditionally.
The proof of `it.hasNext` ensures that there is, in fact, a position that's one character forwards.
This function is faster that `String.Iterator.next` due to avoiding a run-time bounds check.
-/
@[inline] def next' (it : Iterator) (h : it.hasNext) : Iterator :=
match it with
| ⟨s, i⟩ => ⟨s, s.next' i (by simpa only [hasNext, endPos, decide_eq_true_eq, String.atEnd, ge_iff_le, Nat.not_le] using h)⟩
String.Iterator.setCurr
definition
: Iterator → Char → Iterator
/--
Replaces the current character in the string.
Does nothing if the iterator is at the end of the string. If both the replacement character and the
replaced character are 7-bit ASCII characters and the string is not shared, then it is updated
in-place and not copied.
-/
@[inline] def setCurr : Iterator → Char → Iterator
| ⟨s, i⟩, c => ⟨s.set i c, i⟩
String.Iterator.toEnd
definition
: Iterator → Iterator
String.Iterator.extract
definition
: Iterator → Iterator → String
/--
Extracts the substring between the positions of two iterators. The first iterator's position is the
start of the substring, and the second iterator's position is the end.
Returns the empty string if the iterators are for different strings, or if the position of the first
iterator is past the position of the second iterator.
-/
@[inline] def extract : Iterator → Iterator → String
| ⟨s₁, b⟩, ⟨s₂, e⟩ =>
if s₁ ≠ s₂ || b > e then ""
else s₁.extract b e
String.Iterator.forward
definition
: Iterator → Nat → Iterator
/--
Moves the iterator's position forward by the specified number of characters.
The resulting iterator is only valid if the number of characters to skip is less than or equal
to the number of characters left in the iterator.
-/
def forward : Iterator → Nat → Iterator
| it, 0 => it
| it, n+1 => forward it.next n
String.Iterator.remainingToString
definition
: Iterator → String
/--
The remaining characters in an iterator, as a string.
-/
@[inline] def remainingToString : Iterator → String
| ⟨s, i⟩ => s.extract i s.endPos
String.Iterator.nextn
definition
: Iterator → Nat → Iterator
String.Iterator.prevn
definition
: Iterator → Nat → Iterator
String.offsetOfPosAux
definition
(s : String) (pos : Pos) (i : Pos) (offset : Nat) : Nat
def offsetOfPosAux (s : String) (pos : Pos) (i : Pos) (offset : Nat) : Nat :=
if i >= pos then offset
else if h : s.atEnd i then
offset
else
have := Nat.sub_lt_sub_left (Nat.gt_of_not_le (mt decide_eq_true h)) (lt_next s _)
offsetOfPosAux s pos (s.next i) (offset+1)
termination_by s.endPos.1 - i.1
String.offsetOfPos
definition
(s : String) (pos : Pos) : Nat
/--
Returns the character index that corresponds to the provided position (i.e. UTF-8 byte index) in a
string.
If the position is at the end of the string, then the string's length in characters is returned. If
the position is invalid due to pointing at the middle of a UTF-8 byte sequence, then the character
index of the next character after the position is returned.
Examples:
* `"L∃∀N".offsetOfPos ⟨0⟩ = 0`
* `"L∃∀N".offsetOfPos ⟨1⟩ = 1`
* `"L∃∀N".offsetOfPos ⟨2⟩ = 2`
* `"L∃∀N".offsetOfPos ⟨4⟩ = 2`
* `"L∃∀N".offsetOfPos ⟨5⟩ = 3`
* `"L∃∀N".offsetOfPos ⟨50⟩ = 4`
-/
@[inline] def offsetOfPos (s : String) (pos : Pos) : Nat :=
offsetOfPosAux s pos 0 0
String.foldlAux
definition
{α : Type u} (f : α → Char → α) (s : String) (stopPos : Pos) (i : Pos) (a : α) : α
String.foldl
definition
{α : Type u} (f : α → Char → α) (init : α) (s : String) : α
/--
Folds a function over a string from the left, accumulating a value starting with `init`. The
accumulated value is combined with each character in order, using `f`.
Examples:
* `"coffee tea water".foldl (fun n c => if c.isWhitespace then n + 1 else n) 0 = 2`
* `"coffee tea and water".foldl (fun n c => if c.isWhitespace then n + 1 else n) 0 = 3`
* `"coffee tea water".foldl (·.push ·) "" = "coffee tea water"`
-/
@[inline] def foldl {α : Type u} (f : α → Char → α) (init : α) (s : String) : α :=
foldlAux f s s.endPos 0 init
String.foldrAux
definition
{α : Type u} (f : Char → α → α) (a : α) (s : String) (i begPos : Pos) : α
@[specialize] def foldrAux {α : Type u} (f : Char → α → α) (a : α) (s : String) (i begPos : Pos) : α :=
if h : begPos < i then
have := String.prev_lt_of_pos s i <| mt (congrArg String.Pos.byteIdx) <|
Ne.symm <| Nat.ne_of_lt <| Nat.lt_of_le_of_lt (Nat.zero_le _) h
let i := s.prev i
let a := f (s.get i) a
foldrAux f a s i begPos
else a
termination_by i.1
String.foldr
definition
{α : Type u} (f : Char → α → α) (init : α) (s : String) : α
/--
Folds a function over a string from the right, accumulating a value starting with `init`. The
accumulated value is combined with each character in reverse order, using `f`.
Examples:
* `"coffee tea water".foldr (fun c n => if c.isWhitespace then n + 1 else n) 0 = 2`
* `"coffee tea and water".foldr (fun c n => if c.isWhitespace then n + 1 else n) 0 = 3`
* `"coffee tea water".foldr (fun c s => c.push s) "" = "retaw dna aet eeffoc"`
-/
@[inline] def foldr {α : Type u} (f : Char → α → α) (init : α) (s : String) : α :=
foldrAux f init s s.endPos 0
String.anyAux
definition
(s : String) (stopPos : Pos) (p : Char → Bool) (i : Pos) : Bool
String.any
definition
(s : String) (p : Char → Bool) : Bool
/--
Checks whether there is a character in a string for which the Boolean predicate `p` returns `true`.
Short-circuits at the first character for which `p` returns `true`.
Examples:
* `"brown".any (·.isLetter) = true`
* `"brown".any (·.isWhitespace) = false`
* `"brown and orange".any (·.isLetter) = true`
* `"".any (fun _ => false) = false`
-/
@[inline] def any (s : String) (p : Char → Bool) : Bool :=
anyAux s s.endPos p 0
String.all
definition
(s : String) (p : Char → Bool) : Bool
/--
Checks whether the Boolean predicate `p` returns `true` for every character in a string.
Short-circuits at the first character for which `p` returns `false`.
Examples:
* `"brown".all (·.isLetter) = true`
* `"brown and orange".all (·.isLetter) = false`
* `"".all (fun _ => false) = true`
-/
@[inline] def all (s : String) (p : Char → Bool) : Bool :=
!s.any (fun c => !p c)
String.contains
definition
(s : String) (c : Char) : Bool
String.utf8SetAux_of_gt
theorem
(c' : Char) : ∀ (cs : List Char) {i p : Pos}, i > p → utf8SetAux c' cs i p = cs
theorem utf8SetAux_of_gt (c' : Char) : ∀ (cs : List Char) {i p : Pos}, i > p → utf8SetAux c' cs i p = cs
| [], _, _, _ => rfl
| c::cs, i, p, h => by
rw [utf8SetAux, if_neg (mt (congrArg (·.1)) (Ne.symm <| Nat.ne_of_lt h)), utf8SetAux_of_gt c' cs]
exact Nat.lt_of_lt_of_le h (Nat.le_add_right ..)
String.set_next_add
theorem
(s : String) (i : Pos) (c : Char) (b₁ b₂) (h : (s.next i).1 + b₁ = s.endPos.1 + b₂) :
((s.set i c).next i).1 + b₁ = (s.set i c).endPos.1 + b₂
theorem set_next_add (s : String) (i : Pos) (c : Char) (b₁ b₂)
(h : (s.next i).1 + b₁ = s.endPos.1 + b₂) :
((s.set i c).next i).1 + b₁ = (s.set i c).endPos.1 + b₂ := by
simp [next, get, set, endPos, utf8ByteSize] at h ⊢
rw [Nat.add_comm i.1, Nat.add_assoc] at h ⊢
let rec foo : ∀ cs a b₁ b₂,
(utf8GetAux cs a i).utf8Size + b₁ = utf8ByteSize.go cs + b₂ →
(utf8GetAux (utf8SetAux c cs a i) a i).utf8Size + b₁ = utf8ByteSize.go (utf8SetAux c cs a i) + b₂
| [], _, _, _, h => h
| c'::cs, a, b₁, b₂, h => by
unfold utf8SetAux
apply iteInduction (motive := fun p => (utf8GetAux p a i).utf8Size + b₁ = utf8ByteSize.go p + b₂) <;>
intro h' <;> simp [utf8GetAux, h', utf8ByteSize.go] at h ⊢
next =>
rw [Nat.add_assoc, Nat.add_left_comm] at h ⊢; rw [Nat.add_left_cancel h]
next =>
rw [Nat.add_assoc] at h ⊢
refine foo cs (a + c') b₁ (c'.utf8Size + b₂) h
exact foo s.1 0 _ _ h
String.mapAux_lemma
theorem
(s : String) (i : Pos) (c : Char) (h : ¬s.atEnd i) : (s.set i c).endPos.1 - ((s.set i c).next i).1 < s.endPos.1 - i.1
theorem mapAux_lemma (s : String) (i : Pos) (c : Char) (h : ¬s.atEnd i) :
(s.set i c).endPos.1 - ((s.set i c).next i).1 < s.endPos.1 - i.1 :=
suffices (s.set i c).endPos.1 - ((s.set i c).next i).1 = s.endPos.1 - (s.next i).1 by
rw [this]
apply Nat.sub_lt_sub_left (Nat.gt_of_not_le (mt decide_eq_true h)) (lt_next ..)
Nat.sub.elim (motive := (_ = ·)) _ _
(fun _ k e =>
have := set_next_add _ _ _ k 0 e.symm
Nat.sub_eq_of_eq_add <| this.symm.trans <| Nat.add_comm ..)
(fun h => by
have ⟨k, e⟩ := Nat.le.dest h
rw [Nat.succ_add] at e
have : ((s.set i c).next i).1 = _ := set_next_add _ _ c 0 k.succ e.symm
exact Nat.sub_eq_zero_of_le (this ▸ Nat.le_add_right ..))
String.mapAux
definition
(f : Char → Char) (i : Pos) (s : String) : String
String.map
definition
(f : Char → Char) (s : String) : String
String.isNat
definition
(s : String) : Bool
/--
Checks whether the string can be interpreted as the decimal representation of a natural number.
A string can be interpreted as a decimal natural number if it is not empty and all the characters in
it are digits.
Use `String.toNat?` or `String.toNat!` to convert such a string to a natural number.
Examples:
* `"".isNat = false`
* `"0".isNat = true`
* `"5".isNat = true`
* `"05".isNat = true`
* `"587".isNat = true`
* `"-587".isNat = false`
* `" 5".isNat = false`
* `"2+3".isNat = false`
* `"0xff".isNat = false`
-/
@[inline] def isNat (s : String) : Bool :=
!s.isEmpty!s.isEmpty && s.all (·.isDigit)
String.toNat?
definition
(s : String) : Option Nat
/--
Interprets a string as the decimal representation of a natural number, returning it. Returns `none`
if the string does not contain a decimal natural number.
A string can be interpreted as a decimal natural number if it is not empty and all the characters in
it are digits.
Use `String.isNat` to check whether `String.toNat?` would return `some`. `String.toNat!` is an
alternative that panics instead of returning `none` when the string is not a natural number.
Examples:
* `"".toNat? = none`
* `"0".toNat? = some 0`
* `"5".toNat? = some 5`
* `"587".toNat? = some 587`
* `"-587".toNat? = none`
* `" 5".toNat? = none`
* `"2+3".toNat? = none`
* `"0xff".toNat? = none`
-/
def toNat? (s : String) : Option Nat :=
if s.isNat then
some <| s.foldl (fun n c => n*10 + (c.toNat - '0'.toNat)) 0
else
none
String.substrEq
definition
(s1 : String) (pos1 : String.Pos) (s2 : String) (pos2 : String.Pos) (sz : Nat) : Bool
/--
Checks whether substrings of two strings are equal. Substrings are indicated by their starting
positions and a size in _UTF-8 bytes_. Returns `false` if the indicated substring does not exist in
either string.
-/
def substrEq (s1 : String) (pos1 : String.Pos) (s2 : String) (pos2 : String.Pos) (sz : Nat) : Bool :=
pos1.byteIdx + sz ≤ s1.endPos.byteIdx && pos2.byteIdx + sz ≤ s2.endPos.byteIdxpos1.byteIdx + sz ≤ s1.endPos.byteIdx && pos2.byteIdx + sz ≤ s2.endPos.byteIdx && loop pos1 pos2 { byteIdx := pos1.byteIdx + sz }
where
loop (off1 off2 stop1 : Pos) :=
if _h : off1.byteIdx < stop1.byteIdx then
let c₁ := s1.get off1
let c₂ := s2.get off2
c₁ == c₂ && loop (off1 + c₁) (off2 + c₂) stop1
else true
termination_by stop1.1 - off1.1
decreasing_by
have := Nat.sub_lt_sub_left _h (Nat.add_lt_add_left c₁.utf8Size_pos off1.1)
decreasing_tactic
String.isPrefixOf
definition
(p : String) (s : String) : Bool
/--
Checks whether the first string (`p`) is a prefix of the second (`s`).
`String.startsWith` is a version that takes the potential prefix after the string.
Examples:
* `"red".isPrefixOf "red green blue" = true`
* `"green".isPrefixOf "red green blue" = false`
* `"".isPrefixOf "red green blue" = true`
-/
def isPrefixOf (p : String) (s : String) : Bool :=
substrEq p 0 s 0 p.endPos.byteIdx
String.replace
definition
(s pattern replacement : String) : String
/--
In the string `s`, replaces all occurrences of `pattern` with `replacement`.
Examples:
* `"red green blue".replace "e" "" = "rd grn blu"`
* `"red green blue".replace "ee" "E" = "red grEn blue"`
* `"red green blue".replace "e" "E" = "rEd grEEn bluE"`
-/
def replace (s pattern replacement : String) : String :=
if h : pattern.endPos.1 = 0 then s
else
have hPatt := Nat.zero_lt_of_ne_zero h
let rec loop (acc : String) (accStop pos : String.Pos) :=
if h : pos.byteIdx + pattern.endPos.byteIdx > s.endPos.byteIdx then
acc ++ s.extract accStop s.endPos
else
have := Nat.lt_of_lt_of_le (Nat.add_lt_add_left hPatt _) (Nat.ge_of_not_lt h)
if s.substrEq pos pattern 0 pattern.endPos.byteIdx then
have := Nat.sub_lt_sub_left this (Nat.add_lt_add_left hPatt _)
loop (acc ++ s.extract accStop pos ++ replacement) (pos + pattern) (pos + pattern)
else
have := Nat.sub_lt_sub_left this (lt_next s pos)
loop acc accStop (s.next pos)
termination_by s.endPos.1 - pos.1
loop "" 0 0
String.findLineStart
definition
(s : String) (pos : String.Pos) : String.Pos
/--
Returns the position of the beginning of the line that contains the position `pos`.
Lines are ended by `'\n'`, and the returned position is either `0 : String.Pos` or immediately after
a `'\n'` character.
-/
def findLineStart (s : String) (pos : String.Pos) : String.Pos :=
match s.revFindAux (· = '\n') pos with
| none => 0
| somesome n => ⟨n.byteIdx + 1⟩
Substring.isEmpty
definition
(ss : Substring) : Bool
Substring.toString
definition
: Substring → String
Substring.toIterator
definition
: Substring → String.Iterator
/--
Returns an iterator into the underlying string, at the substring's starting position. The ending
position is discarded, so the iterator alone cannot be used to determine whether its current
position is within the original substring.
-/
@[inline] def toIterator : Substring → String.Iterator
| ⟨s, b, _⟩ => ⟨s, b⟩
Substring.get
definition
: Substring → String.Pos → Char
/--
Returns the character at the given position in the substring.
The position is relative to the substring, rather than the underlying string, and no bounds checking
is performed with respect to the substring's end position. If the relative position is not a valid
position in the underlying string, the fallback value `(default : Char)`, which is `'A'`, is
returned. Does not panic.
-/
@[inline] def get : Substring → String.Pos → Char
| ⟨s, b, _⟩, p => s.get (b+p)
Substring.next
definition
: Substring → String.Pos → String.Pos
/--
Returns the next position in a substring after the given position. If the position is at the end of
the substring, it is returned unmodified.
Both the input position and the returned position are interpreted relative to the substring's start
position, not the underlying string.
-/
@[inline] def next : Substring → String.Pos → String.Pos
| ⟨s, b, e⟩, p =>
let absP := b+p
if absP = e then p else { byteIdx := (s.next absP).byteIdx - b.byteIdx }
Substring.lt_next
theorem
(s : Substring) (i : String.Pos) (h : i.1 < s.bsize) : i.1 < (s.next i).1
theorem lt_next (s : Substring) (i : String.Pos) (h : i.1 < s.bsize) :
i.1 < (s.next i).1 := by
simp [next]; rw [if_neg ?a]
case a =>
refine mt (congrArg String.Pos.byteIdx) (Nat.ne_of_lt ?_)
exact (Nat.add_comm .. ▸ Nat.add_lt_of_lt_sub h :)
apply Nat.lt_sub_of_add_lt
rw [Nat.add_comm]; apply String.lt_next
Substring.prev
definition
: Substring → String.Pos → String.Pos
/--
Returns the previous position in a substring, just prior to the given position. If the position is
at the beginning of the substring, it is returned unmodified.
Both the input position and the returned position are interpreted relative to the substring's start
position, not the underlying string.
-/
@[inline] def prev : Substring → String.Pos → String.Pos
| ⟨s, b, _⟩, p =>
let absP := b+p
if absP = b then p else { byteIdx := (s.prev absP).byteIdx - b.byteIdx }
Substring.nextn
definition
: Substring → Nat → String.Pos → String.Pos
/--
Returns the position that's the specified number of characters forward from the given position in a
substring. If the end position of the substring is reached, it is returned.
Both the input position and the returned position are interpreted relative to the substring's start
position, not the underlying string.
-/
def nextn : Substring → Nat → String.Pos → String.Pos
| _, 0, p => p
| ss, i+1, p => ss.nextn i (ss.next p)
Substring.prevn
definition
: Substring → Nat → String.Pos → String.Pos
/--
Returns the position that's the specified number of characters prior to the given position in a
substring. If the start position of the substring is reached, it is returned.
Both the input position and the returned position are interpreted relative to the substring's start
position, not the underlying string.
-/
def prevn : Substring → Nat → String.Pos → String.Pos
| _, 0, p => p
| ss, i+1, p => ss.prevn i (ss.prev p)
Substring.front
definition
(s : Substring) : Char
/--
Returns the first character in the substring.
If the substring is empty, but the substring's start position is a valid position in the underlying
string, then the character at the start position is returned. If the substring's start position is
not a valid position in the string, the fallback value `(default : Char)`, which is `'A'`, is
returned. Does not panic.
-/
@[inline] def front (s : Substring) : Char :=
s.get 0
Substring.posOf
definition
(s : Substring) (c : Char) : String.Pos
/--
Returns the substring-relative position of the first occurrence of `c` in `s`, or `s.bsize` if `c`
doesn't occur.
-/
@[inline] def posOf (s : Substring) (c : Char) : String.Pos :=
match s with
| ⟨s, b, e⟩ => { byteIdx := (String.posOfAux s c e b).byteIdx - b.byteIdx }
Substring.drop
definition
: Substring → Nat → Substring
/--
Removes the specified number of characters (Unicode code points) from the beginning of a substring
by advancing its start position.
If the substring's end position is reached, the start position is not advanced past it.
-/
@[inline] def drop : Substring → Nat → Substring
| ss@⟨s, b, e⟩, n => ⟨s, b + ss.nextn n 0, e⟩
Substring.dropRight
definition
: Substring → Nat → Substring
/--
Removes the specified number of characters (Unicode code points) from the end of a substring
by moving its end position towards its start position.
If the substring's start position is reached, the end position is not retracted past it.
-/
@[inline] def dropRight : Substring → Nat → Substring
| ss@⟨s, b, _⟩, n => ⟨s, b, b + ss.prevn n ⟨ss.bsize⟩⟩
Substring.take
definition
: Substring → Nat → Substring
/--
Retains only the specified number of characters (Unicode code points) at the beginning of a
substring, by moving its end position towards its start position.
If the substring's start position is reached, the end position is not retracted past it.
-/
@[inline] def take : Substring → Nat → Substring
| ss@⟨s, b, _⟩, n => ⟨s, b, b + ss.nextn n 0⟩
Substring.takeRight
definition
: Substring → Nat → Substring
/--
Retains only the specified number of characters (Unicode code points) at the end of a substring, by
moving its start position towards its end position.
If the substring's end position is reached, the start position is not advanced past it.
-/
@[inline] def takeRight : Substring → Nat → Substring
| ss@⟨s, b, e⟩, n => ⟨s, b + ss.prevn n ⟨ss.bsize⟩, e⟩
Substring.atEnd
definition
: Substring → String.Pos → Bool
/--
Checks whether a position in a substring is precisely equal to its ending position.
The position is understood relative to the substring's starting position, rather than the underlying
string's starting position.
-/
@[inline] def atEnd : Substring → String.Pos → Bool
| ⟨_, b, e⟩, p => b + p == e
Substring.extract
definition
: Substring → String.Pos → String.Pos → Substring
/--
Returns the region of the substring delimited by the provided start and stop positions, as a
substring. The positions are interpreted with respect to the substring's start position, rather than
the underlying string.
If the resulting substring is empty, then the resulting substring is a substring of the empty string
`""`. Otherwise, the underlying string is that of the input substring with the beginning and end
positions adjusted.
-/
@[inline] def extract : Substring → String.Pos → String.Pos → Substring
| ⟨s, b, e⟩, b', e' => if b' ≥ e' then ⟨"", 0, 0⟩ else ⟨s, e.min (b+b'), e.min (b+e')⟩
Substring.foldl
definition
{α : Type u} (f : α → Char → α) (init : α) (s : Substring) : α
/--
Folds a function over a substring from the left, accumulating a value starting with `init`. The
accumulated value is combined with each character in order, using `f`.
-/
@[inline] def foldl {α : Type u} (f : α → Char → α) (init : α) (s : Substring) : α :=
match s with
| ⟨s, b, e⟩ => String.foldlAux f s e b init
Substring.foldr
definition
{α : Type u} (f : Char → α → α) (init : α) (s : Substring) : α
/--
Folds a function over a substring from the right, accumulating a value starting with `init`. The
accumulated value is combined with each character in reverse order, using `f`.
-/
@[inline] def foldr {α : Type u} (f : Char → α → α) (init : α) (s : Substring) : α :=
match s with
| ⟨s, b, e⟩ => String.foldrAux f init s e b
Substring.any
definition
(s : Substring) (p : Char → Bool) : Bool
Substring.all
definition
(s : Substring) (p : Char → Bool) : Bool
Substring.contains
definition
(s : Substring) (c : Char) : Bool
Substring.takeWhileAux
definition
(s : String) (stopPos : String.Pos) (p : Char → Bool) (i : String.Pos) : String.Pos
@[specialize] def takeWhileAux (s : String) (stopPos : String.Pos) (p : Char → Bool) (i : String.Pos) : String.Pos :=
if h : i < stopPos then
if p (s.get i) then
have := Nat.sub_lt_sub_left h (String.lt_next s i)
takeWhileAux s stopPos p (s.next i)
else i
else i
termination_by stopPos.1 - i.1
Substring.takeWhile
definition
: Substring → (Char → Bool) → Substring
/--
Retains only the longest prefix of a substring in which a Boolean predicate returns `true` for all
characters by moving the substring's end position towards its start position.
-/
@[inline] def takeWhile : Substring → (Char → Bool) → Substring
| ⟨s, b, e⟩, p =>
let e := takeWhileAux s e p b;
⟨s, b, e⟩
Substring.dropWhile
definition
: Substring → (Char → Bool) → Substring
/--
Removes the longest prefix of a substring in which a Boolean predicate returns `true` for all
characters by moving the substring's start position. The start position is moved to the position of
the first character for which the predicate returns `false`, or to the substring's end position if
the predicate always returns `true`.
-/
@[inline] def dropWhile : Substring → (Char → Bool) → Substring
| ⟨s, b, e⟩, p =>
let b := takeWhileAux s e p b;
⟨s, b, e⟩
Substring.takeRightWhileAux
definition
(s : String) (begPos : String.Pos) (p : Char → Bool) (i : String.Pos) : String.Pos
@[specialize] def takeRightWhileAux (s : String) (begPos : String.Pos) (p : Char → Bool) (i : String.Pos) : String.Pos :=
if h : begPos < i then
have := String.prev_lt_of_pos s i <| mt (congrArg String.Pos.byteIdx) <|
Ne.symm <| Nat.ne_of_lt <| Nat.lt_of_le_of_lt (Nat.zero_le _) h
let i' := s.prev i
let c := s.get i'
if !p c then i
else takeRightWhileAux s begPos p i'
else i
termination_by i.1
Substring.takeRightWhile
definition
: Substring → (Char → Bool) → Substring
/--
Retains only the longest suffix of a substring in which a Boolean predicate returns `true` for all
characters by moving the substring's start position towards its end position.
-/
@[inline] def takeRightWhile : Substring → (Char → Bool) → Substring
| ⟨s, b, e⟩, p =>
let b := takeRightWhileAux s b p e
⟨s, b, e⟩
Substring.dropRightWhile
definition
: Substring → (Char → Bool) → Substring
/--
Removes the longest suffix of a substring in which a Boolean predicate returns `true` for all
characters by moving the substring's end position. The end position is moved just after the position
of the last character for which the predicate returns `false`, or to the substring's start position
if the predicate always returns `true`.
-/
@[inline] def dropRightWhile : Substring → (Char → Bool) → Substring
| ⟨s, b, e⟩, p =>
let e := takeRightWhileAux s b p e
⟨s, b, e⟩
Substring.trimLeft
definition
(s : Substring) : Substring
/--
Removes leading whitespace from a substring by moving its start position to the first non-whitespace
character, or to its end position if there is no non-whitespace character.
“Whitespace” is defined as characters for which `Char.isWhitespace` returns `true`.
-/
@[inline] def trimLeft (s : Substring) : Substring :=
s.dropWhile Char.isWhitespace
Substring.trimRight
definition
(s : Substring) : Substring
/--
Removes trailing whitespace from a substring by moving its end position to the last non-whitespace
character, or to its start position if there is no non-whitespace character.
“Whitespace” is defined as characters for which `Char.isWhitespace` returns `true`.
-/
@[inline] def trimRight (s : Substring) : Substring :=
s.dropRightWhile Char.isWhitespace
Substring.trim
definition
: Substring → Substring
/--
Removes leading and trailing whitespace from a substring by first moving its start position to the
first non-whitespace character, and then moving its end position to the last non-whitespace
character.
If the substring consists only of whitespace, then the resulting substring's start position is moved
to its end position.
“Whitespace” is defined as characters for which `Char.isWhitespace` returns `true`.
Examples:
* `" red green blue ".toSubstring.trim.toString = "red green blue"`
* `" red green blue ".toSubstring.trim.startPos = ⟨1⟩`
* `" red green blue ".toSubstring.trim.stopPos = ⟨15⟩`
* `" ".toSubstring.trim.startPos = ⟨5⟩`
-/
@[inline] def trim : Substring → Substring
| ⟨s, b, e⟩ =>
let b := takeWhileAux s e Char.isWhitespace b
let e := takeRightWhileAux s b Char.isWhitespace e
⟨s, b, e⟩
Substring.isNat
definition
(s : Substring) : Bool
/--
Checks whether the substring can be interpreted as the decimal representation of a natural number.
A substring can be interpreted as a decimal natural number if it is not empty and all the characters
in it are digits.
Use `Substring.toNat?` to convert such a substring to a natural number.
-/
@[inline] def isNat (s : Substring) : Bool :=
s.all fun c => c.isDigit
Substring.toNat?
definition
(s : Substring) : Option Nat
/--
Checks whether the substring can be interpreted as the decimal representation of a natural number,
returning the number if it can.
A substring can be interpreted as a decimal natural number if it is not empty and all the characters
in it are digits.
Use `Substring.isNat` to check whether the substring is such a substring.
-/
def toNat? (s : Substring) : Option Nat :=
if s.isNat then
some <| s.foldl (fun n c => n*10 + (c.toNat - '0'.toNat)) 0
else
none
Substring.beq
definition
(ss1 ss2 : Substring) : Bool
/--
Checks whether two substrings represent equal strings. Usually accessed via the `==` operator.
Two substrings do not need to have the same underlying string or the same start and end positions;
instead, they are equal if they contain the same sequence of characters.
-/
def beq (ss1 ss2 : Substring) : Bool :=
ss1.bsize == ss2.bsize && ss1.str.substrEq ss1.startPos ss2.str ss2.startPos ss1.bsize
Substring.hasBeq
instance
: BEq Substring
Substring.sameAs
definition
(ss1 ss2 : Substring) : Bool
/--
Checks whether two substrings have the same position and content.
The two substrings do not need to have the same underlying string for this check to succeed.
-/
def sameAs (ss1 ss2 : Substring) : Bool :=
ss1.startPos == ss2.startPos && ss1 == ss2
Substring.commonPrefix
definition
(s t : Substring) : Substring
/--
Returns the longest common prefix of two substrings.
The returned substring uses the same underlying string as `s`.
-/
def commonPrefix (s t : Substring) : Substring :=
{ s with stopPos := loop s.startPos t.startPos }
where
/-- Returns the ending position of the common prefix, working up from `spos, tpos`. -/
loop spos tpos :=
if h : spos < s.stopPos ∧ tpos < t.stopPos then
if s.str.get spos == t.str.get tpos then
have := Nat.sub_lt_sub_left h.1 (s.str.lt_next spos)
loop (s.str.next spos) (t.str.next tpos)
else
spos
else
spos
termination_by s.stopPos.byteIdx - spos.byteIdx
Substring.commonSuffix
definition
(s t : Substring) : Substring
/--
Returns the longest common suffix of two substrings.
The returned substring uses the same underlying string as `s`.
-/
def commonSuffix (s t : Substring) : Substring :=
{ s with startPos := loop s.stopPos t.stopPos }
where
/-- Returns the starting position of the common prefix, working down from `spos, tpos`. -/
loop spos tpos :=
if h : s.startPos < spos ∧ t.startPos < tpos then
let spos' := s.str.prev spos
let tpos' := t.str.prev tpos
if s.str.get spos' == t.str.get tpos' then
have : spos' < spos := s.str.prev_lt_of_pos spos (String.Pos.ne_zero_of_lt h.1)
loop spos' tpos'
else
spos
else
spos
termination_by spos.byteIdx
Substring.dropPrefix?
definition
(s : Substring) (pre : Substring) : Option Substring
/--
If `pre` is a prefix of `s`, returns the remainder. Returns `none` otherwise.
The substring `pre` is a prefix of `s` if there exists a `t : Substring` such that
`s.toString = pre.toString ++ t.toString`. If so, the result is the substring of `s` without the
prefix.
-/
def dropPrefix? (s : Substring) (pre : Substring) : Option Substring :=
let t := s.commonPrefix pre
if t.bsize = pre.bsize then
some { s with startPos := t.stopPos }
else
none
Substring.dropSuffix?
definition
(s : Substring) (suff : Substring) : Option Substring
/--
If `suff` is a suffix of `s`, returns the remainder. Returns `none` otherwise.
The substring `suff` is a suffix of `s` if there exists a `t : Substring` such that
`s.toString = t.toString ++ suff.toString`. If so, the result the substring of `s` without the
suffix.
-/
def dropSuffix? (s : Substring) (suff : Substring) : Option Substring :=
let t := s.commonSuffix suff
if t.bsize = suff.bsize then
some { s with stopPos := t.startPos }
else
none
String.drop
definition
(s : String) (n : Nat) : String
/--
Removes the specified number of characters (Unicode code points) from the start of the string.
If `n` is greater than `s.length`, returns `""`.
Examples:
* `"red green blue".drop 4 = "green blue"`
* `"red green blue".drop 10 = "blue"`
* `"red green blue".drop 50 = ""`
-/
@[inline] def drop (s : String) (n : Nat) : String :=
(s.toSubstring.drop n).toString
String.dropRight
definition
(s : String) (n : Nat) : String
/--
Removes the specified number of characters (Unicode code points) from the end of the string.
If `n` is greater than `s.length`, returns `""`.
Examples:
* `"red green blue".dropRight 5 = "red green"`
* `"red green blue".dropRight 11 = "red"`
* `"red green blue".dropRight 50 = ""`
-/
@[inline] def dropRight (s : String) (n : Nat) : String :=
(s.toSubstring.dropRight n).toString
String.take
definition
(s : String) (n : Nat) : String
/--
Creates a new string that contains the first `n` characters (Unicode code points) of `s`.
If `n` is greater than `s.length`, returns `s`.
Examples:
* `"red green blue".take 3 = "red"`
* `"red green blue".take 1 = "r"`
* `"red green blue".take 0 = ""`
* `"red green blue".take 100 = "red green blue"`
-/
@[inline] def take (s : String) (n : Nat) : String :=
(s.toSubstring.take n).toString
String.takeRight
definition
(s : String) (n : Nat) : String
/--
Creates a new string that contains the last `n` characters (Unicode code points) of `s`.
If `n` is greater than `s.length`, returns `s`.
Examples:
* `"red green blue".takeRight 4 = "blue"`
* `"red green blue".takeRight 1 = "e"`
* `"red green blue".takeRight 0 = ""`
* `"red green blue".takeRight 100 = "red green blue"`
-/
@[inline] def takeRight (s : String) (n : Nat) : String :=
(s.toSubstring.takeRight n).toString
String.takeWhile
definition
(s : String) (p : Char → Bool) : String
/--
Creates a new string that contains the longest prefix of `s` in which `p` returns `true` for all
characters.
Examples:
* `"red green blue".takeWhile (·.isLetter) = "red"`
* `"red green blue".takeWhile (· == 'r') = "r"`
* `"red green blue".takeWhile (· != 'n') = "red gree"`
* `"red green blue".takeWhile (fun _ => true) = "red green blue"`
-/
@[inline] def takeWhile (s : String) (p : Char → Bool) : String :=
(s.toSubstring.takeWhile p).toString
String.dropWhile
definition
(s : String) (p : Char → Bool) : String
/--
Creates a new string by removing the longest prefix from `s` in which `p` returns `true` for all
characters.
Examples:
* `"red green blue".dropWhile (·.isLetter) = " green blue"`
* `"red green blue".dropWhile (· == 'r') = "ed green blue"`
* `"red green blue".dropWhile (· != 'n') = "n blue"`
* `"red green blue".dropWhile (fun _ => true) = ""`
-/
@[inline] def dropWhile (s : String) (p : Char → Bool) : String :=
(s.toSubstring.dropWhile p).toString
String.takeRightWhile
definition
(s : String) (p : Char → Bool) : String
/--
Creates a new string that contains the longest suffix of `s` in which `p` returns `true` for all
characters.
Examples:
* `"red green blue".takeRightWhile (·.isLetter) = "blue"`
* `"red green blue".takeRightWhile (· == 'e') = "e"`
* `"red green blue".takeRightWhile (· != 'n') = " blue"`
* `"red green blue".takeRightWhile (fun _ => true) = "red green blue"`
-/
@[inline] def takeRightWhile (s : String) (p : Char → Bool) : String :=
(s.toSubstring.takeRightWhile p).toString
String.dropRightWhile
definition
(s : String) (p : Char → Bool) : String
/--
Creates a new string by removing the longest suffix from `s` in which `p` returns `true` for all
characters.
Examples:
* `"red green blue".dropRightWhile (·.isLetter) = "red green "`
* `"red green blue".dropRightWhile (· == 'e') = "red green blu"`
* `"red green blue".dropRightWhile (· != 'n') = "red green"`
* `"red green blue".dropRightWhile (fun _ => true) = ""`
-/
@[inline] def dropRightWhile (s : String) (p : Char → Bool) : String :=
(s.toSubstring.dropRightWhile p).toString
String.startsWith
definition
(s pre : String) : Bool
/--
Checks whether the first string (`s`) begins with the second (`pre`).
`String.isPrefix` is a version that takes the potential prefix before the string.
Examples:
* `"red green blue".startsWith "red" = true`
* `"red green blue".startsWith "green" = false`
* `"red green blue".startsWith "" = true`
* `"red".startsWith "red" = true`
-/
@[inline] def startsWith (s pre : String) : Bool :=
s.toSubstring.take pre.length == pre.toSubstring
String.endsWith
definition
(s post : String) : Bool
/--
Checks whether the first string (`s`) ends with the second (`post`).
Examples:
* `"red green blue".endsWith "blue" = true`
* `"red green blue".endsWith "green" = false`
* `"red green blue".endsWith "" = true`
* `"red".endsWith "red" = true`
-/
@[inline] def endsWith (s post : String) : Bool :=
s.toSubstring.takeRight post.length == post.toSubstring
String.trimRight
definition
(s : String) : String
/--
Removes trailing whitespace from a string.
“Whitespace” is defined as characters for which `Char.isWhitespace` returns `true`.
Examples:
* `"abc".trimRight = "abc"`
* `" abc".trimRight = " abc"`
* `"abc \t ".trimRight = "abc"`
* `" abc ".trimRight = " abc"`
* `"abc\ndef\n".trimRight = "abc\ndef"`
-/
@[inline] def trimRight (s : String) : String :=
s.toSubstring.trimRight.toString
String.trimLeft
definition
(s : String) : String
/--
Removes leading whitespace from a string.
“Whitespace” is defined as characters for which `Char.isWhitespace` returns `true`.
Examples:
* `"abc".trimLeft = "abc"`
* `" abc".trimLeft = " abc"`
* `"abc \t ".trimLeft = "abc \t "`
* `" abc ".trimLeft = "abc "`
* `"abc\ndef\n".trimLeft = "abc\ndef\n"`
-/
@[inline] def trimLeft (s : String) : String :=
s.toSubstring.trimLeft.toString
String.trim
definition
(s : String) : String
/--
Removes leading and trailing whitespace from a string.
“Whitespace” is defined as characters for which `Char.isWhitespace` returns `true`.
Examples:
* `"abc".trim = "abc"`
* `" abc".trim = "abc"`
* `"abc \t ".trim = "abc"`
* `" abc ".trim = "abc"`
* `"abc\ndef\n".trim = "abc\ndef"`
-/
@[inline] def trim (s : String) : String :=
s.toSubstring.trim.toString
String.nextWhile
definition
(s : String) (p : Char → Bool) (i : String.Pos) : String.Pos
/--
Repeatedly increments a position in a string, as if by `String.next`, while the predicate `p`
returns `true` for the character at the position. Stops incrementing at the end of the string or
when `p` returns `false` for the current character.
Examples:
* `let s := " a "; s.get (s.nextWhile Char.isWhitespace 0) = 'a'`
* `let s := "a "; s.get (s.nextWhile Char.isWhitespace 0) = 'a'`
* `let s := "ba "; s.get (s.nextWhile Char.isWhitespace 0) = 'b'`
-/
@[inline] def nextWhile (s : String) (p : Char → Bool) (i : String.Pos) : String.Pos :=
Substring.takeWhileAux s s.endPos p i
String.nextUntil
definition
(s : String) (p : Char → Bool) (i : String.Pos) : String.Pos
/--
Repeatedly increments a position in a string, as if by `String.next`, while the predicate `p`
returns `false` for the character at the position. Stops incrementing at the end of the string or
when `p` returns `true` for the current character.
Examples:
* `let s := " a "; s.get (s.nextUntil Char.isWhitespace 0) = ' '`
* `let s := " a "; s.get (s.nextUntil Char.isLetter 0) = 'a'`
* `let s := "a "; s.get (s.nextUntil Char.isWhitespace 0) = ' '`
-/
@[inline] def nextUntil (s : String) (p : Char → Bool) (i : String.Pos) : String.Pos :=
nextWhile s (fun c => !p c) i
String.toUpper
definition
(s : String) : String
/--
Replaces each character in `s` with the result of applying `Char.toUpper` to it.
`Char.toUpper` has no effect on characters outside of the range `'a'`–`'z'`.
Examples:
* `"orange".toUpper = "ORANGE"`
* `"abc123".toUpper = "ABC123"`
-/
@[inline] def toUpper (s : String) : String :=
s.map Char.toUpper
String.toLower
definition
(s : String) : String
/--
Replaces each character in `s` with the result of applying `Char.toLower` to it.
`Char.toLower` has no effect on characters outside of the range `'A'`–`'Z'`.
Examples:
* `"ORANGE".toLower = "orange"`
* `"Orange".toLower = "orange"`
* `"ABc123".toLower = "abc123"`
-/
@[inline] def toLower (s : String) : String :=
s.map Char.toLower
String.capitalize
definition
(s : String)
/--
Replaces the first character in `s` with the result of applying `Char.toUpper` to it. Returns the
empty string if the string is empty.
`Char.toUpper` has no effect on characters outside of the range `'a'`–`'z'`.
Examples:
* `"orange".capitalize = "Orange"`
* `"ORANGE".capitalize = "ORANGE"`
* `"".capitalize = ""`
-/
@[inline] def capitalize (s : String) :=
s.set 0 <| s.get 0 |>.toUpper
String.decapitalize
definition
(s : String)
/--
Replaces the first character in `s` with the result of applying `Char.toLower` to it. Returns the
empty string if the string is empty.
`Char.toLower` has no effect on characters outside of the range `'A'`–`'Z'`.
Examples:
* `"Orange".decapitalize = "orange"`
* `"ORANGE".decapitalize = "oRANGE"`
* `"".decapitalize = ""`
-/
@[inline] def decapitalize (s : String) :=
s.set 0 <| s.get 0 |>.toLower
String.dropPrefix?
definition
(s : String) (pre : String) : Option Substring
/--
If `pre` is a prefix of `s`, returns the remainder. Returns `none` otherwise.
The string `pre` is a prefix of `s` if there exists a `t : String` such that `s = pre ++ t`. If so,
the result is `some t`.
Use `String.stripPrefix` to return the string unchanged when `pre` is not a prefix.
Examples:
* `"red green blue".dropPrefix? "red " = some "green blue"`
* `"red green blue".dropPrefix? "reed " = none`
* `"red green blue".dropPrefix? "" = some "red green blue"`
-/
def dropPrefix? (s : String) (pre : String) : Option Substring :=
s.toSubstring.dropPrefix? pre.toSubstring
String.dropSuffix?
definition
(s : String) (suff : String) : Option Substring
/--
If `suff` is a suffix of `s`, returns the remainder. Returns `none` otherwise.
The string `suff` is a suffix of `s` if there exists a `t : String` such that `s = t ++ suff`. If so,
the result is `some t`.
Use `String.stripSuffix` to return the string unchanged when `suff` is not a suffix.
Examples:
* `"red green blue".dropSuffix? " blue" = some "red green"`
* `"red green blue".dropSuffix? " blu " = none`
* `"red green blue".dropSuffix? "" = some "red green blue"`
-/
def dropSuffix? (s : String) (suff : String) : Option Substring :=
s.toSubstring.dropSuffix? suff.toSubstring
String.stripPrefix
definition
(s : String) (pre : String) : String
/--
If `pre` is a prefix of `s`, returns the remainder. Returns `s` unmodified otherwise.
The string `pre` is a prefix of `s` if there exists a `t : String` such that `s = pre ++ t`. If so,
the result is `t`. Otherwise, it is `s`.
Use `String.dropPrefix?` to return `none` when `pre` is not a prefix.
Examples:
* `"red green blue".stripPrefix "red " = "green blue"`
* `"red green blue".stripPrefix "reed " = "red green blue"`
* `"red green blue".stripPrefix "" = "red green blue"`
-/
def stripPrefix (s : String) (pre : String) : String :=
s.dropPrefix? pre |>.map Substring.toString |>.getD s
String.stripSuffix
definition
(s : String) (suff : String) : String
/--
If `suff` is a suffix of `s`, returns the remainder. Returns `s` unmodified otherwise.
The string `suff` is a suffix of `s` if there exists a `t : String` such that `s = t ++ suff`. If so,
the result is `t`. Otherwise, it is `s`.
Use `String.dropSuffix?` to return `none` when `suff` is not a suffix.
Examples:
* `"red green blue".stripSuffix " blue" = "red green"`
* `"red green blue".stripSuffix " blu " = "red green blue"`
* `"red green blue".stripSuffix "" = "red green blue"`
-/
def stripSuffix (s : String) (suff : String) : String :=
s.dropSuffix? suff |>.map Substring.toString |>.getD s
Char.toString
definition
(c : Char) : String
/--
Constructs a singleton string that contains only the provided character.
Examples:
* `'L'.toString = "L"`
* `'"'.toString = "\""`
-/
@[inline] protected def toString (c : Char) : String :=
String.singleton c
Char.length_toString
theorem
(c : Char) : c.toString.length = 1
@[simp] theorem length_toString (c : Char) : c.toString.length = 1 := rfl
String.ext
theorem
{s₁ s₂ : String} (h : s₁.data = s₂.data) : s₁ = s₂
String.ext_iff
theorem
{s₁ s₂ : String} : s₁ = s₂ ↔ s₁.data = s₂.data
theorem ext_iff {s₁ s₂ : String} : s₁ = s₂ ↔ s₁.data = s₂.data := ⟨fun h => h ▸ rfl, ext⟩
String.default_eq
theorem
: default = ""
@[simp] theorem default_eq : default = "" := rfl
String.length_mk
theorem
(s : List Char) : (String.mk s).length = s.length
String.length_empty
theorem
: "".length = 0
@[simp] theorem length_empty : "".length = 0 := rfl
String.length_singleton
theorem
(c : Char) : (String.singleton c).length = 1
@[simp] theorem length_singleton (c : Char) : (String.singleton c).length = 1 := rfl
String.length_push
theorem
(c : Char) : (String.push s c).length = s.length + 1
@[simp] theorem length_push (c : Char) : (String.push s c).length = s.length + 1 := by
rw [push, length_mk, List.length_append, List.length_singleton, Nat.succ.injEq]
rfl
String.length_pushn
theorem
(c : Char) (n : Nat) : (pushn s c n).length = s.length + n
@[simp] theorem length_pushn (c : Char) (n : Nat) : (pushn s c n).length = s.length + n := by
unfold pushn; induction n <;> simp [Nat.repeat, Nat.add_assoc, *]
String.length_append
theorem
(s t : String) : (s ++ t).length = s.length + t.length
@[simp] theorem length_append (s t : String) : (s ++ t).length = s.length + t.length := by
simp only [length, append, List.length_append]
String.data_push
theorem
(s : String) (c : Char) : (s.push c).data = s.data ++ [c]
String.data_append
theorem
(s t : String) : (s ++ t).data = s.data ++ t.data
String.lt_iff
theorem
{s t : String} : s < t ↔ s.data < t.data
theorem lt_iff {s t : String} : s < t ↔ s.data < t.data := .rfl
String.Pos.byteIdx_zero
theorem
: (0 : Pos).byteIdx = 0
@[simp] theorem byteIdx_zero : (0 : Pos).byteIdx = 0 := rfl
String.Pos.byteIdx_mk
theorem
(n : Nat) : byteIdx ⟨n⟩ = n
theorem byteIdx_mk (n : Nat) : byteIdx ⟨n⟩ = n := rfl
String.Pos.mk_zero
theorem
: ⟨0⟩ = (0 : Pos)
String.Pos.mk_byteIdx
theorem
(p : Pos) : ⟨p.byteIdx⟩ = p
@[simp] theorem mk_byteIdx (p : Pos) : ⟨p.byteIdx⟩ = p := rfl
String.Pos.ext
theorem
{i₁ i₂ : Pos} (h : i₁.byteIdx = i₂.byteIdx) : i₁ = i₂
theorem ext {i₁ i₂ : Pos} (h : i₁.byteIdx = i₂.byteIdx) : i₁ = i₂ :=
show ⟨i₁.byteIdx⟩ = (⟨i₂.byteIdx⟩ : Pos) from h ▸ rfl
String.Pos.ext_iff
theorem
{i₁ i₂ : Pos} : i₁ = i₂ ↔ i₁.byteIdx = i₂.byteIdx
theorem ext_iff {i₁ i₂ : Pos} : i₁ = i₂ ↔ i₁.byteIdx = i₂.byteIdx := ⟨fun h => h ▸ rfl, ext⟩
String.Pos.add_byteIdx
theorem
(p₁ p₂ : Pos) : (p₁ + p₂).byteIdx = p₁.byteIdx + p₂.byteIdx
String.Pos.add_eq
theorem
(p₁ p₂ : Pos) : p₁ + p₂ = ⟨p₁.byteIdx + p₂.byteIdx⟩
theorem add_eq (p₁ p₂ : Pos) : p₁ + p₂ = ⟨p₁.byteIdx + p₂.byteIdx⟩ := rfl
String.Pos.sub_byteIdx
theorem
(p₁ p₂ : Pos) : (p₁ - p₂).byteIdx = p₁.byteIdx - p₂.byteIdx
String.Pos.sub_eq
theorem
(p₁ p₂ : Pos) : p₁ - p₂ = ⟨p₁.byteIdx - p₂.byteIdx⟩
theorem sub_eq (p₁ p₂ : Pos) : p₁ - p₂ = ⟨p₁.byteIdx - p₂.byteIdx⟩ := rfl
String.Pos.addChar_byteIdx
theorem
(p : Pos) (c : Char) : (p + c).byteIdx = p.byteIdx + c.utf8Size
String.Pos.addChar_eq
theorem
(p : Pos) (c : Char) : p + c = ⟨p.byteIdx + c.utf8Size⟩
theorem addChar_eq (p : Pos) (c : Char) : p + c = ⟨p.byteIdx + c.utf8Size⟩ := rfl
String.Pos.zero_addChar_byteIdx
theorem
(c : Char) : ((0 : Pos) + c).byteIdx = c.utf8Size
theorem zero_addChar_byteIdx (c : Char) : ((0 : Pos) + c).byteIdx = c.utf8Size := by
simp only [addChar_byteIdx, byteIdx_zero, Nat.zero_add]
String.Pos.zero_addChar_eq
theorem
(c : Char) : (0 : Pos) + c = ⟨c.utf8Size⟩
theorem zero_addChar_eq (c : Char) : (0 : Pos) + c = ⟨c.utf8Size⟩ := by rw [← zero_addChar_byteIdx]
String.Pos.addChar_right_comm
theorem
(p : Pos) (c₁ c₂ : Char) : p + c₁ + c₂ = p + c₂ + c₁
theorem addChar_right_comm (p : Pos) (c₁ c₂ : Char) : p + c₁ + c₂ = p + c₂ + c₁ := by
apply ext
repeat rw [pos_add_char]
apply Nat.add_right_comm
String.Pos.ne_of_lt
theorem
{i₁ i₂ : Pos} (h : i₁ < i₂) : i₁ ≠ i₂
String.Pos.ne_of_gt
theorem
{i₁ i₂ : Pos} (h : i₁ < i₂) : i₂ ≠ i₁
String.Pos.byteIdx_addString
theorem
(p : Pos) (s : String) : (p + s).byteIdx = p.byteIdx + s.utf8ByteSize
@[simp] theorem byteIdx_addString (p : Pos) (s : String) :
(p + s).byteIdx = p.byteIdx + s.utf8ByteSize := rfl
String.Pos.addString_byteIdx
abbrev
@[deprecated byteIdx_addString (since := "2025-03-18")]
abbrev addString_byteIdx := @byteIdx_addString
String.Pos.addString_eq
theorem
(p : Pos) (s : String) : p + s = ⟨p.byteIdx + s.utf8ByteSize⟩
theorem addString_eq (p : Pos) (s : String) : p + s = ⟨p.byteIdx + s.utf8ByteSize⟩ := rfl
String.Pos.byteIdx_zero_addString
theorem
(s : String) : ((0 : Pos) + s).byteIdx = s.utf8ByteSize
theorem byteIdx_zero_addString (s : String) : ((0 : Pos) + s).byteIdx = s.utf8ByteSize := by
simp only [byteIdx_addString, byteIdx_zero, Nat.zero_add]
String.Pos.zero_addString_byteIdx
abbrev
@[deprecated byteIdx_zero_addString (since := "2025-03-18")]
abbrev zero_addString_byteIdx := @byteIdx_zero_addString
String.Pos.zero_addString_eq
theorem
(s : String) : (0 : Pos) + s = ⟨s.utf8ByteSize⟩
theorem zero_addString_eq (s : String) : (0 : Pos) + s = ⟨s.utf8ByteSize⟩ := by
rw [← byteIdx_zero_addString]
String.Pos.le_iff
theorem
{i₁ i₂ : Pos} : i₁ ≤ i₂ ↔ i₁.byteIdx ≤ i₂.byteIdx
theorem le_iff {i₁ i₂ : Pos} : i₁ ≤ i₂ ↔ i₁.byteIdx ≤ i₂.byteIdx := .rfl
String.Pos.mk_le_mk
theorem
{i₁ i₂ : Nat} : Pos.mk i₁ ≤ Pos.mk i₂ ↔ i₁ ≤ i₂
@[simp] theorem mk_le_mk {i₁ i₂ : Nat} : Pos.mkPos.mk i₁ ≤ Pos.mk i₂ ↔ i₁ ≤ i₂ := .rfl
String.Pos.lt_iff
theorem
{i₁ i₂ : Pos} : i₁ < i₂ ↔ i₁.byteIdx < i₂.byteIdx
theorem lt_iff {i₁ i₂ : Pos} : i₁ < i₂ ↔ i₁.byteIdx < i₂.byteIdx := .rfl
String.Pos.mk_lt_mk
theorem
{i₁ i₂ : Nat} : Pos.mk i₁ < Pos.mk i₂ ↔ i₁ < i₂
@[simp] theorem mk_lt_mk {i₁ i₂ : Nat} : Pos.mkPos.mk i₁ < Pos.mk i₂ ↔ i₁ < i₂ := .rfl
String.get!_eq_get
theorem
(s : String) (p : Pos) : get! s p = get s p
String.lt_next'
theorem
(s : String) (p : Pos) : p < next s p
String.prev_zero
theorem
(s : String) : prev s 0 = 0
String.get'_eq
theorem
(s : String) (p : Pos) (h) : get' s p h = get s p
String.next'_eq
theorem
(s : String) (p : Pos) (h) : next' s p h = next s p
String.singleton_eq
theorem
(c : Char) : singleton c = ⟨[c]⟩
theorem singleton_eq (c : Char) : singleton c = ⟨[c]⟩ := rfl
String.data_singleton
theorem
(c : Char) : (singleton c).data = [c]
String.append_empty
theorem
(s : String) : s ++ "" = s
@[simp] theorem append_empty (s : String) : s ++ "" = s := ext (List.append_nil _)
String.empty_append
theorem
(s : String) : "" ++ s = s
@[simp] theorem empty_append (s : String) : "" ++ s = s := rfl
String.append_assoc
theorem
(s₁ s₂ s₃ : String) : (s₁ ++ s₂) ++ s₃ = s₁ ++ (s₂ ++ s₃)
theorem append_assoc (s₁ s₂ s₃ : String) : (s₁ ++ s₂) ++ s₃ = s₁ ++ (s₂ ++ s₃) :=
ext (List.append_assoc ..)
Substring.prev_zero
theorem
(s : Substring) : s.prev 0 = 0
@[simp] theorem prev_zero (s : Substring) : s.prev 0 = 0 := by simp [prev, Pos.add_eq, Pos.byteIdx_zero]
Substring.prevn_zero
theorem
(s : Substring) : ∀ n, s.prevn n 0 = 0
@[simp] theorem prevn_zero (s : Substring) : ∀ n, s.prevn n 0 = 0
| 0 => rfl
| n+1 => by simp [prevn, prevn_zero s n]