Documentation

Init.Data.String.Basic

Equations
  • s.asString = { data := s }
Instances For
    Equations
    Equations
    @[extern lean_string_dec_lt]
    instance String.decLt (s₁ : String) (s₂ : String) :
    Decidable (s₁ < s₂)
    Equations
    • s₁.decLt s₂ = s₁.data.hasDecidableLt s₂.data
    @[extern lean_string_length]
    Equations
    • x.length = match x with | { data := s } => s.length
    Instances For
      @[extern lean_string_push]

      The internal implementation uses dynamic arrays and will perform destructive updates if the String is not shared.

      Equations
      • x✝.push x = match x✝, x with | { data := s }, c => { data := s ++ [c] }
      Instances For
        @[extern lean_string_append]

        The internal implementation uses dynamic arrays and will perform destructive updates if the String is not shared.

        Equations
        • x✝.append x = match x✝, x with | { data := a }, { data := b } => { data := a ++ b }
        Instances For

          O(n) in the runtime, where n is the length of the String

          Equations
          • s.toList = s.data
          Instances For
            @[extern lean_string_is_valid_pos]

            Returns true if p is a valid UTF-8 position in the string s, meaning that p ≤ s.endPos and p lies on a UTF-8 character boundary. This has an O(1) implementation in the runtime.

            Equations
            Instances For
              Equations
              Instances For
                Equations
                Instances For
                  @[extern lean_string_utf8_get]
                  def String.get (s : String) (p : String.Pos) :

                  Return character at position p. If p is not a valid position returns (default : Char). See utf8GetAux for the reference implementation.

                  Equations
                  Instances For
                    Equations
                    Instances For
                      @[extern lean_string_utf8_get_opt]
                      Equations
                      Instances For
                        @[extern lean_string_utf8_get_bang]

                        Similar to get, but produces a panic error message if p is not a valid String.Pos.

                        Equations
                        Instances For
                          Equations
                          Instances For
                            @[extern lean_string_utf8_set]
                            Equations
                            • x✝¹.set x✝ x = match x✝¹, x✝, x with | { data := s }, i, c => { data := String.utf8SetAux c s 0 i }
                            Instances For
                              def String.modify (s : String) (i : String.Pos) (f : CharChar) :
                              Equations
                              • s.modify i f = s.set i (f (s.get i))
                              Instances For
                                @[extern lean_string_utf8_next]
                                Equations
                                • s.next p = let c := s.get p; p + c
                                Instances For
                                  Equations
                                  Instances For
                                    @[extern lean_string_utf8_prev]
                                    Equations
                                    • x✝.prev x = match x✝, x with | { data := s }, p => if p = 0 then 0 else String.utf8PrevAux s 0 p
                                    Instances For
                                      Equations
                                      • s.front = s.get 0
                                      Instances For
                                        Equations
                                        • s.back = s.get (s.prev s.endPos)
                                        Instances For
                                          @[extern lean_string_utf8_at_end]
                                          Equations
                                          • x✝.atEnd x = match x✝, x with | s, p => decide (p.byteIdx s.utf8ByteSize)
                                          Instances For
                                            @[extern lean_string_utf8_get_fast]
                                            def String.get' (s : String) (p : String.Pos) (h : ¬s.atEnd p = true) :

                                            Similar to get but runtime does not perform bounds check.

                                            Equations
                                            Instances For
                                              @[extern lean_string_utf8_next_fast]
                                              def String.next' (s : String) (p : String.Pos) (h : ¬s.atEnd p = true) :

                                              Similar to next but runtime does not perform bounds check.

                                              Equations
                                              • s.next' p h = let c := s.get p; p + c
                                              Instances For
                                                @[simp]
                                                theorem String.pos_lt_eq (p₁ : String.Pos) (p₂ : String.Pos) :
                                                (p₁ < p₂) = (p₁.byteIdx < p₂.byteIdx)
                                                @[simp]
                                                theorem String.pos_add_char (p : String.Pos) (c : Char) :
                                                (p + c).byteIdx = p.byteIdx + String.csize c
                                                theorem String.lt_next (s : String) (i : String.Pos) :
                                                i.byteIdx < (s.next i).byteIdx
                                                theorem String.utf8PrevAux_lt_of_pos (cs : List Char) (i : String.Pos) (p : String.Pos) :
                                                p 0(String.utf8PrevAux cs i p).byteIdx < p.byteIdx
                                                theorem String.prev_lt_of_pos (s : String) (i : String.Pos) (h : i 0) :
                                                (s.prev i).byteIdx < i.byteIdx
                                                def String.posOfAux (s : String) (c : Char) (stopPos : String.Pos) (pos : String.Pos) :
                                                Equations
                                                • s.posOfAux c stopPos pos = if h : pos < stopPos then if (s.get pos == c) = true then pos else let_fun this := ; s.posOfAux c stopPos (s.next pos) else pos
                                                Instances For
                                                  @[inline]
                                                  Equations
                                                  • s.posOf c = s.posOfAux c s.endPos 0
                                                  Instances For
                                                    Equations
                                                    • s.revPosOfAux c pos = if h : pos = 0 then none else let_fun this := ; let pos := s.prev pos; if (s.get pos == c) = true then some pos else s.revPosOfAux c pos
                                                    Instances For
                                                      Equations
                                                      • s.revPosOf c = s.revPosOfAux c s.endPos
                                                      Instances For
                                                        def String.findAux (s : String) (p : CharBool) (stopPos : String.Pos) (pos : String.Pos) :
                                                        Equations
                                                        • s.findAux p stopPos pos = if h : pos < stopPos then if p (s.get pos) = true then pos else let_fun this := ; s.findAux p stopPos (s.next pos) else pos
                                                        Instances For
                                                          @[inline]
                                                          def String.find (s : String) (p : CharBool) :
                                                          Equations
                                                          • s.find p = s.findAux p s.endPos 0
                                                          Instances For
                                                            Equations
                                                            • s.revFindAux p pos = if h : pos = 0 then none else let_fun this := ; let pos := s.prev pos; if p (s.get pos) = true then some pos else s.revFindAux p pos
                                                            Instances For
                                                              Equations
                                                              • s.revFind p = s.revFindAux p s.endPos
                                                              Instances For
                                                                @[reducible, inline]
                                                                Equations
                                                                • p₁.min p₂ = { byteIdx := p₁.byteIdx.min p₂.byteIdx }
                                                                Instances For

                                                                  Returns the first position where the two strings differ.

                                                                  Equations
                                                                  Instances For
                                                                    Equations
                                                                    • One or more equations did not get rendered due to their size.
                                                                    Instances For
                                                                      @[extern lean_string_utf8_extract]
                                                                      Equations
                                                                      • x✝¹.extract x✝ x = match x✝¹, x✝, x with | { data := s }, b, e => if b.byteIdx e.byteIdx then "" else { data := String.extract.go₁ s 0 b e }
                                                                      Instances For
                                                                        Equations
                                                                        Instances For
                                                                          Equations
                                                                          Instances For
                                                                            @[specialize #[]]
                                                                            def String.splitAux (s : String) (p : CharBool) (b : String.Pos) (i : String.Pos) (r : List String) :
                                                                            Equations
                                                                            • One or more equations did not get rendered due to their size.
                                                                            Instances For
                                                                              @[specialize #[]]
                                                                              def String.split (s : String) (p : CharBool) :
                                                                              Equations
                                                                              • s.split p = s.splitAux p 0 0 []
                                                                              Instances For

                                                                                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.

                                                                                Equations
                                                                                • One or more equations did not get rendered due to their size.
                                                                                Instances For

                                                                                  Splits a string s on occurrences of the separator sep. When sep is empty, it returns [s]; when sep occurs in overlapping patterns, the first match is taken. There will always be exactly n+1 elements in the returned list if there were n nonoverlapping matches of sep in the string. The default separator is " ". The separators are not included in the returned substrings.

                                                                                  "here is some text ".splitOn = ["here", "is", "some", "text", ""]
                                                                                  "here is some text ".splitOn "some" = ["here is ", " text "]
                                                                                  "here is some text ".splitOn "" = ["here is some text "]
                                                                                  "ababacabac".splitOn "aba" = ["", "bac", "c"]
                                                                                  
                                                                                  Equations
                                                                                  • s.splitOn sep = if (sep == "") = true then [s] else s.splitOnAux sep 0 0 0 []
                                                                                  Instances For
                                                                                    Equations
                                                                                    Instances For
                                                                                      def String.pushn (s : String) (c : Char) (n : Nat) :
                                                                                      Equations
                                                                                      Instances For
                                                                                        Equations
                                                                                        • s.isEmpty = (s.endPos == 0)
                                                                                        Instances For
                                                                                          Equations
                                                                                          Instances For
                                                                                            Equations
                                                                                            Instances For
                                                                                              Equations
                                                                                              Instances For
                                                                                                Equations
                                                                                                Instances For

                                                                                                  Iterator over the characters (Char) of a String.

                                                                                                  Typically created by s.iter, where s is a String.

                                                                                                  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 arbitrary values if the iterator is not valid. The functions in the String.Iterator API should rule out the creation of invalid iterators, with two exceptions:

                                                                                                  • s : String

                                                                                                    The string the iterator is for.

                                                                                                  • The current position.

                                                                                                    This position is not necessarily valid for the string, for instance if one keeps calling Iterator.next when Iterator.atEnd is true. If the position is not valid, then the current character is (default : Char), similar to String.get on an invalid position.

                                                                                                  Instances For

                                                                                                    Creates an iterator at the beginning of a string.

                                                                                                    Equations
                                                                                                    • s.mkIterator = { s := s, i := 0 }
                                                                                                    Instances For
                                                                                                      @[reducible, inline]

                                                                                                      Creates an iterator at the beginning of a string.

                                                                                                      Equations
                                                                                                      Instances For

                                                                                                        The size of a string iterator is the number of bytes remaining.

                                                                                                        Equations
                                                                                                        theorem String.Iterator.sizeOf_eq (i : String.Iterator) :
                                                                                                        sizeOf i = i.s.utf8ByteSize - i.i.byteIdx

                                                                                                        The string the iterator is for.

                                                                                                        Equations
                                                                                                        Instances For

                                                                                                          Number of bytes remaining in the iterator.

                                                                                                          Equations
                                                                                                          • x.remainingBytes = match x with | { s := s, i := i } => s.endPos.byteIdx - i.byteIdx
                                                                                                          Instances For

                                                                                                            The current position.

                                                                                                            This position is not necessarily valid for the string, for instance if one keeps calling Iterator.next when Iterator.atEnd is true. If the position is not valid, then the current character is (default : Char), similar to String.get on an invalid position.

                                                                                                            Equations
                                                                                                            Instances For

                                                                                                              The character at the current position.

                                                                                                              On an invalid position, returns (default : Char).

                                                                                                              Equations
                                                                                                              • x.curr = match x with | { s := s, i := i } => s.get i
                                                                                                              Instances For

                                                                                                                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. Iterator.atEnd is false; otherwise, the resulting iterator will be invalid.

                                                                                                                Equations
                                                                                                                • x.next = match x with | { s := s, i := i } => { s := s, i := s.next i }
                                                                                                                Instances For

                                                                                                                  Decreases the iterator's position.

                                                                                                                  If the position is zero, this function is the identity.

                                                                                                                  Equations
                                                                                                                  • x.prev = match x with | { s := s, i := i } => { s := s, i := s.prev i }
                                                                                                                  Instances For

                                                                                                                    True if the iterator is past the string's last character.

                                                                                                                    Equations
                                                                                                                    • x.atEnd = match x with | { s := s, i := i } => decide (i.byteIdx s.endPos.byteIdx)
                                                                                                                    Instances For

                                                                                                                      True if the iterator is not past the string's last character.

                                                                                                                      Equations
                                                                                                                      • x.hasNext = match x with | { s := s, i := i } => decide (i.byteIdx < s.endPos.byteIdx)
                                                                                                                      Instances For

                                                                                                                        True if the position is not zero.

                                                                                                                        Equations
                                                                                                                        • x.hasPrev = match x with | { s := s, i := i } => decide (i.byteIdx > 0)
                                                                                                                        Instances For

                                                                                                                          Replaces the current character in the string.

                                                                                                                          Does nothing if the iterator is at the end of the string. If the iterator contains the only reference to its string, this function will mutate the string in-place instead of allocating a new one.

                                                                                                                          Equations
                                                                                                                          • x✝.setCurr x = match x✝, x with | { s := s, i := i }, c => { s := s.set i c, i := i }
                                                                                                                          Instances For

                                                                                                                            Moves the iterator's position to the end of the string.

                                                                                                                            Note that i.toEnd.atEnd is always true.

                                                                                                                            Equations
                                                                                                                            • x.toEnd = match x with | { s := s, i := i } => { s := s, i := s.endPos }
                                                                                                                            Instances For

                                                                                                                              Extracts the substring between the positions of two iterators.

                                                                                                                              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.

                                                                                                                              Equations
                                                                                                                              • x✝.extract x = match x✝, x with | { s := s₁, i := b }, { s := s₂, i := e } => if (decide (s₁ s₂) || decide (b > e)) = true then "" else s₁.extract b e
                                                                                                                              Instances For

                                                                                                                                Moves the iterator's position several characters forward.

                                                                                                                                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.

                                                                                                                                Equations
                                                                                                                                • x.forward 0 = x
                                                                                                                                • x.forward n.succ = x.next.forward n
                                                                                                                                Instances For

                                                                                                                                  The remaining characters in an iterator, as a string.

                                                                                                                                  Equations
                                                                                                                                  • x.remainingToString = match x with | { s := s, i := i } => s.extract i s.endPos
                                                                                                                                  Instances For

                                                                                                                                    Moves the iterator's position several characters forward.

                                                                                                                                    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.

                                                                                                                                    Equations
                                                                                                                                    • x.nextn 0 = x
                                                                                                                                    • x.nextn n.succ = x.next.nextn n
                                                                                                                                    Instances For

                                                                                                                                      Moves the iterator's position several characters back.

                                                                                                                                      If asked to go back more characters than available, stops at the beginning of the string.

                                                                                                                                      Equations
                                                                                                                                      • x.prevn 0 = x
                                                                                                                                      • x.prevn n.succ = x.prev.prevn n
                                                                                                                                      Instances For
                                                                                                                                        def String.offsetOfPosAux (s : String) (pos : String.Pos) (i : String.Pos) (offset : Nat) :
                                                                                                                                        Equations
                                                                                                                                        • s.offsetOfPosAux pos i offset = if i pos then offset else if h : s.atEnd i = true then offset else let_fun this := ; s.offsetOfPosAux pos (s.next i) (offset + 1)
                                                                                                                                        Instances For
                                                                                                                                          Equations
                                                                                                                                          • s.offsetOfPos pos = s.offsetOfPosAux pos 0 0
                                                                                                                                          Instances For
                                                                                                                                            @[specialize #[]]
                                                                                                                                            def String.foldlAux {α : Type u} (f : αCharα) (s : String) (stopPos : String.Pos) (i : String.Pos) (a : α) :
                                                                                                                                            α
                                                                                                                                            Equations
                                                                                                                                            Instances For
                                                                                                                                              @[inline]
                                                                                                                                              def String.foldl {α : Type u} (f : αCharα) (init : α) (s : String) :
                                                                                                                                              α
                                                                                                                                              Equations
                                                                                                                                              Instances For
                                                                                                                                                @[specialize #[]]
                                                                                                                                                def String.foldrAux {α : Type u} (f : Charαα) (a : α) (s : String) (i : String.Pos) (begPos : String.Pos) :
                                                                                                                                                α
                                                                                                                                                Equations
                                                                                                                                                Instances For
                                                                                                                                                  @[inline]
                                                                                                                                                  def String.foldr {α : Type u} (f : Charαα) (init : α) (s : String) :
                                                                                                                                                  α
                                                                                                                                                  Equations
                                                                                                                                                  Instances For
                                                                                                                                                    @[specialize #[]]
                                                                                                                                                    def String.anyAux (s : String) (stopPos : String.Pos) (p : CharBool) (i : String.Pos) :
                                                                                                                                                    Equations
                                                                                                                                                    • s.anyAux stopPos p i = if h : i < stopPos then if p (s.get i) = true then true else let_fun this := ; s.anyAux stopPos p (s.next i) else false
                                                                                                                                                    Instances For
                                                                                                                                                      @[inline]
                                                                                                                                                      def String.any (s : String) (p : CharBool) :
                                                                                                                                                      Equations
                                                                                                                                                      • s.any p = s.anyAux s.endPos p 0
                                                                                                                                                      Instances For
                                                                                                                                                        @[inline]
                                                                                                                                                        def String.all (s : String) (p : CharBool) :
                                                                                                                                                        Equations
                                                                                                                                                        • s.all p = !s.any fun (c : Char) => !p c
                                                                                                                                                        Instances For
                                                                                                                                                          def String.contains (s : String) (c : Char) :
                                                                                                                                                          Equations
                                                                                                                                                          • s.contains c = s.any fun (a : Char) => a == c
                                                                                                                                                          Instances For
                                                                                                                                                            theorem String.utf8SetAux_of_gt (c' : Char) (cs : List Char) {i : String.Pos} {p : String.Pos} :
                                                                                                                                                            i > pString.utf8SetAux c' cs i p = cs
                                                                                                                                                            theorem String.set_next_add (s : String) (i : String.Pos) (c : Char) (b₁ : Nat) (b₂ : Nat) (h : (s.next i).byteIdx + b₁ = s.endPos.byteIdx + b₂) :
                                                                                                                                                            ((s.set i c).next i).byteIdx + b₁ = (s.set i c).endPos.byteIdx + b₂
                                                                                                                                                            theorem String.mapAux_lemma (s : String) (i : String.Pos) (c : Char) (h : ¬s.atEnd i = true) :
                                                                                                                                                            (s.set i c).endPos.byteIdx - ((s.set i c).next i).byteIdx < s.endPos.byteIdx - i.byteIdx
                                                                                                                                                            @[specialize #[]]
                                                                                                                                                            def String.mapAux (f : CharChar) (i : String.Pos) (s : String) :
                                                                                                                                                            Equations
                                                                                                                                                            Instances For
                                                                                                                                                              @[inline]
                                                                                                                                                              def String.map (f : CharChar) (s : String) :
                                                                                                                                                              Equations
                                                                                                                                                              Instances For
                                                                                                                                                                Equations
                                                                                                                                                                • s.isNat = (!s.isEmpty && s.all fun (x : Char) => x.isDigit)
                                                                                                                                                                Instances For
                                                                                                                                                                  Equations
                                                                                                                                                                  Instances For
                                                                                                                                                                    def String.substrEq (s1 : String) (off1 : String.Pos) (s2 : String) (off2 : String.Pos) (sz : Nat) :

                                                                                                                                                                    Return true iff the substring of byte size sz starting at position off1 in s1 is equal to that starting at off2 in s2.. False if either substring of that byte size does not exist.

                                                                                                                                                                    Equations
                                                                                                                                                                    • One or more equations did not get rendered due to their size.
                                                                                                                                                                    Instances For
                                                                                                                                                                      def String.substrEq.loop (s1 : String) (s2 : String) (off1 : String.Pos) (off2 : String.Pos) (stop1 : String.Pos) :
                                                                                                                                                                      Equations
                                                                                                                                                                      • One or more equations did not get rendered due to their size.
                                                                                                                                                                      Instances For

                                                                                                                                                                        Return true iff p is a prefix of s

                                                                                                                                                                        Equations
                                                                                                                                                                        • p.isPrefixOf s = p.substrEq 0 s 0 p.endPos.byteIdx
                                                                                                                                                                        Instances For
                                                                                                                                                                          def String.replace (s : String) (pattern : String) (replacement : String) :

                                                                                                                                                                          Replace all occurrences of pattern in s with replacement.

                                                                                                                                                                          Equations
                                                                                                                                                                          • s.replace pattern replacement = if h : pattern.endPos.byteIdx = 0 then s else let_fun hPatt := ; String.replace.loop s pattern replacement hPatt "" 0 0
                                                                                                                                                                          Instances For
                                                                                                                                                                            def String.replace.loop (s : String) (pattern : String) (replacement : String) (hPatt : 0 < pattern.endPos.byteIdx) (acc : String) (accStop : String.Pos) (pos : String.Pos) :
                                                                                                                                                                            Equations
                                                                                                                                                                            • One or more equations did not get rendered due to their size.
                                                                                                                                                                            Instances For

                                                                                                                                                                              Return the beginning of the line that contains character pos.

                                                                                                                                                                              Equations
                                                                                                                                                                              • s.findLineStart pos = match s.revFindAux (fun (x : Char) => decide (x = '\n')) pos with | none => 0 | some n => { byteIdx := n.byteIdx + 1 }
                                                                                                                                                                              Instances For
                                                                                                                                                                                @[inline]
                                                                                                                                                                                Equations
                                                                                                                                                                                • ss.isEmpty = (ss.bsize == 0)
                                                                                                                                                                                Instances For
                                                                                                                                                                                  @[inline]
                                                                                                                                                                                  Equations
                                                                                                                                                                                  • x.toString = match x with | { str := s, startPos := b, stopPos := e } => s.extract b e
                                                                                                                                                                                  Instances For
                                                                                                                                                                                    @[inline]
                                                                                                                                                                                    Equations
                                                                                                                                                                                    • x.toIterator = match x with | { str := s, startPos := b, stopPos := stopPos } => { s := s, i := b }
                                                                                                                                                                                    Instances For
                                                                                                                                                                                      @[inline]

                                                                                                                                                                                      Return the codepoint at the given offset into the substring.

                                                                                                                                                                                      Equations
                                                                                                                                                                                      • x✝.get x = match x✝, x with | { str := s, startPos := b, stopPos := stopPos }, p => s.get (b + p)
                                                                                                                                                                                      Instances For
                                                                                                                                                                                        @[inline]

                                                                                                                                                                                        Given an offset of a codepoint into the substring, return the offset there of the next codepoint.

                                                                                                                                                                                        Equations
                                                                                                                                                                                        • x✝.next x = match x✝, x with | { str := s, startPos := b, stopPos := e }, p => let absP := b + p; if absP = e then p else { byteIdx := (s.next absP).byteIdx - b.byteIdx }
                                                                                                                                                                                        Instances For
                                                                                                                                                                                          theorem Substring.lt_next (s : Substring) (i : String.Pos) (h : i.byteIdx < s.bsize) :
                                                                                                                                                                                          i.byteIdx < (s.next i).byteIdx
                                                                                                                                                                                          @[inline]

                                                                                                                                                                                          Given an offset of a codepoint into the substring, return the offset there of the previous codepoint.

                                                                                                                                                                                          Equations
                                                                                                                                                                                          • x✝.prev x = match x✝, x with | { str := s, startPos := b, stopPos := stopPos }, p => let absP := b + p; if absP = b then p else { byteIdx := (s.prev absP).byteIdx - b.byteIdx }
                                                                                                                                                                                          Instances For
                                                                                                                                                                                            Equations
                                                                                                                                                                                            • x✝.nextn 0 x = x
                                                                                                                                                                                            • x✝.nextn i.succ x = x✝.nextn i (x✝.next x)
                                                                                                                                                                                            Instances For
                                                                                                                                                                                              Equations
                                                                                                                                                                                              • x✝.prevn 0 x = x
                                                                                                                                                                                              • x✝.prevn i.succ x = x✝.prevn i (x✝.prev x)
                                                                                                                                                                                              Instances For
                                                                                                                                                                                                @[inline]
                                                                                                                                                                                                Equations
                                                                                                                                                                                                • s.front = s.get 0
                                                                                                                                                                                                Instances For
                                                                                                                                                                                                  @[inline]

                                                                                                                                                                                                  Return the offset into s of the first occurrence of c in s, or s.bsize if c doesn't occur.

                                                                                                                                                                                                  Equations
                                                                                                                                                                                                  • s.posOf c = match s with | { str := s, startPos := b, stopPos := e } => { byteIdx := (s.posOfAux c e b).byteIdx - b.byteIdx }
                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                    @[inline]
                                                                                                                                                                                                    Equations
                                                                                                                                                                                                    • x✝.drop x = match x✝, x with | ss@h:{ str := s, startPos := b, stopPos := e }, n => { str := s, startPos := b + ss.nextn n 0, stopPos := e }
                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                      @[inline]
                                                                                                                                                                                                      Equations
                                                                                                                                                                                                      • x✝.dropRight x = match x✝, x with | ss@h:{ str := s, startPos := b, stopPos := stopPos }, n => { str := s, startPos := b, stopPos := b + ss.prevn n { byteIdx := ss.bsize } }
                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                        @[inline]
                                                                                                                                                                                                        Equations
                                                                                                                                                                                                        • x✝.take x = match x✝, x with | ss@h:{ str := s, startPos := b, stopPos := stopPos }, n => { str := s, startPos := b, stopPos := b + ss.nextn n 0 }
                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                          @[inline]
                                                                                                                                                                                                          Equations
                                                                                                                                                                                                          • x✝.takeRight x = match x✝, x with | ss@h:{ str := s, startPos := b, stopPos := e }, n => { str := s, startPos := b + ss.prevn n { byteIdx := ss.bsize }, stopPos := e }
                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                            @[inline]
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            • x✝.atEnd x = match x✝, x with | { str := str, startPos := b, stopPos := e }, p => b + p == e
                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                              @[inline]
                                                                                                                                                                                                              Equations
                                                                                                                                                                                                              • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                  • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                    @[inline]
                                                                                                                                                                                                                    def Substring.foldl {α : Type u} (f : αCharα) (init : α) (s : Substring) :
                                                                                                                                                                                                                    α
                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                      @[inline]
                                                                                                                                                                                                                      def Substring.foldr {α : Type u} (f : Charαα) (init : α) (s : Substring) :
                                                                                                                                                                                                                      α
                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                        @[inline]
                                                                                                                                                                                                                        def Substring.any (s : Substring) (p : CharBool) :
                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                        • s.any p = match s with | { str := s, startPos := b, stopPos := e } => s.anyAux e p b
                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                          @[inline]
                                                                                                                                                                                                                          def Substring.all (s : Substring) (p : CharBool) :
                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                          • s.all p = !s.any fun (c : Char) => !p c
                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                            • s.contains c = s.any fun (a : Char) => a == c
                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                              @[specialize #[]]
                                                                                                                                                                                                                              def Substring.takeWhileAux (s : String) (stopPos : String.Pos) (p : CharBool) (i : String.Pos) :
                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                @[inline]
                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                • x✝.takeWhile x = match x✝, x with | { str := s, startPos := b, stopPos := e }, p => let e := Substring.takeWhileAux s e p b; { str := s, startPos := b, stopPos := e }
                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                  @[inline]
                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                  • x✝.dropWhile x = match x✝, x with | { str := s, startPos := b, stopPos := e }, p => let b := Substring.takeWhileAux s e p b; { str := s, startPos := b, stopPos := e }
                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                    @[specialize #[]]
                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                    • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                      @[inline]
                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                      • x✝.takeRightWhile x = match x✝, x with | { str := s, startPos := b, stopPos := e }, p => let b := Substring.takeRightWhileAux s b p e; { str := s, startPos := b, stopPos := e }
                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                        @[inline]
                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                        • x✝.dropRightWhile x = match x✝, x with | { str := s, startPos := b, stopPos := e }, p => let e := Substring.takeRightWhileAux s b p e; { str := s, startPos := b, stopPos := e }
                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                          @[inline]
                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                            @[inline]
                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                              @[inline]
                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                              • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                • s.isNat = s.all fun (c : Char) => c.isDigit
                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                    def Substring.beq (ss1 : Substring) (ss2 : Substring) :
                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                    • ss1.beq ss2 = (ss1.bsize == ss2.bsize && ss1.str.substrEq ss1.startPos ss2.str ss2.startPos ss1.bsize)
                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                      def String.drop (s : String) (n : Nat) :
                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                      • s.drop n = (s.toSubstring.drop n).toString
                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                        • s.dropRight n = (s.toSubstring.dropRight n).toString
                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                          def String.take (s : String) (n : Nat) :
                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                          • s.take n = (s.toSubstring.take n).toString
                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                            • s.takeRight n = (s.toSubstring.takeRight n).toString
                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                              def String.takeWhile (s : String) (p : CharBool) :
                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                              • s.takeWhile p = (s.toSubstring.takeWhile p).toString
                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                def String.dropWhile (s : String) (p : CharBool) :
                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                • s.dropWhile p = (s.toSubstring.dropWhile p).toString
                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                  • s.takeRightWhile p = (s.toSubstring.takeRightWhile p).toString
                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                    • s.dropRightWhile p = (s.toSubstring.dropRightWhile p).toString
                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                      def String.startsWith (s : String) (pre : String) :
                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                      • s.startsWith pre = (s.toSubstring.take pre.length == pre.toSubstring)
                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                        def String.endsWith (s : String) (post : String) :
                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                        • s.endsWith post = (s.toSubstring.takeRight post.length == post.toSubstring)
                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                          • s.trimRight = s.toSubstring.trimRight.toString
                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                            • s.trimLeft = s.toSubstring.trimLeft.toString
                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                              • s.trim = s.toSubstring.trim.toString
                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                @[inline]
                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                  @[inline]
                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                  • s.nextUntil p i = s.nextWhile (fun (c : Char) => !p c) i
                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                        • s.capitalize = s.set 0 (s.get 0).toUpper
                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                          • s.decapitalize = s.set 0 (s.get 0).toLower
                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                            Instances For