How do I use Agda's implementation of delimited continuations? - standard-library

We can implement a delimited continuation monad in Agda rather easily.
There is, however, no need to, as the Agda "standard library" has an implementation of a delimited continuation monad. What confuses me about this implementation, though, is the addition of an extra parameter to the DCont type.
DCont : ∀ {i f} {I : Set i} → (I → Set f) → IFun I f
DCont K = DContT K Identity
My question is: why is the extra parameter K there? And how would I use the DContIMonadDCont instance? Can I open it in such a way that I'll get something akin to the below reference implementation in (global) scope?
All my attempts to use it are leading to unsolvable metas.
Reference implementation of delimited continuations not using the Agda "standard library".
DCont : Set → Set → Set → Set
DCont r i a = (a → i) → r
return : ∀ {r a} → a → DCont r r a
return x = λ k → k x
_>>=_ : ∀ {r i j a b} → DCont r i a → (a → DCont i j b) → DCont r j b
c >>= f = λ k → c (λ x → f x k)
join : ∀ {r i j a} → DCont r i (DCont i j a) → DCont r j a
join c = c >>= id
shift : ∀ {r o i j a} → ((a → DCont i i o) → DCont r j j) → DCont r o a
shift f = λ k → f (λ x → λ k′ → k′ (k x)) id
reset : ∀ {r i a} → DCont a i i → DCont r r a
reset a = λ k → k (a id)

Let me answer your second and third questions first. Looking at how DContT is defined:
DContT K M r₂ r₁ a = (a → M (K r₁)) → M (K r₂)
We can recover the requested definition by specifying M = id and K = id (M also has to be a monad, but we have the Identity monad). DCont already fixes M to be id, so we are left with K.
import Category.Monad.Continuation as Cont
open import Function
DCont : Set → Set → Set → Set
DCont = Cont.DCont id
Now, we can open the RawIMonadDCont module provided we have an instance of the corresponding record. And luckily, we do: Category.Monad.Continuation has one such record under the name DContIMonadDCont.
module ContM {ℓ} =
Cont.RawIMonadDCont (Cont.DContIMonadDCont {f = ℓ} id)
And that's it. Let's make sure the required operations are really there:
return : ∀ {r a} → a → DCont r r a
return = ContM.return
_>>=_ : ∀ {r i j a b} → DCont r i a → (a → DCont i j b) → DCont r j b
_>>=_ = ContM._>>=_
join : ∀ {r i j a} → DCont r i (DCont i j a) → DCont r j a
join = ContM.join
shift : ∀ {r o i j a} → ((a → DCont i i o) → DCont r j j) → DCont r o a
shift = ContM.shift
reset : ∀ {r i a} → DCont a i i → DCont r r a
reset = ContM.reset
And indeed, this typechecks. You can also check if the implementation matches. For example, using C-c C-n (normalize) on shift, we get:
λ {.r} {.o} {.i} {.j} {.a} e k → e (λ a f → f (k a)) (λ x → x)
Modulo renaming and some implicit parameters, this is exactly implementation of the shift in your question.
Now the first question. The extra parameter is there to allow additional dependency on the indices. I haven't used delimited continuations in this way, so let me reach for an example somewhere else. Consider this indexed writer:
open import Data.Product
IWriter : {I : Set} (K : I → I → Set) (i j : I) → Set → Set
IWriter K i j A = A × K i j
If we have some sort of indexed monoid, we can write a monad instance for IWriter:
record IMonoid {I : Set} (K : I → I → Set) : Set where
field
ε : ∀ {i} → K i i
_∙_ : ∀ {i j k} → K i j → K j k → K i k
module IWriterMonad {I} {K : I → I → Set} (mon : IMonoid K) where
open IMonoid mon
return : ∀ {A} {i : I} →
A → IWriter K i i A
return a = a , ε
_>>=_ : ∀ {A B} {i j k : I} →
IWriter K i j A → (A → IWriter K j k B) → IWriter K i k B
(a , w₁) >>= f with f a
... | (b , w₂) = b , w₁ ∙ w₂
Now, how is this useful? Imagine you wanted to use the writer to produce a message log or something of the same ilk. With usual boring lists, this is not a problem; but if you wanted to use vectors, you are stuck. How to express that type of the log can change? With the indexed version, you could do something like this:
open import Data.Nat
open import Data.Unit
open import Data.Vec
hiding (_>>=_)
open import Function
K : ℕ → ℕ → Set
K i j = Vec ℕ i → Vec ℕ j
K-m : IMonoid K
K-m = record
{ ε = id
; _∙_ = λ f g → g ∘ f
}
open IWriterMonad K-m
tell : ∀ {i j} → Vec ℕ i → IWriter K j (i + j) ⊤
tell v = _ , _++_ v
test : ∀ {i} → IWriter K i (5 + i) ⊤
test =
tell [] >>= λ _ →
tell (4 ∷ 5 ∷ []) >>= λ _ →
tell (1 ∷ 2 ∷ 3 ∷ [])
Well, that was a lot of (ad-hoc) code to make a point. I haven't given it much thought, so I'm fairly sure there's nicer/more principled approach, but it illustrates that such dependency allows your code to be more expressive.
Now, you could apply the same thing to DCont, for example:
test : Cont.DCont (Vec ℕ) 2 3 ℕ
test c = tail (c 2)
If we apply the definitions, the type reduces to (ℕ → Vec ℕ 3) → Vec ℕ 2. Not very convincing example, I know. But perhaps you can some up with something more useful now that you know what this parameter does.

Related

Why is Agda refusing Set₁ → Set₁?

Here's the code:
Continuation : Set → Set₁ → Set₁
Continuation R X = (X → R) → R
Selection : Set → Set₁ → Set₁
Selection R X = (X → R) → X
dagger : {R : Set} → Selection R → Continuation R
dagger S X k = k (S k)
I need continuation and selection to use a higher universe for other reasons. But then the definition of dagger raises an error:
(Set₁ → Set₁) should be a sort, but it isn't when checking that the inferred type of an application
Set₁ → Set₁
matches the expected type _16
In Agda _→_ is the function space not of any category but that of types and
functions. My guess is that you want the following:
dagger : {R : Set} → ∀ X → Selection R X → Continuation R X
dagger X S k = k (S k)
i.e. the function space of indexed types and index-preserving functions.
Alternatively you could make this clear by using the stdlib's
Relation.Unary notion of morphism and write:
open import Relation.Unary
dagger : {R : Set} → Selection R ⊆ Continuation R
dagger S k = k (S k)

Agda Store Comonad

I'm just starting with Agda but know some Haskell and would like to know how to define the Store Comonad in Agda.
This is what I have until now:
open import Category.Comonad
open import Data.Product
Store : Set → Set → ((Set → Set) × Set)
Store s a = ((λ s → a) , s)
StoreComonad : RawComonad (λ s a → (Store s a))
StoreComonad = record
{ extract (Store s a) = extract s a
; extend f (Store s a = Store (extend (λ s' a' →­ f (Store s' a')) s) a
} where open RawComonad
For now I'm getting the following error:
Parse error
=<ERROR>
extract s a
; extend f (Sto...
I'm not too sure what it is I'm doing wrong. Any help would be appreciated! Thanks!
EDIT
I think I'm getting closer. I replaced the fields in the record using matching lambdas:
Store : Set → Set → ((Set → Set) × Set)
Store s a = ((λ s → a) , s)
StoreComonad : RawComonad (λ s a → (Store s a))
StoreComonad = record
{ extract = λ st → (proj₁ st) (proj₂ st)
; duplicate = λ st → Store (λ s → Store (proj₁ st) s) (proj₂ st)
; extend = λ g st → g (duplicate st)
} where open RawComonad
RawComonad is from https://github.com/agda/agda-stdlib/blob/master/src/Category/Comonad.agda
and has the signature
record RawComonad (W : Set f → Set f) : Set (suc f)
and Store is based on type Store s a = (s -> a, s) in Haskell.
Now the error I'm getting is:
(a : Set) → Σ (Set → Set) (λ x → Set) !=< Set
when checking that the expression λ a → Store s a has type Set
I'm wondering if this error has to do with this line:
StoreComonad : RawComonad (λ s a → (Store s a))
I find that the compilation error messages in Agda don't give many clues or I haven't yet been able to understand them well.
Your problem is that λ s a → (Store s a) (or, eta-contracted, Store) is not a comonad; its type (or, for your Haskell intuition, we could say its kind) is not right.
However, for any choice of s, Store s is! So let's write that:
StoreComonad : ∀ {s : Set} → RawComonad (Store s)
The rest of the definition of StoreComonad will follow easily.
As an aside, you can make the definition of StoreComonad nicer by using pattern-matching lambdas instead of using explicit projections (and please do read that link because it seems you have mixed up normal lambdas with pattern-matching ones); e.g:
extract = λ { (f , a) → f a }
and so on.
Wow, ok, I think silence was the answer I needed. I managed to advance quite a bit on defining the Store Comonad:
S : Set
S = ℕ
Store : Set → Set
Store A = ((S → A) × S)
pos : ∀ {A : Set} → Store A → S
pos = λ st → proj₂ st
peek : ∀ {A : Set} → S → Store A → A
peek = λ s → λ st → (proj₁ st) s
fmap : ∀ {A : Set} {B : Set} → (A → B) → Store A → Store B
fmap = λ f → λ st → ((f ∘ proj₁ st) , proj₂ st)
duplicate' : ∀ {A : Set} → (Store A) → Store (Store A)
duplicate' = λ st → (λ s' → proj₁ st , s') , proj₂ st
StoreComonad : RawComonad Store
StoreComonad = record
{ extract = λ st → (proj₁ st) (proj₂ st)
; extend = λ g st → fmap g (duplicate' st)
} where open RawComonad
along the way I learned that C-c-C-l and C-c-C-r with ? can be quite helpful in trying to find the type that is needed to fill the ?. I used ? for proving some examples before but hadn't tried using it to find how to write a type.
What's left.. I would like to make S not just a Nat.

Why do function composition and application have a dependent implementation in Agda?

Why do function composition (∘) and application ($) have the implementation as available in https://github.com/agda/agda-stdlib/blob/master/src/Function.agda#L74-L76?
Copied here for convenience:
_∘_ : ∀ {a b c}
{A : Set a} {B : A → Set b} {C : {x : A} → B x → Set c} →
(∀ {x} (y : B x) → C y) → (g : (x : A) → B x) →
((x : A) → C (g x))
f ∘ g = λ x → f (g x)
_∘'_ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} →
(B → C) → (A → B) → (A → C)
f ∘' g = λ x → f (g x)
_$_ : ∀ {a b} {A : Set a} {B : A → Set b} →
((x : A) → B x) → ((x : A) → B x)
f $ x = f x
_$'_ : ∀ {a b} {A : Set a} {B : Set b} →
(A → B) → (A → B)
f $' x = f x
I initially thought the rationale behind this was that $ would be able to handle higher order types that $' wouldn't be able to handle. For example, consider A=Nat, B=List, f is ::, where B depends on A. But after a lot of testing, I couldn't come up with an example that would show that the implementation of $' is not sufficient. What scenarios does $ handle that $' isn't able to handle? (Similarly, what scenarios does ∘ handle that ∘' doesn't?
open import Agda.Builtin.Nat public
open import Agda.Primitive public
--data List {a} (A : Set a) : Set a where
-- [] : List A
-- _∷_ : (x : A) (xs : List A) → List A
data Vec {a} (A : Set a) : Nat → Set a where
[] : Vec A zero
_∷_ : ∀ {n} (x : A) (xs : Vec A n) → Vec A (suc n)
tail : ∀ {a n} {A : Set a} → Vec A (suc n) → Vec A n
tail (x ∷ s) = s
_$_ : ∀ {a b} {A : Set a} {B : A → Set b} →
((x : A) → B x) → ((x : A) → B x)
f $ x = f x
_$'_ : ∀ {a b} {A : Set a} {B : Set b} →
(A → B) → (A → B)
f $' x = f x
_∘_ : ∀ {a b c}
{A : Set a} {B : A → Set b} {C : {x : A} → B x → Set c} →
(∀ {x} (y : B x) → C y) → (g : (x : A) → B x) →
((x : A) → C (g x))
f ∘ g = λ x → f (g x)
_∘'_ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} →
(B → C) → (A → B) → (A → C)
f ∘' g = λ x → f (g x)
Vecc : ∀ {a} → Nat → (A : Set a) → (Set a)
Vecc x y = Vec y x
data Pair {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where
_,_ : (x : A) → (y : B x) → Pair A B
-- Dependent Pair attempt
--fst : ∀ {a b} {A : Set a} {B : A → Set b} → Pair A B → A
--fst (a , b) = a
--
--f : Pair Nat $' Vec Nat
--f = _,_ zero $' []
--
--g : Pair (Pair Nat $' Vec Nat) $' λ x → Nat
--g = _,_ (_,_ zero $' []) $' zero
-- Some other attempt
--f : ∀ {a n} {A : Set a} → Vec A ((suc ∘' suc) n) → Vec A n
--f {a} = tail {a} ∘' tail {a}
-- Vec attempt
--f : ∀ {a} (A : Set a) → (Set a)
--f {a} = Vecc {a} (suc zero) ∘' Vecc {a} (suc zero)
--
--h = f Nat
--
--x : h
--x = (zero ∷ []) ∷ []
-- List attempt
--f : ∀ {a} (A : Set a) → (Set a)
--f {a} = List {a} ∘' List {a}
--
--g : ∀ {a} (A : Set a) → (Set a)
--g {a} = List {a} ∘ List {a}
--
--h = f Nat
--i = g Nat
--
--x : h
--x = (zero ∷ []) ∷ []
∘′ and $′ don't work with dependent functions. You simply didn't try any tests with dependent functions. For f $ x examples, f must be dependent, for f ∘ g, either of the functions must be dependent. Example:
open import Data.Nat
open import Data.Vec
open import Function
open import Relation.Binary.PropositionalEquality
replicate' : {A : Set} → A → (n : ℕ) → Vec A n
replicate' a n = replicate a
refl' : {A : Set}(a : A) → a ≡ a
refl' a = refl
-- fail1 : Vec ℕ 10
-- fail1 = replicate' 10 $′ 10
ok1 : Vec ℕ 10
ok1 = replicate' 10 $ 10
-- fail2 : ∀ n → replicate' 10 n ≡ replicate' 10 n
-- fail2 = refl' ∘′ replicate' 10
ok2 : ∀ n → replicate' 10 n ≡ replicate' 10 n
ok2 = refl' ∘ replicate' 10
One works with dependent functions, the other doesn't, as Andras Kovacs mentioned.
The important difference is that for non-dependent functions stronger proofs can be constructed. For example:
eq : {A B} -> f : (A -> B) -> x y : A -> x == y -> (f x) == (f y)
eq f x .x refl = refl
Here we can construct equality of f x and f y. But we can't do the same for dependent functions - because there is no way to prove B x == B y. So there is only a weaker proof that f x can be "cast" to f y.
transport : {A} {B : A -> Set} -> f : (x : A -> B x) -> x y : A -> x == y -> f x -> f y
transport f x .x refl fx = fx
(Actually, transport is usually defined as B x -> B y, not for a dependent function; but I just can't come up with a better name)

How to convert the J axiom to the fixed-argument form?

I was trying to prove that true ≡ false -> Empty assuming the J axiom. It is defined as:
J : Type
J = forall
{A : Set}
{C : (x y : A) → (x ≡ y) → Set} →
(c : ∀ x → C x x refl) →
(x y : A) →
(p : x ≡ y) →
C x y p
My attempt went like this:
bad : J → true ≡ false -> Empty
bad j e = j Bool (λ { true _ _ => Unit; false _ _ => Empty }) _
Now, to proceed with the proof, I needed a term c : ∀ x -> C x x refl. Since I instantiated C, it becomes c : ∀ x -> (λ { true _ _ => Unit; false _ _ => Empty } x x refl. Then I got stuck. c can't reduce further because we don't know the value of x. I wasn't able to complete this proof. But there is a different version of J:
J' : Type
J' = forall
{A : Set}
{x : A}
{C : (y : A) → (x ≡ y) → Set} →
(c : C x refl) →
(y : A) →
(p : x ≡ y) →
C y p
With this one, this problem is solved, because t can be fixed to be true. This makes the c argument reduce to Unit, which we can provide. My question is: can we convert the former version to the later? That is, can we build a term fix_x : J → J'? Does that hold in general (i.e., can indices be converted to parameters)?
First, regarding true ≡ false -> Empty: this is unprovable if you can only eliminate into Set0 with J, so you need an universe polymorphic or large definition. I write some preliminaries here:
{-# OPTIONS --without-K #-}
open import Relation.Binary.PropositionalEquality
open import Level
data Bool : Set where true false : Bool
data Empty : Set where
record Unit : Set where
constructor tt
JTy : ∀ {i j} → Set _
JTy {i}{j} =
{A : Set i}
(P : (x y : A) → (x ≡ y) → Set j) →
(pr : ∀ x → P x x refl) →
{x y : A} →
(p : x ≡ y) →
P x y p
J : ∀ {i}{j} → JTy {i}{j}
J P pr {x} refl = pr x
J₀ = J {zero}{zero}
Now, transport or subst is the only needed thing for true ≡ false -> Empty:
transp : ∀ {i j}{A : Set i}(P : A → Set j){x y} → x ≡ y → P x → P y
transp P = J (λ x y _ → P x -> P y) (λ _ px → px)
true≢false : true ≡ false → Empty
true≢false e = transp (λ {true → Unit; false → Empty}) e tt
Considering now proving the pointed J' from J, I know about three solutions, and each uses different features from the ambient theory.
The simplest one is to use universes to abstract over the induction motive:
JTy' : ∀ {i j} → Set _
JTy' {i}{j} =
{A : Set i}
{x : A}
(P : ∀ y → x ≡ y → Set j)
(pr : P x refl)
{y : A}
(p : x ≡ y)
→ P y p
JTy→JTy' : (∀ {i j} → JTy {i}{j}) → ∀ {i}{j} → JTy' {i}{j}
JTy→JTy' J {i} {j} {A} {x} P pr {y} e =
J (λ x y e → (P : ∀ y → x ≡ y → Set j) → P x refl → P y e)
(λ x P pr → pr) e P pr
If we only want to use a fixed universe level, then it is a bit more complicated. The following solution, sometimes called "contractible singletons", needs Σ-types, but nothing else:
open import Data.Product
JTy→JTy'withΣ : JTy {zero}{zero} → JTy' {zero}{zero}
JTy→JTy'withΣ J {A} {x} P pr {y} e =
J (λ {(x , r) (y , e) _ → P x r → P y e})
(λ _ px → px)
(J (λ x y e → (x , refl) ≡ (y , e))
(λ _ → refl)
e)
pr
There is a solution which doesn't even need Σ-s, but requires the beta rule for J, which says that J P pr {x} refl = pr x. It doesn't matter whether this rule holds definitionally or just as a propositional equality, but the construction is simpler when it holds definitionally, so let's do that. Note that I don't use any universe other than Set0.
transp₀ = transp {zero}{zero}
transp2 : ∀ {A : Set}{B : A → Set}(C : ∀ a → B a → Set)
{x y : A}(e : x ≡ y){b} → C x b → C y (transp₀ B e b)
transp2 {A}{B} C {x}{y} e {b} cxb =
J₀ (λ x y e → ∀ b → C x b → C y (transp₀ B e b)) (λ _ _ cxb → cxb) e b cxb
JTy→JTy'noΣU : JTy' {zero}{zero}
JTy→JTy'noΣU {A} {x} P pr {y} e =
transp₀ (P y) (J₀ (λ x y e → transp₀ (x ≡_) e refl ≡ e) (λ _ → refl) e)
(transp2 {A} {λ y → x ≡ y} P e pr)
Philosophically, the third version is the most "conservative", since it only assumes J. The addition of the beta rule is not really an extra thing, since it is always assumed to hold (definitionally or propositionally) for _≡_.
can indices be converted to parameters?
If you have propositional equality, then all indices can be converted to parameters, and fixed in constructors using equality proofs.

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

Resources