From 716411e0c02f5900d7ef81d1ca08f8b83dc13351 Mon Sep 17 00:00:00 2001 From: Jacques Comeaux Date: Thu, 30 Oct 2025 14:03:53 -0500 Subject: Add empty list and append natural transformations --- NaturalTransformation/Instance/EmptyList.agda | 23 +++++++++++++++++++++++ 1 file changed, 23 insertions(+) create mode 100644 NaturalTransformation/Instance/EmptyList.agda (limited to 'NaturalTransformation/Instance/EmptyList.agda') diff --git a/NaturalTransformation/Instance/EmptyList.agda b/NaturalTransformation/Instance/EmptyList.agda new file mode 100644 index 0000000..2164d6c --- /dev/null +++ b/NaturalTransformation/Instance/EmptyList.agda @@ -0,0 +1,23 @@ +{-# OPTIONS --without-K --safe #-} + +module NaturalTransformation.Instance.EmptyList where + +import Function.Construct.Constant as Const + +open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) +open import Categories.Functor using (Functor) +open import Categories.Category.Instance.SingletonSet using (SingletonSetoid) +open import Categories.Functor.Construction.Constant using (const) +open import Data.List.Relation.Binary.Pointwise using (refl) +open import Data.List using ([]) +open import Level using (0ℓ) +open import Functor.Instance.List {0ℓ} {0ℓ} using (List) +open import Relation.Binary using (Setoid) + +module List = Functor List + +⊤⇒[] : NaturalTransformation (const SingletonSetoid) List +⊤⇒[] = ntHelper record + { η = λ X → Const.function SingletonSetoid (List.₀ X) [] + ; commute = λ {_} {B} f → let module B = Setoid B in refl B.refl + } -- cgit v1.2.3