Lifting a boolean statement to a proposition - agda

Given a list, a "property" (a function to Bools), and a proof any f xs ≡ true, I want to get a value of Any (λ x → T (f x)) ns, i.e. I want a to convert a proof from any to a proof of Any.
So far, I have been able to sketch the proof, but I am stuck at a seemingly simple statement (please see the second to last line below):
open import Data.Bool
open import Data.Bool.Properties
open import Data.List
open import Data.List.Relation.Unary.Any using (Any; here; there)
open import Data.Unit
open import Relation.Binary.PropositionalEquality
≡→T : ∀ {b : Bool} → b ≡ true → T b
≡→T refl = tt
any-val : ∀ {a} {A : Set a} (f) (ns : List A)
→ any f ns ≡ true
→ Any (λ x → T (f x)) ns
any-val f [] ()
any-val f (n ∷ ns) any-f-⟨n∷ns⟩≡true with f n
... | true = here (≡→T ?)
... | false = there (any-val f ns any-f-⟨n∷ns⟩≡true)
Apparently, I need to prove that f n ≡ true, which should be trivial to prove because we are in the true branch of the with statement f ≡ true.
What is supposed to go in there? What am I understanding incorrectly?

The key is to use the inspect idiom, which I already described thoroughly in another answer here:
`with f x` matches `false`, but cannot construct `f x == false`
This gives the following code, which I cleaned up, including imports:
open import Data.Bool
open import Data.List
open import Data.List.Relation.Unary.Any using (Any; here; there)
open import Data.Unit
open import Relation.Binary.PropositionalEquality
open import Function
≡→T : ∀ {b : Bool} → b ≡ true → T b
≡→T refl = tt
any-val : ∀ {a} {A : Set a}
f (ns : List A) → any f ns ≡ true → Any (T ∘ f) ns
any-val f (x ∷ l) p with f x | inspect f x
... | false | _ = there (any-val f l p)
... | true | [ fx≡true ] = here (≡→T fx≡true)
On a side note, be sure that you have Agda updated to the last version so that you don't have to handle the empty case any-val f [] (), as in your code.

Related

agda: How do I tell the type system that two types are equal?

Say I have type dependent on Nat
data MyType : (i : ℕ) → Set where
cons : (i : ℕ) → MyType i
and a function:
combine : {i₁ i₂ : ℕ} → MyType i₁ → MyType i₂ → MyType (i₁ ⊔ i₂)
Now I'm trying to write a function:
create-combined : (i : ℕ) → MyType i
create-combined i = combine {i} {i} (cons i) (cons i)
Unfortunately, I get the error message:
i ⊔ i != i of type ℕ
when checking that the inferred type of an application
MyType (i ⊔ i)
matches the expected type
MyType i
I know I can prove i ⊔ i ≡ i, but I don't know how to give that proof to the type system in order to get this to resolve.
How do I get this to compile?
You can use subst from Relation.Binary.PropositionalEquality.
You give it:
the proof that i ⊔ i ≡ i
the dependent type that should be transformed with the proof, in your case MyType
the term to be transformed, in your case combine {i} {i} (cons i) (cons i)
Or you can use the rewrite keyword as well on your proof, which is usually to be preferred.
Here is an (artificial) example of both possibilities applied on vector concatenation:
module ConsVec where
open import Data.Vec
open import Data.Nat
open import Data.Nat.Properties
open import Relation.Binary.PropositionalEquality
_++₁_ : ∀ {a} {A : Set a} {m n} → Vec A m → Vec A n → Vec A (n + m)
_++₁_ {m = m} {n} v₁ v₂ = subst (Vec _) (+-comm m n) (v₁ ++ v₂)
_++₂_ : ∀ {a} {A : Set a} {m n} → Vec A m → Vec A n → Vec A (n + m)
_++₂_ {m = m} {n} v₁ v₂ rewrite sym (+-comm m n) = v₁ ++ v₂

Termination checking failed

I was trying to train on this kata about longest common subsequences of a list which I slightly modified so that it worked with my versions of agda and the standard library (Agda 2.6.2, stdlib 1.7) which results in this code
{-# OPTIONS --safe #-}
module pg where
open import Data.List
open import Data.Nat
open import Data.Product
open import Relation.Nullary
open import Relation.Binary.PropositionalEquality
open import Relation.Binary
data Subseq {n} { A : Set n } : List A → List A → Set where
subseq-nil : Subseq [] []
subseq-take : ∀ a xs ys → Subseq xs ys → Subseq (a ∷ xs) (a ∷ ys)
subseq-drop : ∀ a xs ys → Subseq xs ys → Subseq xs (a ∷ ys)
is-lcs : ∀ {n} {A : Set n} → List A → List A → List A → Set n
is-lcs zs xs ys =
(Subseq zs xs × Subseq zs ys) ×
(∀ ts → Subseq ts xs → Subseq ts ys → length ts ≤ length zs)
longest : ∀ {n} {A : Set n} → List A → List A → List A
longest s1 s2 with length s1 ≤? length s2
... | yes _ = s2
... | no _ = s1
lcs : ∀ {n} {A : Set n} → Decidable {A = A} _≡_ → List A → List A → List A
lcs _ [] _ = []
lcs _ _ [] = []
lcs dec (x ∷ xs) (y ∷ ys) with dec x y
... | yes _ = x ∷ lcs dec xs ys
... | no _ = longest (lcs dec (x ∷ xs) ys) (lcs dec xs (y ∷ ys))
Unfortunetaly, Agda fails to recognize that lcs is a terminating function, which I honestly don't understand : the recursive calls are made on structurally smaller arguments if I get it right ?
If anyone can explain me what the problem is here, it would help tremendously. Thanks in advance !
Try to avoid using with when dealing with termination. The termination checker is successful when your code is refactored as follows (note that I removed your first two definitions because they are irrelevant in your question, maybe you should edit it accordingly):
{-# OPTIONS --safe #-}
open import Data.List
open import Data.Nat
open import Relation.Nullary
open import Relation.Binary.PropositionalEquality
open import Relation.Binary
open import Data.Bool using (if_then_else_)
module Term where
longest : ∀ {a} {A : Set a} → List A → List A → List A
longest s1 s2 with length s1 ≤? length s2
... | yes _ = s2
... | no _ = s1
lcs : ∀ {a} {A : Set a} → Decidable {A = A} _≡_ → List A → List A → List A
lcs _ [] _ = []
lcs _ _ [] = []
lcs dec (x ∷ xs) (y ∷ ys) = if does (dec x y) then
x ∷ lcs dec xs ys else
longest (lcs dec (x ∷ xs) ys) (lcs dec xs (y ∷ ys))
Apparently, this is a known limitation of the with abstraction combined with the termination checker as noted in the wiki:
https://agda.readthedocs.io/en/v2.6.2.1/language/with-abstraction.html#termination-checking
Here is a similar question:
Failing termination check with a with-abstraction

How can I implement a `rotate` function on Vec by using `splitAt`?

Question
I'm trying to implement a rotate function on Vec, which moves every element n positions to the left, looping around. I could implement that function by using splitAt. Here is a sketch:
open import Data.Nat
open import Data.Nat.DivMod
open import Data.Fin
open import Data.Vec
open import Relation.Nullary.Decidable
open import Relation.Binary.PropositionalEquality
rotateLeft : {A : Set} -> {w : ℕ} -> {w≢0 : False (w ≟ 0)} -> ℕ -> Vec A w -> Vec A w
rotateLeft {A} {w} n vec =
let parts = splitAt (toℕ (n mod w)) {n = ?} vec
in ?
The problem is that splitAt requires two inputs, m and n, such that the size of the vector is m + n. Since the size of the vector here is w, I need to find a k such that k + toℕ (n mod w) = w. I couldn't find any standard function handy for that. What is the best way to proceed?
Some possibilities?
Perhaps it would be helpful if k = n mod w gave me a proof that k < w, that way I could try implementing a function diff : ∀ {k w} -> k < w -> ∃ (λ a : Nat) -> a + k = w. Another possibility would be to just receive a and b as inputs, rather than the bits to shift and size of the vector, but I'm not sure that is the best interface.
Update
I've implemented the following:
add-diff : (a : ℕ) -> (b : Fin (suc a)) -> toℕ b + (a ℕ-ℕ b) ≡ a
add-diff zero zero = refl
add-diff zero (suc ())
add-diff (suc a) zero = refl
add-diff (suc a) (suc b) = cong suc (aaa a b)
Now I just need a proof that ∀ {n m} -> n mod m < m.
Here's what I came up with.
open import Data.Vec
open import Data.Nat
open import Data.Nat.DivMod
open import Data.Fin hiding (_+_)
open import Data.Product
open import Relation.Binary.PropositionalEquality
open import Data.Nat.Properties using (+-comm)
difference : ∀ m (n : Fin m) → ∃ λ o → m ≡ toℕ n + o
difference zero ()
difference (suc m) zero = suc m , refl
difference (suc m) (suc n) with difference m n
difference (suc m) (suc n) | o , p1 = o , cong suc p1
rotate-help : ∀ {A : Set} {m} (n : Fin m) → Vec A m → Vec A m
rotate-help {A} {m = m} n vec with difference m n
... | o , p rewrite p with splitAt (toℕ n) vec
... | xs , ys , _ = subst (Vec A) (+-comm o (toℕ n)) (ys ++ xs)
rotate : ∀ {A : Set} {m} (n : ℕ) → Vec A m → Vec A m
rotate {m = zero} n v = v
rotate {m = suc m} n v = rotate-help (n mod suc m) v
After talking with adamse on IRC, I've came up with this:
open import Data.Fin hiding (_+_)
open import Data.Vec
open import Data.Nat
open import Data.Nat.Properties
open import Data.Nat.DivMod
open import Data.Empty
open import Data.Product
open import Relation.Binary.PropositionalEquality
open import Relation.Nullary.Decidable
diff : {a : ℕ} → {b : Fin a} → ∃ λ c → toℕ b + c ≡ a
diff {zero} {()}
diff {suc a} {zero} = suc a , refl
diff {suc a} {suc b} with diff {a} {b}
... | c , prf = c , cong suc prf
rotateLeft : {A : Set} → {w : ℕ} → {w≢0 : False (w ≟ 0)} → ℕ → Vec A w → Vec A w
rotateLeft {A} {w} {w≢0} n v =
let m = _mod_ n w {w≢0}
d = diff {w} {m}
d₁ = proj₁ d
d₂ = proj₂ d
d₃ = subst (λ x → x ≡ w) (+-comm (toℕ (n mod w)) d₁) d₂
v₁ = subst (λ x → Vec A x) (sym d₂) v
sp = splitAt {A = A} (toℕ m) {n = d₁} v₁
xs = proj₁ (proj₂ sp)
ys = proj₁ sp
in subst (λ x → Vec A x) d₃ (xs ++ ys)
Which is nowhere as elegant as his implementation (partly because I'm still learning Agda's syntax so I opt to just use functions), but works. Now I should return a more refined type, I believe. (Can't thank him enough!)
For your last question to just prove k < w, since k = toℕ (n mod w), you can use bounded from Data.Fin.Properties:
bounded : ∀ {n} (i : Fin n) → toℕ i ℕ< n

Agda's standard library Data.AVL.Sets containing Data.String as values

I am trying to figure out how to use Agda's standard library implementation of finite sets based on AVL trees in the Data.AVL.Sets module. I was able to do so successfully using ℕ as the values with the following code.
import Data.AVL.Sets
open import Data.Nat.Properties as ℕ
open import Relation.Binary using (module StrictTotalOrder)
open Data.AVL.Sets (StrictTotalOrder.isStrictTotalOrder ℕ.strictTotalOrder)
test = singleton 5
Now I want to achieve the same thing but with Data.String as the values. There doesn't seem to be a corresponding Data.String.Properties module, but Data.String exports strictTotalOrder : StrictTotalOrder _ _ _ which I thought looked appropriate.
However, just strictly replacing the modules according to this assumption fails.
import Data.AVL.Sets
open import Data.String as String
open import Relation.Binary using (module StrictTotalOrder)
open Data.AVL.Sets (StrictTotalOrder.isStrictTotalOrder String.strictTotalOrder)
Produces the error
.Relation.Binary.List.Pointwise.Rel
(StrictTotalOrder._≈_ .Data.Char.strictTotalOrder) (toList x) (toList x₁)
!= x .Relation.Binary.Core.Dummy.≡ x₁ of type Set
when checking that the expression
StrictTotalOrder.isStrictTotalOrder String.strictTotalOrder
has type
Relation.Binary.IsStrictTotalOrder .Relation.Binary.Core.Dummy._≡_
__<__3
which I find difficult to unpack in detail since I have no idea what the Core.Dummy stuff is. It seems that there is some problem with the pointwise definition of the total order for Strings, but I can't figure it out.
If you look at Data.AVL.Sets, you can see that it is parameterised by a strict total order associated to the equivalence relation _≡_ (defined in Relation.Binary.PropositionalEquality):
module Data.AVL.Sets
{k ℓ} {Key : Set k} {_<_ : Rel Key ℓ}
(isStrictTotalOrder : IsStrictTotalOrder _≡_ _<_)
where
Now we can have a look at how the strict total order on Strings is defined. We first convert the Strings to List Chars and then compare them based on the strict lexicographic ordering for lists:
strictTotalOrder =
On.strictTotalOrder
(StrictLex.<-strictTotalOrder Char.strictTotalOrder)
toList
If we dig into the code for StrictLex.<-strictTotalOrder, we can see that the equivalence relation associated to our List of Chars is built using the pointwise lifting Pointwise.isEquivalence of whatever the equivalence relation for Chars is.
But Pointwise.isEquivalence is defined in term of this datatype:
data Rel {a b ℓ} {A : Set a} {B : Set b}
(_∼_ : REL A B ℓ) : List A → List B → Set (a ⊔ b ⊔ ℓ) where
[] : Rel _∼_ [] []
_∷_ : ∀ {x xs y ys} (x∼y : x ∼ y) (xs∼ys : Rel _∼_ xs ys) →
Rel _∼_ (x ∷ xs) (y ∷ ys)
So when Agda expects a strict total order associated to _≡_, we instead provided it with a strict total order associated to Rel _ on toList which has no chance of unifying.
How do we move on from here? Well, you could define your own strict total order on strings. Alternatively, you can try to turn the current one into one where _≡_ is the equivalence used. This is what I am going to do in the rest of this post.
So, I want to reuse an IsStrictTotalOrder R O with a different equivalence relation R′. The trick is to notice that if can transport values from R a b to R′ a b then, I should be fine! So I introduce a notion of RawIso A B which states that we can always transport values from A to B and vice-versa:
record RawIso {ℓ : Level} (A B : Set ℓ) : Set ℓ where
field
push : A → B
pull : B → A
open RawIso public
Then we can prove that RawIsos preserve a lot of properties:
RawIso-IsEquivalence :
{ℓ ℓ′ : Level} {A : Set ℓ} {R R′ : Rel A ℓ′} →
(iso : {a b : A} → RawIso (R a b) (R′ a b)) →
IsEquivalence R → IsEquivalence R′
RawIso-IsEquivalence = ...
RawIso-Trichotomous :
{ℓ ℓ′ ℓ′′ : Level} {A : Set ℓ} {R R′ : Rel A ℓ′} {O : Rel A ℓ′′} →
(iso : {a b : A} → RawIso (R a b) (R′ a b)) →
Trichotomous R O → Trichotomous R′ O
RawIso-Trichotomous = ...
RawIso-Respects₂ :
{ℓ ℓ′ ℓ′′ : Level} {A : Set ℓ} {R R′ : Rel A ℓ′} {O : Rel A ℓ′′} →
(iso : {a b : A} → RawIso (R a b) (R′ a b)) →
O Respects₂ R → O Respects₂ R′
RawIso-Respects₂ = ...
All these lemmas can be combined to prove that given a strict total order, we can build a new one via a RawIso:
RawIso-IsStrictTotalOrder :
{ℓ ℓ′ ℓ′′ : Level} {A : Set ℓ} {R R′ : Rel A ℓ′} {O : Rel A ℓ′′} →
(iso : {a b : A} → RawIso (R a b) (R′ a b)) →
IsStrictTotalOrder R O → IsStrictTotalOrder R′ O
RawIso-IsStrictTotalOrder = ...
Now that we know we can transport strict total orders along these RawIsos, we simply need to prove that the equivalence relation used by the strict total order defined in Data.String is in RawIso with propositional equality. It's (almost) simply a matter of unfolding the definitions. The only problem is that equality on characters is defined by first converting them to natural numbers and then using propositional equality. But the toNat function used has no stated property (compare e.g. to toList and fromList which are stated to be inverses)! I threw in this hack and I think it should be fine but if someone has a better solution, I'd love to know it!
toNat-injective : {c d : Char} → toNat c ≡ toNat d → c ≡ d
toNat-injective {c} pr with toNat c
toNat-injective refl | ._ = trustMe -- probably unsafe
where open import Relation.Binary.PropositionalEquality.TrustMe
Anyway, now that you have this you can unfold the definitions and prove:
rawIso : {a b : String} →
RawIso ((Ptwise.Rel (_≡_ on toNat) on toList) a b) (a ≡ b)
rawIso {a} {b} = record { push = `push ; pull = `pull } where
`push : {a b : String} → (Ptwise.Rel (_≡_ on toNat) on toList) a b → a ≡ b
`push {a} {b} pr =
begin
a ≡⟨ sym (fromList∘toList a) ⟩
fromList (toList a) ≡⟨ cong fromList (aux pr) ⟩
fromList (toList b) ≡⟨ fromList∘toList b ⟩
b
∎ where
aux : {xs ys : List Char} → Ptwise.Rel (_≡_ on toNat) xs ys → xs ≡ ys
aux = Ptwise.rec (λ {xs} {ys} _ → xs ≡ ys)
(cong₂ _∷_ ∘ toNat-injective) refl
`pull : {a b : String} → a ≡ b → (Ptwise.Rel (_≡_ on toNat) on toList) a b
`pull refl = Ptwise.refl refl
Which allows you to
stringSTO : IsStrictTotalOrder _ _
stringSTO = StrictTotalOrder.isStrictTotalOrder String.strictTotalOrder
open Data.AVL.Sets (RawIso-IsStrictTotalOrder rawIso stringSTO)
Phew!
I have uploaded a raw gist so that you can easily access the code, see the imports, etc.

Agda: parse a string with numbers

I am trying to parse a string with natural numbers in Agda.
e.g., the result of stringListToℕ "1,2,3" should be Just (1 ∷ 2 ∷ 3 ∷ [])
My current code is not quite right or by any means nice, but it works.
However it returns the type:
Maybe (List (Maybe ℕ))
The Question is:
How to implement the function stringListToℕ in a nice way (compared to my code);
it should have the type Maybe (List ℕ)
(optional, not important) How can I convert the type Maybe (List (Maybe ℕ)) to Maybe (List ℕ)?
My Code:
charToℕ : Char → Maybe ℕ
charToℕ '0' = just 0
charToℕ '1' = just 1
charToℕ '2' = just 2
charToℕ '3' = just 3
charToℕ '4' = just 4
charToℕ '5' = just 5
charToℕ '6' = just 6
charToℕ '7' = just 7
charToℕ '8' = just 8
charToℕ '9' = just 9
charToℕ _ = nothing
stringToℕ' : List Char → (acc : ℕ) → Maybe ℕ
stringToℕ' [] acc = just acc
stringToℕ' (x ∷ xs) acc = charToℕ x >>= λ n → stringToℕ' xs ( 10 * acc + n )
stringToℕ : String → Maybe ℕ
stringToℕ s = stringToℕ' (toList s) 0
isComma : Char → Bool
isComma h = h Ch.== ','
notComma : Char → Bool
notComma ',' = false
notComma _ = true
{-# NO_TERMINATION_CHECK #-}
split : List Char → List (List Char)
split [] = []
split s = l ∷ split (drop (length(l) + 1) s)
where l : List Char
l = takeWhile notComma s
isNothing' : Maybe ℕ → Bool
isNothing' nothing = true
isNothing' _ = false
isNothing : List (Maybe ℕ) → Bool
isNothing l = any isNothing' l
-- wrong type, should be String -> Maybe (List N)
stringListToℕ : String → Maybe (List (Maybe ℕ))
stringListToℕ s = if (isNothing res) then nothing else just res
where res : List (Maybe ℕ)
res = map stringToℕ (map fromList( split (Data.String.toList s)))
test1 = stringListToℕ "1,2,3"
-- => just (just 1 ∷ just 2 ∷ just 3 ∷ [])
EDIT
I tried to write a conversion function using from-just, but this gives a error when type checking:
conv : Maybe (List (Maybe ℕ)) → Maybe (List ℕ)
conv (just xs) = map from-just xs
conv _ = nothing
the error is:
Cannot instantiate the metavariable _143 to solution
(Data.Maybe.From-just (_145 xs) x) since it contains the variable x
which is not in scope of the metavariable or irrelevant in the
metavariable but relevant in the solution
when checking that the expression from-just has type
Maybe (_145 xs) → _143 xs
I took the liberty of rewriting your split function into something more general which also works with the termination check:
open import Data.List
open import Data.Product
open import Function
splitBy : ∀ {a} {A : Set a} → (A → Bool) → List A → List (List A)
splitBy {A = A} p = uncurry′ _∷_ ∘ foldr step ([] , [])
where
step : A → List A × List (List A) → List A × List (List A)
step x (cur , acc) with p x
... | true = x ∷ cur , acc
... | false = [] , cur ∷ acc
Also, stringToℕ "" should most likely be nothing, unless you really want:
stringListToℕ "1,,2" ≡ just (1 ∷ 0 ∷ 2 ∷ [])
Let's rewrite it a bit (note that helper is your original stringToℕ function):
stringToℕ : List Char → Maybe ℕ
stringToℕ [] = nothing
stringToℕ list = helper list 0
where {- ... -}
And now we can put it all together. For simplicity I'm using List Char everywhere, sprinkle with fromList/toList as necessary):
let x1 = s : List Char -- start
let x2 = splitBy notComma x1 : List (List Char) -- split at commas
let x3 = map stringToℕ x2 : List (Maybe ℕ) -- map our ℕ-conversion
let x4 = sequence x3 : Maybe (List ℕ) -- turn Maybe inside out
You can find sequence in Data.List; we also have to specify which monad instance we want to use. Data.Maybe exports its monad instance under the name monad. Final code:
open import Data.Char
open import Data.List
open import Data.Maybe
open import Data.Nat
open import Function
stringListToℕ : List Char → Maybe (List ℕ)
stringListToℕ = sequence Data.Maybe.monad ∘ map stringToℕ ∘ splitBy notComma
And a small test:
open import Relation.Binary.PropositionalEquality
test : stringListToℕ ('1' ∷ '2' ∷ ',' ∷ '3' ∷ []) ≡ just (12 ∷ 3 ∷ [])
test = refl
Considering your second question: there are many ways to turn a Maybe (List (Maybe ℕ)) into a Maybe (List ℕ), for example:
silly : Maybe (List (Maybe ℕ)) → Maybe (List ℕ)
silly _ = nothing
Right, this doesn't do much. We'd like the conversion to preserve the elements if they are all just. isNothing already does this part of checking but it cannot get rid of the inner Maybe layer.
from-just could work since we know that when we use it, all elements of the List must be just x for some x. The problem is that conv in its current form is just wrong - from-just works as a function of type Maybe A → A only when the Maybe value is just x! We could very well do something like this:
test₂ : Maybe (List ℕ)
test₂ = conv ∘ just $ nothing ∷ just 1 ∷ []
And since from-list behaves as a Maybe A → ⊤ when given nothing, we are esentially trying to construct a heterogeneous list with elements of type both ⊤ and ℕ.
Let's scrap this solution, I'll show a much simpler one (in fact, it should resemble the first part of this answer).
We are given a Maybe (List (Maybe ℕ)) and we gave two goals:
take the inner List (Maybe ℕ) (if any), check if all elements are just x and in this case put them all into a list wrapped in a just, otherwise return nothing
squash the doubled Maybe layer into one
Well, the second point sounds familiar - that's something monads can do! We get:
join : {A : Set} → Maybe (Maybe A) → Maybe A
join mm = mm >>= λ x → x
where
open RawMonad Data.Maybe.monad
This function could work with any monad but we'll be fine with Maybe.
And for the first part, we need a way to turn a List (Maybe ℕ) into a Maybe (List ℕ) - that is, we want to swap the layers while propagating the possible error (i.e. nothing) into the outer layer. Haskell has specialized typeclass for this kind of stuff (Traversable from Data.Traversable), this question has some excellent answers if you'd like to know more. Basically, it's all about rebuilding the structure while collecting the "side effects". We'll be fine with the version that works just for Lists and we're back at sequence again.
There's still one piece missing, let's look at what we have so far:
sequence-maybe : List (Maybe ℕ) → Maybe (List ℕ)
sequence-maybe = sequence Data.Maybe.monad
join : Maybe (Maybe (List ℕ)) → Maybe (List ℕ)
-- substituting A with List ℕ
We need to apply sequence-maybe inside one Maybe layer. That's where the Maybe functor instance comes into play (you could do it with a monad instance alone, but it's more convenient). With this functor instance, we can lift an ordinary function of type a → b into a function of type Maybe a → Maybe b. And finally:
open import Category.Functor
open import Data.Maybe
final : Maybe (List (Maybe ℕ)) → Maybe (List ℕ)
final mlm = join (sequence-maybe <$> mlm)
where
open RawFunctor functor
I had a go at it trying not to be clever and using simple recursive functions rather than stdlib magic. parse xs m ns parses xs by recording the (possibly empty) prefix already read in m while keeping the list of numbers already parsed in the accumulator ns.
If a parsing failure happens (non recognized character, two consecutive ,, etc.) everything is thrown away and we return nothing.
module parseList where
open import Data.Nat
open import Data.List
open import Data.Maybe
open import Data.Char
open import Data.String
isDigit : Char → Maybe ℕ
isDigit '0' = just 0
isDigit '1' = just 1
isDigit '2' = just 2
isDigit '3' = just 3
isDigit _ = nothing
attach : Maybe ℕ → ℕ → ℕ
attach nothing n = n
attach (just m) n = 10 * m + n
Quote : List Char → Maybe (List ℕ)
Quote xs = parse xs nothing []
where
parse : List Char → Maybe ℕ → List ℕ → Maybe (List ℕ)
parse [] nothing ns = just ns
parse [] (just n) ns = just (n ∷ ns)
parse (',' ∷ tl) (just n) ns = parse tl nothing (n ∷ ns)
parse (hd ∷ tl) m ns with isDigit hd
... | nothing = nothing
... | just n = parse tl (just (attach m n)) ns
stringListToℕ : String → Maybe (List ℕ)
stringListToℕ xs with Quote (toList xs)
... | nothing = nothing
... | just ns = just (reverse ns)
open import Relation.Binary.PropositionalEquality
test : stringListToℕ ("12,3") ≡ just (12 ∷ 3 ∷ [])
test = refl
Here is the Code from Vitus as a running example that uses the Agda Prelude
module Parse where
open import Prelude
-- Install Prelude
---- clone this git repo:
---- https://github.com/fkettelhoit/agda-prelude
-- Configure Prelude
--- press Meta/Alt and the letter X together
--- type "customize-group" (i.e. in the mini buffer)
--- type "agda2"
--- expand the Entry "Agda2 Include Dirs:"
--- add the directory
open import Data.Product using (uncurry′)
open import Data.Maybe using ()
open import Data.List using (sequence)
splitBy : ∀ {a} {A : Set a} → (A → Bool) → List A → List (List A)
splitBy {A = A} p = uncurry′ _∷_ ∘ foldr step ([] , [])
where
step : A → List A × List (List A) → List A × List (List A)
step x (cur , acc) with p x
... | true = x ∷ cur , acc
... | false = [] , cur ∷ acc
charsToℕ : List Char → Maybe ℕ
charsToℕ [] = nothing
charsToℕ list = stringToℕ (fromList list)
notComma : Char → Bool
notComma c = not (c == ',')
-- Finally:
charListToℕ : List Char → Maybe (List ℕ)
charListToℕ = Data.List.sequence Data.Maybe.monad ∘ map charsToℕ ∘ splitBy notComma
stringListToℕ : String → Maybe (List ℕ)
stringListToℕ = charListToℕ ∘ toList
-- Test
test1 : charListToℕ ('1' ∷ '2' ∷ ',' ∷ '3' ∷ []) ≡ just (12 ∷ 3 ∷ [])
test1 = refl
test2 : stringListToℕ "12,33" ≡ just (12 ∷ 33 ∷ [])
test2 = refl
test3 : stringListToℕ ",,," ≡ nothing
test3 = refl
test4 : stringListToℕ "abc,def" ≡ nothing
test4 = refl

Resources