diff --git a/CHANGELOG.md b/CHANGELOG.md index 081dda5b66..34ee3bc046 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -48,6 +48,15 @@ New modules * `Data.List.Relation.Binary.Permutation.Declarative{.Properties}` for the least congruence on `List` making `_++_` commutative, and its equivalence with the `Setoid` definition. +* New module defining Naperian functors, 'logarithms of containers' (Hancock/McBride) +``` +Effect.Functor.Naperian +``` +defining +```agda + record RawNaperian (F : Set a → Set b) (c : Level) : Set _ + record Naperian (F : Set a → Set b) (c : Level) (S : Setoid a ℓ) : Set _ +``` Additions to existing modules ----------------------------- diff --git a/src/Data/Vec/Effectful.agda b/src/Data/Vec/Effectful.agda index 33abd14902..b4d39f3f7f 100644 --- a/src/Data/Vec/Effectful.agda +++ b/src/Data/Vec/Effectful.agda @@ -14,14 +14,17 @@ open import Data.Vec.Base as Vec hiding (_⊛_) open import Data.Vec.Properties open import Effect.Applicative as App using (RawApplicative) open import Effect.Functor as Fun using (RawFunctor) +open import Effect.Functor.Naperian as Nap using (RawNaperian; PropositionalNaperian) open import Effect.Monad using (RawMonad; module Join; RawMonadT; mkRawMonad) import Function.Identity.Effectful as Id open import Function.Base using (flip; _∘_) -open import Level using (Level) +open import Level using (Level; 0ℓ) +open import Relation.Binary.Bundles using (Setoid) +open import Relation.Binary.PropositionalEquality private variable - a : Level + a b : Level A : Set a n : ℕ @@ -33,8 +36,24 @@ functor = record { _<$>_ = map } -applicative : RawApplicative (λ (A : Set a) → Vec A n) -applicative {n = n} = record +rawNaperian : RawNaperian (λ (A : Set a) → Vec A n) 0ℓ +rawNaperian {n = n} = record + { rawFunctor = functor + ; Log = Fin n + ; index = lookup + ; tabulate = tabulate + } + +naperian : PropositionalNaperian (λ (A : Set a) → Vec A n) 0ℓ +naperian A = record + { rawNaperian = rawNaperian + ; index-tabulate = lookup∘tabulate + ; natural-tabulate = λ f k l → cong (flip lookup l) (tabulate-∘ f k) + ; natural-index = lookup-map + } + +rawApplicative : RawApplicative (λ (A : Set a) → Vec A n) +rawApplicative {n = n} = record { rawFunctor = functor ; pure = replicate n ; _<*>_ = Vec._⊛_ @@ -42,7 +61,7 @@ applicative {n = n} = record monad : RawMonad (λ (A : Set a) → Vec A n) monad = record - { rawApplicative = applicative + { rawApplicative = rawApplicative ; _>>=_ = DiagonalBind._>>=_ } @@ -67,10 +86,9 @@ module TraversableA {f g F} (App : RawApplicative {f} {g} F) where forA = flip mapA module TraversableM {m n M} (Mon : RawMonad {m} {n} M) where - open RawMonad Mon - open TraversableA rawApplicative public + open TraversableA (RawMonad.rawApplicative Mon) public renaming ( sequenceA to sequenceM ; mapA to mapM @@ -90,7 +108,7 @@ lookup-functor-morphism i = record -- lookup is an applicative functor morphism. -lookup-morphism : (i : Fin n) → App.Morphism (applicative {a}) Id.applicative +lookup-morphism : (i : Fin n) → App.Morphism (rawApplicative {a}) Id.applicative lookup-morphism i = record { functorMorphism = lookup-functor-morphism i ; op-pure = lookup-replicate i diff --git a/src/Effect/Functor/Naperian.agda b/src/Effect/Functor/Naperian.agda new file mode 100644 index 0000000000..9f08d1a164 --- /dev/null +++ b/src/Effect/Functor/Naperian.agda @@ -0,0 +1,82 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Naperian functor +-- +-- Definitions of Naperian Functors, as named by Hancock and McBride, +-- and subsequently documented by Jeremy Gibbons +-- in the article "APLicative Programming with Naperian Functors" +-- which appeared at ESOP 2017. +-- https://link.springer.com/chapter/10.1007/978-3-662-54434-1_21 +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module Effect.Functor.Naperian where + +open import Effect.Functor using (RawFunctor) +open import Effect.Applicative using (RawApplicative) +open import Level using (Level; suc; _⊔_) +open import Relation.Binary.Bundles using (Setoid) +open import Relation.Binary.PropositionalEquality.Properties as ≡ using (setoid) +open import Function.Base using (_∘_; const) + +private + variable + a b c ℓ : Level + A : Set a + +-- From the paper: +-- "Functor f is Naperian if there is a type p of ‘positions’ such that fa≃p→a; +-- then p behaves a little like a logarithm of f +-- in particular, if f and g are both Naperian, +-- then Log(f×g)≃Logf+Logg and Log(f.g) ≃ Log f × Log g" + +-- RawNaperian contains just the functions, not the proofs +module _ (F : Set a → Set b) c where + record RawNaperian : Set (suc (a ⊔ c) ⊔ b) where + field + rawFunctor : RawFunctor F + Log : Set c + index : F A → (Log → A) + tabulate : (Log → A) → F A + open RawFunctor rawFunctor public + + -- Full Naperian has the coherence conditions too. + + record Naperian (S : Setoid a ℓ) : Set (suc (a ⊔ c) ⊔ b ⊔ ℓ) where + field + rawNaperian : RawNaperian + open RawNaperian rawNaperian public + open module S = Setoid S + private + FS : Setoid b (c ⊔ ℓ) + FS = record + { _≈_ = λ (fx fy : F Carrier) → ∀ (l : Log) → index fx l ≈ index fy l + ; isEquivalence = record + { refl = λ _ → refl + ; sym = λ eq l → sym (eq l) + ; trans = λ i≈j j≈k l → trans (i≈j l) (j≈k l) + } + } + module FS = Setoid FS + field + index-tabulate : (f : Log → Carrier) → ((l : Log) → index (tabulate f) l ≈ f l) + natural-tabulate : (f : Carrier → Carrier) (k : Log → Carrier) → (tabulate (f ∘ k)) FS.≈ (f <$> (tabulate k)) + natural-index : (l : Log) (f : Carrier → Carrier) (as : F Carrier) → (index (f <$> as) l) ≈ f (index as l) + + tabulate-index : (fx : F Carrier) → tabulate (index fx) FS.≈ fx + tabulate-index = index-tabulate ∘ index + + PropositionalNaperian : Set (suc (a ⊔ c) ⊔ b) + PropositionalNaperian = ∀ A → Naperian (≡.setoid A) + + rawApplicative : RawNaperian → RawApplicative F + rawApplicative rn = + record + { rawFunctor = rawFunctor + ; pure = tabulate ∘ const + ; _<*>_ = λ a b → tabulate (λ i → (index a i) (index b i)) + } + where + open RawNaperian rn