2020-05-18 15:59:07 +03:00
|
|
|
module Data.String.Extra
|
|
|
|
|
Fix import loading
This was taking too long, and adding too many things, because it was
going too deep in the name of having everything accessible at the REPL
and for the compiler. So, it's done a bit differently now, only chasing
everything on a "full" load (i.e., final load at the REPL)
This has some effects:
+ As systems get bigger, load time gets better (on my machine, checking
Idris.Main now takes 52s from scratch, down from 76s)
+ You might find import errors that you didn't previously get, because
things were being imported that shouldn't have been. The new way is
correct!
An unfortunate effect is that sometimes you end up getting "undefined
name" errors even if you didn't explicitly use the name, because
sometimes a module uses a name from another module in a type, which then
gets exported, and eventually needs to be reduced. This mostly happens
because there is a compile time check that should be done which I
haven't implemented yet. That is, public export definitions should only
be allowed to use names that are also public export. I'll get to this
soon.
2020-05-27 17:49:03 +03:00
|
|
|
import Data.List
|
2021-03-17 17:07:52 +03:00
|
|
|
import Data.List1
|
Fix import loading
This was taking too long, and adding too many things, because it was
going too deep in the name of having everything accessible at the REPL
and for the compiler. So, it's done a bit differently now, only chasing
everything on a "full" load (i.e., final load at the REPL)
This has some effects:
+ As systems get bigger, load time gets better (on my machine, checking
Idris.Main now takes 52s from scratch, down from 76s)
+ You might find import errors that you didn't previously get, because
things were being imported that shouldn't have been. The new way is
correct!
An unfortunate effect is that sometimes you end up getting "undefined
name" errors even if you didn't explicitly use the name, because
sometimes a module uses a name from another module in a type, which then
gets exported, and eventually needs to be reduced. This mostly happens
because there is a compile time check that should be done which I
haven't implemented yet. That is, public export definitions should only
be allowed to use names that are also public export. I'll get to this
soon.
2020-05-27 17:49:03 +03:00
|
|
|
import Data.Nat
|
2021-06-28 15:48:37 +03:00
|
|
|
import Data.String
|
2020-05-18 15:59:07 +03:00
|
|
|
|
|
|
|
%default total
|
|
|
|
|
|
|
|
infixl 5 +>
|
|
|
|
infixr 5 <+
|
|
|
|
|
|
|
|
||| Adds a character to the end of the specified string.
|
|
|
|
|||
|
|
|
|
||| ```idris example
|
|
|
|
||| strSnoc "AB" 'C'
|
|
|
|
||| ```
|
|
|
|
||| ```idris example
|
|
|
|
||| strSnoc "" 'A'
|
|
|
|
||| ```
|
|
|
|
public export
|
|
|
|
strSnoc : String -> Char -> String
|
|
|
|
strSnoc s c = s ++ (singleton c)
|
|
|
|
|
|
|
|
||| Alias of `strSnoc`
|
|
|
|
|||
|
|
|
|
||| ```idris example
|
|
|
|
||| "AB" +> 'C'
|
|
|
|
||| ```
|
|
|
|
public export
|
|
|
|
(+>) : String -> Char -> String
|
|
|
|
(+>) = strSnoc
|
|
|
|
|
|
|
|
||| Alias of `strCons`
|
|
|
|
|||
|
|
|
|
||| ```idris example
|
|
|
|
||| 'A' <+ "AB"
|
|
|
|
||| ```
|
|
|
|
public export
|
|
|
|
(<+) : Char -> String -> String
|
|
|
|
(<+) = strCons
|
|
|
|
|
|
|
|
||| Take the first `n` characters from a string. Returns the whole string
|
|
|
|
||| if it's too short.
|
|
|
|
public export
|
|
|
|
take : (n : Nat) -> (input : String) -> String
|
|
|
|
take n str = substr Z n str
|
|
|
|
|
|
|
|
||| Take the last `n` characters from a string. Returns the whole string
|
|
|
|
||| if it's too short.
|
|
|
|
public export
|
|
|
|
takeLast : (n : Nat) -> (input : String) -> String
|
|
|
|
takeLast n str with (length str)
|
|
|
|
takeLast n str | len with (isLTE n len)
|
|
|
|
takeLast n str | len | Yes prf = substr (len `minus` n) len str
|
|
|
|
takeLast n str | len | No contra = str
|
|
|
|
|
|
|
|
||| Remove the first `n` characters from a string. Returns the empty string if
|
|
|
|
||| the input string is too short.
|
|
|
|
public export
|
|
|
|
drop : (n : Nat) -> (input : String) -> String
|
|
|
|
drop n str = substr n (length str) str
|
|
|
|
|
|
|
|
||| Remove the last `n` characters from a string. Returns the empty string if
|
|
|
|
||| the input string is too short.
|
|
|
|
public export
|
|
|
|
dropLast : (n : Nat) -> (input : String) -> String
|
|
|
|
dropLast n str = reverse (drop n (reverse str))
|
|
|
|
|
|
|
|
||| Remove the first and last `n` characters from a string. Returns the empty
|
|
|
|
||| string if the input string is too short.
|
|
|
|
public export
|
|
|
|
shrink : (n : Nat) -> (input : String) -> String
|
|
|
|
shrink n str = dropLast n (drop n str)
|
|
|
|
|
|
|
|
||| Concatenate the strings from a `Foldable` containing strings, separated by
|
|
|
|
||| the given string.
|
|
|
|
public export
|
|
|
|
join : (sep : String) -> Foldable t => (xs : t String) -> String
|
|
|
|
join sep xs = drop (length sep)
|
|
|
|
(foldl (\acc, x => acc ++ sep ++ x) "" xs)
|
|
|
|
|
|
|
|
||| Get a character from a string if the string is long enough.
|
|
|
|
public export
|
|
|
|
index : (n : Nat) -> (input : String) -> Maybe Char
|
|
|
|
index n str with (unpack str)
|
|
|
|
index n str | [] = Nothing
|
|
|
|
index Z str | (x :: xs) = Just x
|
|
|
|
index (S n) str | (x :: xs) = index n str | xs
|
|
|
|
|
|
|
|
||| Indent each line of a given string by `n` spaces.
|
|
|
|
public export
|
|
|
|
indentLines : (n : Nat) -> String -> String
|
2021-07-17 16:54:23 +03:00
|
|
|
indentLines n str = unlines $ map (indent n) $ lines str
|
2021-04-22 10:30:56 +03:00
|
|
|
|
|
|
|
||| Left-justify a string to the given length, using the
|
|
|
|
||| specified fill character on the right.
|
|
|
|
export
|
|
|
|
justifyLeft : Nat -> Char -> String -> String
|
2021-08-06 12:03:13 +03:00
|
|
|
justifyLeft n c s = s ++ replicate (n `minus` length s) c
|
2021-04-22 10:30:56 +03:00
|
|
|
|
|
|
|
||| Right-justify a string to the given length, using the
|
|
|
|
||| specified fill character on the left.
|
|
|
|
export
|
|
|
|
justifyRight : Nat -> Char -> String -> String
|
2021-08-06 12:03:13 +03:00
|
|
|
justifyRight n c s = replicate (n `minus` length s) c ++ s
|