{-# OPTIONS --without-K --safe #-}
module Data.String.Base where
open import Level using (zero)
open import Data.Bool.Base using (true; false)
open import Data.Char.Base as Char using (Char)
open import Data.List.Base as List using (List; [_]; _∷_; [])
open import Data.List.NonEmpty.Base as NE using (List⁺)
open import Data.List.Relation.Binary.Pointwise.Base using (Pointwise)
open import Data.List.Relation.Binary.Lex.Core using (Lex-<; Lex-≤)
open import Data.Maybe.Base as Maybe using (Maybe)
open import Data.Nat.Base using (ℕ; _∸_; ⌊_/2⌋; ⌈_/2⌉)
open import Data.Product using (proj₁; proj₂)
open import Function.Base using (_on_; _∘′_; _∘_)
open import Relation.Binary.Core using (Rel)
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl)
open import Relation.Unary using (Pred; Decidable)
open import Relation.Nullary using (does)
import Agda.Builtin.String as String
open String public using ( String )
renaming
( primStringUncons to uncons
; primStringToList to toList
; primStringFromList to fromList
; primShowString to show
)
infix 4 _≈_
_≈_ : Rel String zero
_≈_ = Pointwise _≡_ on toList
infix 4 _<_
_<_ : Rel String zero
_<_ = Lex-< _≡_ Char._<_ on toList
infix 4 _≤_
_≤_ : Rel String zero
_≤_ = Lex-≤ _≡_ Char._<_ on toList
head : String → Maybe Char
head = Maybe.map proj₁ ∘′ uncons
tail : String → Maybe String
tail = Maybe.map proj₂ ∘′ uncons
fromChar : Char → String
fromChar = fromList ∘′ [_]
fromList⁺ : List⁺ Char → String
fromList⁺ = fromList ∘′ NE.toList
infixr 5 _++_
_++_ : String → String → String
_++_ = String.primStringAppend
length : String → ℕ
length = List.length ∘ toList
replicate : ℕ → Char → String
replicate n = fromList ∘ List.replicate n
concat : List String → String
concat = List.foldr _++_ ""
intersperse : String → List String → String
intersperse sep = concat ∘′ (List.intersperse sep)
unwords : List String → String
unwords = intersperse " "
unlines : List String → String
unlines = intersperse "\n"
parens : String → String
parens s = "(" ++ s ++ ")"
braces : String → String
braces s = "{" ++ s ++ "}"
infixr 5 _<+>_
_<+>_ : String → String → String
"" <+> b = b
a <+> "" = a
a <+> b = a ++ " " ++ b
padLeft : Char → ℕ → String → String
padLeft c n str with n ∸ length str
... | 0 = str
... | l = replicate l c ++ str
padRight : Char → ℕ → String → String
padRight c n str with n ∸ length str
... | 0 = str
... | l = str ++ replicate l c
padBoth : Char → Char → ℕ → String → String
padBoth cₗ cᵣ n str with n ∸ length str
... | 0 = str
... | l = replicate ⌊ l /2⌋ cₗ ++ str ++ replicate ⌈ l /2⌉ cᵣ
data Alignment : Set where
Left Center Right : Alignment
fromAlignment : Alignment → ℕ → String → String
fromAlignment Left = padRight ' '
fromAlignment Center = padBoth ' ' ' '
fromAlignment Right = padLeft ' '