Skip to content

Commit d80dfdb

Browse files
lemasterolemasteroMatthewDaggitt
authored andcommitted
Add map to Data.String.Base (#2208)
* add map to Data.String * Add test for Data.String map * CHANGELOG.md add map to list of new functions in Data.String.Base * precise import of Data.String to avoid conflict on map * Fix import statements --------- Co-authored-by: lemastero <[email protected]> Co-authored-by: MatthewDaggitt <[email protected]>
1 parent 52926d4 commit d80dfdb

File tree

5 files changed

+14
-3
lines changed

5 files changed

+14
-3
lines changed

CHANGELOG.md

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -80,3 +80,8 @@ Additions to existing modules
8080
pred-injective : .{{NonZero m}} → .{{NonZero n}} → pred m ≡ pred n → m ≡ n
8181
pred-cancel-≡ : pred m ≡ pred n → ((m ≡ 0 × n ≡ 1) ⊎ (m ≡ 1 × n ≡ 0)) ⊎ m ≡ n
8282
```
83+
84+
* Added new functions in `Data.String.Base`:
85+
```agda
86+
map : (Char → Char) → String → String
87+
```

src/Data/String/Base.agda

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -185,3 +185,9 @@ lines = linesByᵇ ('\n' Char.≈ᵇ_)
185185
-- `lines` preserves empty lines
186186
_ : lines "\nabc\n\nb\n\n\n""""abc""""b""""" ∷ []
187187
_ = refl
188+
189+
map : (Char Char) String String
190+
map f = fromList ∘ List.map f ∘ toList
191+
192+
_ : map Char.toUpper "abc""ABC"
193+
_ = refl

src/Data/Tree/Rose/Show.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -15,7 +15,7 @@ open import Data.DifferenceList as DList renaming (DiffList to DList) using ()
1515
open import Data.List.Base as List using (List; []; [_]; _∷_; _∷ʳ_)
1616
open import Data.Nat.Base using (ℕ; _∸_)
1717
open import Data.Product.Base using (_×_; _,_)
18-
open import Data.String.Base hiding (show)
18+
open import Data.String.Base using (String; _++_)
1919
open import Data.Tree.Rose using (Rose; node; map; fromBinary)
2020
open import Function.Base using (flip; _∘′_; id)
2121

src/IO.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@ module IO where
1111
open import Codata.Musical.Notation
1212
open import Codata.Musical.Costring
1313
open import Data.Unit.Polymorphic.Base
14-
open import Data.String.Base
14+
open import Data.String.Base using (String)
1515
import Data.Unit.Base as Unit0
1616
open import Function.Base using (_∘_; flip)
1717
import IO.Primitive as Prim

src/Text/Printf/Generic.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -15,7 +15,7 @@ open import Data.Maybe.Base hiding (map)
1515
open import Data.Nat.Base using (ℕ)
1616
open import Data.Product.Base hiding (map)
1717
open import Data.Product.Nary.NonDependent
18-
open import Data.String.Base
18+
open import Data.String.Base using (String)
1919
open import Data.Sum.Base using (_⊎_; inj₁; inj₂)
2020
open import Data.Unit.Base using (⊤)
2121
open import Function.Nary.NonDependent

0 commit comments

Comments
 (0)