With expression non evaluation - agda

I am trying to define a CoList without the delay constructors. I am running into a problem where I use a with expression but agda doesn't refine the type of a subcase.
module Failing where
open import Data.Unit
open import Data.Empty
open import Data.Maybe
open import Data.Nat
open import Data.Vec hiding (head ; tail ; map ; take)
record CoList (A : Set) : Set where
coinductive
field
head : Maybe A
tail : maybe (λ _ → ⊤) ⊥ head -> CoList A
open CoList
nil : ∀ {A} -> CoList A
head nil = nothing
tail nil ()
cons : ∀ {A} -> A -> CoList A -> CoList A
head (cons x xs) = just x
tail (cons x xs) tt = xs
take : ∀ {A} -> CoList A -> (n : ℕ) -> Maybe (Vec A n)
take l zero = just []
take l (suc n) with head l
... | nothing = nothing
... | just x = map (λ xs → x ∷ xs) (take (tail l {!!}) n)
The type of that hole is maybe (λ _ → ⊤) ⊥ (head l) but because of the with expression I would expect the type to be ⊤. I expect this because I withed on the head l and in that case head l = just x. If I try to fill the whole with tt agda mode gives me the following error:
⊤ !=< (maybe (λ _ → ⊤) ⊥ (head l)) of type Set
when checking that the expression tt has type
(maybe (λ _ → ⊤) ⊥ (head l))
I answered the question below, so now I am curious is there a better way to encode this list without the delay constructor?

You can think of with t as replacing t by whatever you match against, in the types of both function arguments and the goal. However, head l does not appear in your goal type when you perform the with — a goal whose type involves head l only appears later, once you have partially constructed the solution. This is the reason why your initial attempt doesn't work.
The inspect idiom, as demonstrated in your answer, is indeed the usual solution for this sort of problem.
As for encodings of coinductive types with 'more than one constructor', there are two (closely related) approaches that I'm aware of:
A mutual inductive/coinductive type:
data CoList′ (A : Set) : Set
record CoList (A : Set) : Set
data CoList′ A where
[] : CoList′ A
_∷_ : A → CoList A → CoList′ A
record CoList A where
coinductive
field
unfold : CoList′ A
open CoList
repeat : ∀ {A} → A → CoList A
repeat x .unfold = x ∷ repeat x
take : ∀ {A} → ℕ → CoList A → List A
take zero _ = []
take (suc n) xs with unfold xs
... | [] = []
... | x ∷ xs′ = x ∷ take n xs′
Taking the cofixpoint explicitly:
data CoList′ (A : Set) (CoList : Set) : Set where
[] : CoList′ A CoList
_∷_ : A → CoList → CoList′ A CoList
record CoList (A : Set) : Set where
coinductive
field
unfold : CoList′ A (CoList A)
open CoList
repeat : ∀ {A} → A → CoList A
repeat x .unfold = x ∷ repeat x
take : ∀ {A} → ℕ → CoList A → List A
take zero _ = []
take (suc n) xs with unfold xs
... | [] = []
... | x ∷ xs′ = x ∷ take n xs′

One solution I found is to use the inspect idiom. Apparently with abstractions in agda don't propagate equalities. The inspect idiom makes the equality obvious.
data Uncons (A : Set) : Set where
Nil : Uncons A
Cons : A -> CoList A -> Uncons A
uncons : ∀ {A} -> CoList A -> Uncons A
uncons l with head l | inspect head l
uncons l | nothing | _ = Nil
uncons l | just x | [ p ] = Cons x (tail l (subst (maybe (λ _ -> ⊤) ⊥) (sym p) tt))
take : ∀ {A} -> CoList A -> (n : ℕ) -> Maybe (Vec A n)
take l zero = just []
take l (suc n) with uncons l
... | Nil = nothing
... | Cons x xs = map (λ rest → x ∷ rest) (take xs n)

Related

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

Eliminating erased argument with only one valid case

I have defined infinite streams as follows:
record Stream (A : Set) : Set where
coinductive
field head : A
field tail : Stream A
and an inductive type which shows that some element in a stream eventually satisfies a predicate:
data Eventually {A} (P : A -> Set) (xs : Stream A) : Set where
here : P (head xs) -> Eventually P xs
there : Eventually P (tail xs) -> Eventually P xs
I would like to write a function which skips over elements of the stream until the head of the stream satisfies a predicate. To ensure termination, we must know that an element eventually satisfies the predicate, else we could loop forever. Hence, the definition of Eventually must be passed as an argument. Furthermore, the function should not computationally depend on the Eventually predicate, as it is just there to prove termination, so I would like it to be an erased argument.
dropUntil : {A : Set} {P : A -> Set} (decide : ∀ x → Dec (P x)) → (xs : Stream A) → #0 Eventually P xs → Stream A
dropUntil decide xs ev with decide (head xs)
... | yes prf = xs
... | no contra = dropUntil decide (tail xs) ?
Here is the problem - I would like to fill in the hole in the definition. From contra in scope, we know that the head of the stream does not satisfy P, and hence by definition of eventually, some element in the tail of the stream must satisfy P. If Eventually wasn't erased in this context, we could simply pattern match on the predicate, and prove the here case impossible. Normally in these scenarios I would write an erased auxiliary function, on the lines of:
#0 eventuallyInv : ∀ {A} {P : A → Set} {xs : Stream A} → Eventually P xs → ¬ P (head xs) → Eventually P (tail xs)
eventuallyInv (here x) contra with contra x
... | ()
eventuallyInv (there ev) contra = ev
The problem with this approach is that the Eventually proof is the structurally recursive argument in dropUntil, and calling this auxiliary function does not pass the termination checker as Agda does not "look inside" the function definition.
Another approach I tried is inlining the above erased function into the definition of dropUntil. Unfortunately, I had no luck with this approach either - using the definition of case ... of like described here https://agda.readthedocs.io/en/v2.5.2/language/with-abstraction.html did not pass the termination checker either.
I have written an equivalent program in Coq which is accepted (using Prop rather than erased types), so I am confident that my reasoning is correct. The main reason why Coq accepted the definition and Agda doesn't is that Coq's termination checker expands function definitions, and hence the "auxiliary erased function" approach succeeds.
EDIT:
This is my attempt using sized types, however it does not pass the termination checker and I can't figure out why.
record Stream (A : Set) : Set where
coinductive
field
head : A
tail : Stream A
open Stream
data Eventually {A} (P : A → Set) (xs : Stream A) : Size → Set where
here : ∀ {i} → P (head xs) → Eventually P xs (↑ i)
there : ∀ {i} → Eventually P (tail xs) i → Eventually P xs (↑ i)
#0 eventuallyInv : ∀ {A P i} {xs : Stream A} → Eventually P xs (↑ i) → ¬ P (head xs) → Eventually P (tail xs) i
eventuallyInv (here p) ¬p with ¬p p
... | ()
eventuallyInv (there ev) ¬p = ev
dropUntil : ∀ {A P i} → (∀ x → Dec (P x)) → (xs : Stream A) → #0 Eventually P xs (↑ i) → Stream A
dropUntil decide xs ev with decide (head xs)
... | yes p = xs
... | no ¬p = dropUntil decide (tail xs) (eventuallyInv ev ¬p)
In your case you can work with a weaker notion of Eventually which matches what dropUntil actually needs to know. It's also single constructor so you can match on it even when erased.
data Eventually' {A} (P : A -> Set) (xs : Stream A) : Set where
next : (¬ P (head xs) → Eventually' P (tail xs)) → Eventually' P xs
eventuallyInv : ∀ {A} {P : A → Set} {xs : Stream A} → (ev : Eventually P xs) → Eventually' P xs
eventuallyInv (here p) = next \ np → ⊥-elim (np p)
eventuallyInv (there ev) = next \ np → eventuallyInv ev
dropUntil' : {A : Set} {P : A -> Set} (decide : ∀ x → Dec (P x)) → (xs : Stream A) → #0 Eventually' P xs → Stream A
dropUntil' decide xs (next ev) with decide (head xs)
... | yes prf = xs
... | no contra = dropUntil' decide (tail xs) (ev contra)
dropUntil : {A : Set} {P : A -> Set} (decide : ∀ x → Dec (P x)) → (xs : Stream A) → #0 Eventually P xs → Stream A
dropUntil decide xs ev = dropUntil' decide xs (eventuallyInv ev)

How to inspect the value inside a product (Σ)?

I am trying to wrap my head around the Σ type by playing with a filterVec function, analogous to the filter function available for lists. See my implementation below:
-- Some imports we will need later
open import Data.Bool
open import Data.Nat
open import Data.Product
open import Data.Vec
-- The filterVec function
filterVec : {n : ℕ} -> (ℕ -> Bool) -> Vec ℕ n -> Σ ℕ (λ length → Vec ℕ length)
filterVec _ [] = 0 , []
filterVec f (x ∷ xs) with filterVec f xs
... | length , filtered = if f x then (suc length , x ∷ filtered) else (length , filtered)
Happy with my function, I decided to test it
dummyVec : Vec ℕ 5
dummyVec = 1 ∷ 2 ∷ 3 ∷ 4 ∷ 5 ∷ []
dummyFn : Vec ℕ 5
dummyFn with filterVec (λ _ → true) dummyVec
dummyFn | length , xs = {!!} -- I would like to return xs here, but Agda throws a type error
For some reason, Agda is unable to tell that length is 5, so it doesn't allow me to return xs. However, the code below does work:
open import Relation.Binary.PropositionalEquality
dummyProof : proj₁ (filterVec (λ _ → true) dummyVec) ≡ 5
dummyProof = refl
I am utterly confused by this behavior. In the case of dummyFn, Agda cannot tell that the length is 5, but in the case of dummyProof, Agda proves it without any problem.
What am I missing? How can I get dummyFn to work?
with binds a value to a name. If you want to compute, you can use let:
dummyFn′ : Vec ℕ 5
dummyFn′ = let length , xs = filterVec (λ _ → true) dummyVec
in xs
Let's look at the following example:
dummy : Set
dummy with 5
dummy | x = {!!}
Should the x magically reduce to 5 in the hole? No. You've given 5 another name — x, and those expressions are neither judgementally nor propositionally equal.
Now in your example:
dummyFn : Vec ℕ 5
dummyFn with filterVec (λ _ → true) dummyVec
dummyFn | length , xs = {!!}
You've given the length name to the length of xs. And it does not reduce to 5. It's a name. If you want both to give a name and to remember what the name is for, there is the inspect idiom for this:
dummyFn′′ : Vec ℕ 5
dummyFn′′ with filterVec (λ _ → true) dummyVec | inspect (filterVec (λ _ → true)) dummyVec
dummyFn′′ | length , xs | [ refl ] = xs

Irrelevant implicits: Why doesn't agda infer this proof?

Recently I made a type for finite sets in Agda with the following implementation:
open import Relation.Nullary
open import Relation.Nullary.Negation
open import Data.Empty
open import Data.Unit
open import Relation.Binary.PropositionalEquality
open import Data.Nat
suc-inj : (n m : ℕ) → (suc n) ≡ (suc m) → n ≡ m
suc-inj n .n refl = refl
record Eq (A : Set) : Set₁ where
constructor mkEqInst
field
_decide≡_ : (a b : A) → Dec (a ≡ b)
open Eq {{...}}
mutual
data FinSet (A : Set) {{_ : Eq A }} : Set where
ε : FinSet A
_&_ : (a : A) → (X : FinSet A) → .{ p : ¬ (a ∈ X)} → FinSet A
_∈_ : {A : Set} → {{p : Eq A}} → (a : A) → FinSet A → Set
a ∈ ε = ⊥
a ∈ (b & B) with (a decide≡ b)
... | yes _ = ⊤
... | no _ = a ∈ B
_∉_ : {A : Set} → {{p : Eq A}} → (a : A) → FinSet A → Set
_∉_ a X = ¬ (a ∈ X)
decide∈ : {A : Set} → {{_ : Eq A}} → (a : A) → (X : FinSet A) → Dec (a ∈ X)
decide∈ a ε = no (λ z → z)
decide∈ a (b & X) with (a decide≡ b)
decide∈ a (b & X) | yes _ = yes tt
... | no _ = decide∈ a X
decide∉ : {A : Set} → {{_ : Eq A}} → (a : A) → (X : FinSet A) → Dec (a ∉ X)
decide∉ a X = ¬? (decide∈ a X)
instance
eqℕ : Eq ℕ
eqℕ = mkEqInst decide
where decide : (a b : ℕ) → Dec (a ≡ b)
decide zero zero = yes refl
decide zero (suc b) = no (λ ())
decide (suc a) zero = no (λ ())
decide (suc a) (suc b) with (decide a b)
... | yes p = yes (cong suc p)
... | no p = no (λ x → p ((suc-inj a b) x))
However, when I test this type out with the following:
test : FinSet ℕ
test = _&_ zero ε
Agda for some reason can't infer the implicit argument of type ¬ ⊥! However, auto of course finds the proof of this trivial proposition: λ x → x : ¬ ⊥.
My question is this: Since I've marked the implicit proof as irrelevant, why can't Agda simply run auto to find the proof of ¬ ⊥ during type checking? Presumably, whenever filling in other implicit arguments, it might matter exactly what proof Agda finda, so it shouldn't just run auto, but if the proof has been marked irrelevant, like it my case, why can't Agda find a proof?
Note: I have a better implementation of this, where I implement ∉ directly, and Agda can find the relevant proof, but I want to understand in general why Agda can't automatically find these sorts of proofs for implicit arguments. Is there any way in the current implementation of Agda to get these "auto implicits" like I want here? Or is there some theoretical reason why this would be a bad idea?
There's no fundamental reason why irrelevant arguments couldn't be solved by proof search, however the fear is that in many cases it would be slow and/or not find a solution.
A more user-directed thing would be to allow the user to specify that a certain argument should be inferred using a specific tactic, but that has not been implemented either. In your case you would provide a tactic that tries to solve the goal with (\ x -> x).
If you give a more direct definition of ∉, then the implicit argument gets type ⊤ instead of ¬ ⊥. Agda can fill in arguments of type ⊤ automatically by eta-expansion, so your code just works:
open import Relation.Nullary
open import Relation.Nullary.Negation
open import Data.Empty
open import Data.Unit
open import Relation.Binary.PropositionalEquality
open import Data.Nat
suc-inj : (n m : ℕ) → (suc n) ≡ (suc m) → n ≡ m
suc-inj n .n refl = refl
record Eq (A : Set) : Set₁ where
constructor mkEqInst
field
_decide≡_ : (a b : A) → Dec (a ≡ b)
open Eq {{...}}
mutual
data FinSet (A : Set) {{_ : Eq A}} : Set where
ε : FinSet A
_&_ : (a : A) → (X : FinSet A) → .{p : (a ∉ X)} → FinSet A
_∉_ : {A : Set} → {{p : Eq A}} → (a : A) → FinSet A → Set
a ∉ ε = ⊤
a ∉ (b & X) with (a decide≡ b)
... | yes _ = ⊥
... | no _ = a ∉ X
decide∉ : {A : Set} → {{_ : Eq A}} → (a : A) → (X : FinSet A) → Dec (a ∉ X)
decide∉ a ε = yes tt
decide∉ a (b & X) with (a decide≡ b)
... | yes _ = no (λ z → z)
... | no _ = decide∉ a X
instance
eqℕ : Eq ℕ
eqℕ = mkEqInst decide
where decide : (a b : ℕ) → Dec (a ≡ b)
decide zero zero = yes refl
decide zero (suc b) = no (λ ())
decide (suc a) zero = no (λ ())
decide (suc a) (suc b) with (decide a b)
... | yes p = yes (cong suc p)
... | no p = no (λ x → p ((suc-inj a b) x))
test : FinSet ℕ
test = _&_ zero ε

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