diff options
Diffstat (limited to 'Functor/Instance')
| -rw-r--r-- | Functor/Instance/Decorate.agda | 2 | ||||
| -rw-r--r-- | Functor/Instance/Nat/Preimage.agda | 65 | ||||
| -rw-r--r-- | Functor/Instance/Nat/Pull.agda | 63 | ||||
| -rw-r--r-- | Functor/Instance/Nat/Push.agda | 64 | ||||
| -rw-r--r-- | Functor/Instance/Nat/System.agda | 96 |
5 files changed, 289 insertions, 1 deletions
diff --git a/Functor/Instance/Decorate.agda b/Functor/Instance/Decorate.agda index e87f20c..30cec87 100644 --- a/Functor/Instance/Decorate.agda +++ b/Functor/Instance/Decorate.agda @@ -40,7 +40,7 @@ module DecoratedCospans = Category DecoratedCospans module mc𝒞 = CocartesianMonoidal 𝒞.U 𝒞.cocartesian -- For every cospan there exists a free decorated cospan --- i.e. the original cospan with the empty decoration +-- i.e. the original cospan with the discrete decoration private variable diff --git a/Functor/Instance/Nat/Preimage.agda b/Functor/Instance/Nat/Preimage.agda new file mode 100644 index 0000000..7da00f4 --- /dev/null +++ b/Functor/Instance/Nat/Preimage.agda @@ -0,0 +1,65 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Instance.Nat.Preimage where + +open import Categories.Category.Instance.Nat using (Natop) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Functor using (Functor) +open import Data.Fin.Base using (Fin) +open import Data.Bool.Base using (Bool) +open import Data.Nat.Base using (ℕ) +open import Data.Subset.Functional using (Subset) +open import Data.Vec.Functional.Relation.Binary.Equality.Setoid using (≋-setoid) +open import Function.Base using (id; _∘_) +open import Function.Bundles using (Func; _⟶ₛ_) +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Setoid using (setoid; _∙_) +open import Level using (0ℓ) +open import Relation.Binary using (Rel; Setoid) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_) + +open Functor +open Func + +_≈_ : {X Y : Setoid 0ℓ 0ℓ} → Rel (X ⟶ₛ Y) 0ℓ +_≈_ {X} {Y} = Setoid._≈_ (setoid X Y) +infixr 4 _≈_ + +private + variable A B C : ℕ + +-- action on objects (Subset n) +Subsetₛ : ℕ → Setoid 0ℓ 0ℓ +Subsetₛ = ≋-setoid (≡.setoid Bool) + +-- action of Preimage on morphisms (contravariant) +Preimage₁ : (Fin A → Fin B) → Subsetₛ B ⟶ₛ Subsetₛ A +to (Preimage₁ f) i = i ∘ f +cong (Preimage₁ f) x≗y = x≗y ∘ f + +-- Preimage respects identity +Preimage-identity : Preimage₁ id ≈ Id (Subsetₛ A) +Preimage-identity {A} = Setoid.refl (Subsetₛ A) + +-- Preimage flips composition +Preimage-homomorphism + : {A B C : ℕ} + (f : Fin A → Fin B) + (g : Fin B → Fin C) + → Preimage₁ (g ∘ f) ≈ Preimage₁ f ∙ Preimage₁ g +Preimage-homomorphism {A} _ _ = Setoid.refl (Subsetₛ A) + +-- Preimage respects equality +Preimage-resp-≈ + : {f g : Fin A → Fin B} + → f ≗ g + → Preimage₁ f ≈ Preimage₁ g +Preimage-resp-≈ f≗g {v} = ≡.cong v ∘ f≗g + +-- the Preimage functor +Preimage : Functor Natop (Setoids 0ℓ 0ℓ) +F₀ Preimage = Subsetₛ +F₁ Preimage = Preimage₁ +identity Preimage = Preimage-identity +homomorphism Preimage {f = f} {g} {v} = Preimage-homomorphism g f {v} +F-resp-≈ Preimage = Preimage-resp-≈ diff --git a/Functor/Instance/Nat/Pull.agda b/Functor/Instance/Nat/Pull.agda new file mode 100644 index 0000000..5b74399 --- /dev/null +++ b/Functor/Instance/Nat/Pull.agda @@ -0,0 +1,63 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Instance.Nat.Pull where + +open import Categories.Category.Instance.Nat using (Natop) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Functor using (Functor) +open import Data.Fin.Base using (Fin) +open import Data.Nat.Base using (ℕ) +open import Function.Base using (id; _∘_) +open import Function.Bundles using (Func; _⟶ₛ_) +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Setoid using (setoid; _∙_) +open import Level using (0ℓ) +open import Relation.Binary using (Rel; Setoid) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_) +open import Data.Circuit.Value using (Value) +open import Data.System.Values Value using (Values) + +open Functor +open Func + +_≈_ : {X Y : Setoid 0ℓ 0ℓ} → Rel (X ⟶ₛ Y) 0ℓ +_≈_ {X} {Y} = Setoid._≈_ (setoid X Y) +infixr 4 _≈_ + +private + variable A B C : ℕ + + +-- action on objects is Values n (Vector Value n) + +-- action of Pull on morphisms (contravariant) +Pull₁ : (Fin A → Fin B) → Values B ⟶ₛ Values A +to (Pull₁ f) i = i ∘ f +cong (Pull₁ f) x≗y = x≗y ∘ f + +-- Pull respects identity +Pull-identity : Pull₁ id ≈ Id (Values A) +Pull-identity {A} = Setoid.refl (Values A) + +-- Pull flips composition +Pull-homomorphism + : {A B C : ℕ} + (f : Fin A → Fin B) + (g : Fin B → Fin C) + → Pull₁ (g ∘ f) ≈ Pull₁ f ∙ Pull₁ g +Pull-homomorphism {A} _ _ = Setoid.refl (Values A) + +-- Pull respects equality +Pull-resp-≈ + : {f g : Fin A → Fin B} + → f ≗ g + → Pull₁ f ≈ Pull₁ g +Pull-resp-≈ f≗g {v} = ≡.cong v ∘ f≗g + +-- the Pull functor +Pull : Functor Natop (Setoids 0ℓ 0ℓ) +F₀ Pull = Values +F₁ Pull = Pull₁ +identity Pull = Pull-identity +homomorphism Pull {f = f} {g} {v} = Pull-homomorphism g f {v} +F-resp-≈ Pull = Pull-resp-≈ diff --git a/Functor/Instance/Nat/Push.agda b/Functor/Instance/Nat/Push.agda new file mode 100644 index 0000000..53d0bef --- /dev/null +++ b/Functor/Instance/Nat/Push.agda @@ -0,0 +1,64 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Instance.Nat.Push where + +open import Categories.Functor using (Functor) +open import Categories.Category.Instance.Nat using (Nat) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Data.Circuit.Merge using (merge; merge-cong₁; merge-cong₂; merge-⁅⁆; merge-preimage) +open import Data.Fin.Base using (Fin) +open import Data.Fin.Preimage using (preimage; preimage-cong₁) +open import Data.Nat.Base using (ℕ) +open import Data.Subset.Functional using (⁅_⁆) +open import Function.Base using (id; _∘_) +open import Function.Bundles using (Func; _⟶ₛ_) +open import Function.Construct.Identity using () renaming (function to Id) +open import Function.Construct.Setoid using (setoid; _∙_) +open import Level using (0ℓ) +open import Relation.Binary using (Rel; Setoid) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_) +open import Data.Circuit.Value using (Value) +open import Data.System.Values Value using (Values) + +open Func +open Functor + +private + variable A B C : ℕ + +_≈_ : {X Y : Setoid 0ℓ 0ℓ} → Rel (X ⟶ₛ Y) 0ℓ +_≈_ {X} {Y} = Setoid._≈_ (setoid X Y) +infixr 4 _≈_ + +-- action of Push on objects is Values n (Vector Value n) + +-- action of Push on morphisms (covariant) +Push₁ : (Fin A → Fin B) → Values A ⟶ₛ Values B +to (Push₁ f) v = merge v ∘ preimage f ∘ ⁅_⁆ +cong (Push₁ f) x≗y = merge-cong₁ x≗y ∘ preimage f ∘ ⁅_⁆ + +-- Push respects identity +Push-identity : Push₁ id ≈ Id (Values A) +Push-identity {_} {v} = merge-⁅⁆ v + +-- Push respects composition +Push-homomorphism + : {f : Fin A → Fin B} + {g : Fin B → Fin C} + → Push₁ (g ∘ f) ≈ Push₁ g ∙ Push₁ f +Push-homomorphism {f = f} {g} {v} = merge-preimage f v ∘ preimage g ∘ ⁅_⁆ + +-- Push respects equality +Push-resp-≈ + : {f g : Fin A → Fin B} + → f ≗ g + → Push₁ f ≈ Push₁ g +Push-resp-≈ f≗g {v} = merge-cong₂ v ∘ preimage-cong₁ f≗g ∘ ⁅_⁆ + +-- the Push functor +Push : Functor Nat (Setoids 0ℓ 0ℓ) +F₀ Push = Values +F₁ Push = Push₁ +identity Push = Push-identity +homomorphism Push = Push-homomorphism +F-resp-≈ Push = Push-resp-≈ diff --git a/Functor/Instance/Nat/System.agda b/Functor/Instance/Nat/System.agda new file mode 100644 index 0000000..730f90d --- /dev/null +++ b/Functor/Instance/Nat/System.agda @@ -0,0 +1,96 @@ +{-# OPTIONS --without-K --safe #-} + +module Functor.Instance.Nat.System {ℓ} where + +open import Categories.Category.Instance.Nat using (Nat) +open import Categories.Category.Instance.Setoids using (Setoids) +open import Categories.Functor.Core using (Functor) +open import Data.Circuit.Value using (Value) +open import Data.Fin.Base using (Fin) +open import Data.Nat.Base using (ℕ) +open import Data.Product.Base using (_,_; _×_) +open import Data.System using (System; ≤-System; Systemₛ) +open import Data.System.Values Value using (module ≋) +open import Function.Bundles using (Func; _⟶ₛ_) +open import Function.Base using (id; _∘_) +open import Function.Construct.Setoid using (_∙_) +open import Functor.Instance.Nat.Pull using (Pull₁; Pull-resp-≈) +open import Functor.Instance.Nat.Push using (Push₁; Push-identity; Push-homomorphism; Push-resp-≈) +open import Level using (suc) +open import Relation.Binary.PropositionalEquality as ≡ using (_≗_) + +-- import Relation.Binary.Reasoning.Setoid as ≈-Reasoning +import Function.Construct.Identity as Id + +open Func +open ≤-System +open Functor + +private + variable A B C : ℕ + +map : (Fin A → Fin B) → System {ℓ} A → System B +map f X = record + { S = S + ; fₛ = fₛ ∙ Pull₁ f + ; fₒ = Push₁ f ∙ fₒ + } + where + open System X + +≤-cong : (f : Fin A → Fin B) {X Y : System A} → ≤-System Y X → ≤-System (map f Y) (map f X) +⇒S (≤-cong f x≤y) = ⇒S x≤y +≗-fₛ (≤-cong f x≤y) = ≗-fₛ x≤y ∘ to (Pull₁ f) +≗-fₒ (≤-cong f x≤y) = cong (Push₁ f) ∘ ≗-fₒ x≤y + +System₁ : (Fin A → Fin B) → Systemₛ A ⟶ₛ Systemₛ B +to (System₁ f) = map f +cong (System₁ f) (x≤y , y≤x) = ≤-cong f x≤y , ≤-cong f y≤x + +id-x≤x : {X : System A} → ≤-System (map id X) X +⇒S (id-x≤x) = Id.function _ +≗-fₛ (id-x≤x {_} {x}) i s = System.refl x +≗-fₒ (id-x≤x {A} {x}) s = Push-identity + +x≤id-x : {x : System A} → ≤-System x (map id x) +⇒S x≤id-x = Id.function _ +≗-fₛ (x≤id-x {A} {x}) i s = System.refl x +≗-fₒ (x≤id-x {A} {x}) s = ≋.sym Push-identity + + +System-homomorphism + : {f : Fin A → Fin B} + {g : Fin B → Fin C} + {X : System A} + → ≤-System (map (g ∘ f) X) (map g (map f X)) × ≤-System (map g (map f X)) (map (g ∘ f) X) +System-homomorphism {f = f} {g} {X} = left , right + where + open System X + left : ≤-System (map (g ∘ f) X) (map g (map f X)) + left .⇒S = Id.function S + left .≗-fₛ i s = refl + left .≗-fₒ s = Push-homomorphism + right : ≤-System (map g (map f X)) (map (g ∘ f) X) + right .⇒S = Id.function S + right .≗-fₛ i s = refl + right .≗-fₒ s = ≋.sym Push-homomorphism + +System-resp-≈ + : {f g : Fin A → Fin B} + → f ≗ g + → {X : System A} + → (≤-System (map f X) (map g X)) × (≤-System (map g X) (map f X)) +System-resp-≈ {A} {B} {f = f} {g} f≗g {X} = both f≗g , both (≡.sym ∘ f≗g) + where + open System X + both : {f g : Fin A → Fin B} → f ≗ g → ≤-System (map f X) (map g X) + both f≗g .⇒S = Id.function S + both f≗g .≗-fₛ i s = cong fₛ (Pull-resp-≈ f≗g {i}) + both {f} {g} f≗g .≗-fₒ s = Push-resp-≈ f≗g + +Sys : Functor Nat (Setoids (suc ℓ) ℓ) +Sys .F₀ = Systemₛ +Sys .F₁ = System₁ +Sys .identity = id-x≤x , x≤id-x +Sys .homomorphism {x = X} = System-homomorphism {X = X} +Sys .F-resp-≈ = System-resp-≈ |
