├── .gitignore ├── .mailmap ├── AUTHORS ├── Categories ├── 2-Category.agda ├── 2-Category │ └── Categories.agda ├── Adjunction.agda ├── Adjunction │ ├── Composition.agda │ └── CompositionLaws.agda ├── Agda.agda ├── Agda │ └── ISetoids │ │ ├── Cocomplete.agda │ │ ├── Cocomplete │ │ └── Helpers.agda │ │ └── Complete.agda ├── Arrow.agda ├── Bicategory.agda ├── Bicategory │ └── Helpers.agda ├── Bifunctor.agda ├── Bifunctor │ └── NaturalTransformation.agda ├── Categories.agda ├── Categories │ └── Products.agda ├── Category.agda ├── CategorySolver.agda ├── Closed.agda ├── Cocone.agda ├── Cocones.agda ├── Coend.agda ├── Coequalizer.agda ├── Colimit.agda ├── Comma.agda ├── Comma │ └── Functors.agda ├── Comonad.agda ├── Comonad │ └── Cofree.agda ├── Cone.agda ├── Cones.agda ├── Congruence.agda ├── Coproduct.agda ├── DinaturalTransformation.agda ├── Discrete.agda ├── End.agda ├── Enriched.agda ├── Equalizer.agda ├── Equivalence │ └── Strong.agda ├── Fam.agda ├── Fibration.agda ├── Free.agda ├── Free │ ├── Core.agda │ └── Functor.agda ├── Functor.agda ├── Functor │ ├── Algebra.agda │ ├── Algebras.agda │ ├── Coalgebra.agda │ ├── Coalgebras.agda │ ├── Constant.agda │ ├── Core.agda │ ├── Diagonal.agda │ ├── Discrete.agda │ ├── Equivalence │ │ └── Strong.agda │ ├── Hom.agda │ ├── Monoidal.agda │ ├── Product.agda │ ├── Properties.agda │ └── Representable.agda ├── FunctorCategory.agda ├── Globe.agda ├── GlobularSet.agda ├── Graphs.agda ├── Graphs │ └── Underlying.agda ├── Grothendieck.agda ├── Groupoid.agda ├── Groupoid │ ├── Coproduct.agda │ └── Product.agda ├── Lan.agda ├── Lift.agda ├── Limit.agda ├── Monad.agda ├── Monad │ ├── Algebra.agda │ ├── Algebras.agda │ ├── Coproduct.agda │ ├── EilenbergMoore.agda │ ├── Kleisli.agda │ └── Strong.agda ├── Monoidal.agda ├── Monoidal │ ├── Braided.agda │ ├── Braided │ │ └── Helpers.agda │ ├── Cartesian.agda │ ├── Cartesian │ │ └── Pentagon.agda │ ├── CartesianClosed.agda │ ├── Closed.agda │ ├── Helpers.agda │ ├── IntConstruction.agda │ ├── Symmetric.agda │ └── Traced.agda ├── Morphism │ ├── Cartesian.agda │ └── Indexed.agda ├── Morphisms.agda ├── NaturalIsomorphism.agda ├── NaturalTransformation.agda ├── NaturalTransformation │ └── Core.agda ├── Normalise.agda ├── Object │ ├── BinaryCoproducts.agda │ ├── BinaryProducts.agda │ ├── BinaryProducts │ │ ├── Abstract.agda │ │ └── N-ary.agda │ ├── Coproduct.agda │ ├── Coproducts.agda │ ├── Exponential.agda │ ├── Exponentiating.agda │ ├── Exponentiating │ │ ├── Adjunction.agda │ │ └── Functor.agda │ ├── Indexed.agda │ ├── IndexedProduct.agda │ ├── IndexedProducts.agda │ ├── Initial.agda │ ├── Product.agda │ ├── Product │ │ └── Morphisms.agda │ ├── Products.agda │ ├── Products │ │ ├── N-ary.agda │ │ └── Properties.agda │ ├── SubobjectClassifier.agda │ ├── Terminal.agda │ ├── Terminal │ │ ├── Exponentials.agda │ │ ├── Exponentiating.agda │ │ └── Products.agda │ └── Zero.agda ├── Opposite.agda ├── Power.agda ├── Power │ ├── Functorial.agda │ └── NaturalTransformation.agda ├── Preorder.agda ├── Presheaf.agda ├── Presheaves.agda ├── Product.agda ├── Product │ ├── Projections.agda │ └── Properties.agda ├── Profunctor.agda ├── Pullback.agda ├── Pushout.agda ├── Ran.agda ├── RigCategory.agda ├── Slice.agda ├── Square.agda ├── SubCategory.agda ├── Support.agda ├── Support │ ├── EqReasoning.agda │ ├── Equivalence.agda │ ├── Experimental.agda │ ├── FinSet.agda │ ├── IProduct.agda │ ├── Nat.agda │ ├── PropositionalEquality.agda │ ├── SetoidFunctions.agda │ ├── SetoidPi.agda │ ├── StarEquality.agda │ └── ZigZag.agda ├── Terminal.agda ├── Topos.agda ├── WithFamilies.agda └── Yoneda.agda ├── Everything.agda ├── Graphs ├── Graph.agda └── GraphMorphism.agda ├── LICENSE ├── README └── categories.agda-lib /.gitignore: -------------------------------------------------------------------------------- 1 | .DS_Store 2 | *.agdai 3 | *.agda~ 4 | \#*\# 5 | .\#* 6 | -------------------------------------------------------------------------------- /.mailmap: -------------------------------------------------------------------------------- 1 | ## This file allows joining different accounts of a single person. 2 | ## Cf for instance: git shortlog -nse. More details via: man git shortlog 3 | 4 | # having the same name name on a line will fix capitalization 5 | Daniel Peebles Daniel Peebles 6 | James Cook james.cook 7 | James Deikun James Deikun 8 | 9 | -------------------------------------------------------------------------------- /AUTHORS: -------------------------------------------------------------------------------- 1 | In no particular order: 2 | Daniel Peebles 3 | James Deikun 4 | Ulf Norell 5 | Dan Doel 6 | Darius Jahandarie 7 | James Cook 8 | -------------------------------------------------------------------------------- /Categories/2-Category/Categories.agda: -------------------------------------------------------------------------------- 1 | module Categories.2-Category.Categories where 2 | 3 | open import Level 4 | open import Data.Product 5 | 6 | open import Categories.Category 7 | open import Categories.2-Category 8 | open import Categories.Functor using (module Functor) renaming (id to idF; _∘_ to _∘F_) 9 | open import Categories.Functor.Constant 10 | open import Categories.FunctorCategory 11 | open import Categories.Bifunctor using (Bifunctor) 12 | open import Categories.NaturalTransformation 13 | open import Categories.Product using (Product) 14 | 15 | Categories : ∀ o ℓ e → 2-Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) (o ⊔ e) 16 | Categories o ℓ e = record 17 | { Obj = Category o ℓ e 18 | ; _⇒_ = Functors 19 | ; id = λ {A} → Constant {D = Functors A A} idF 20 | ; —∘— = my-∘ 21 | ; assoc = {!!} 22 | ; identityˡ = λ {A B} η → Heterogeneous.≡⇒∼ {C = Functors A B} {f = proj₂ η ∘₁ id} {proj₂ η} (Category.identityʳ B) 23 | ; identityʳ = λ {A B f g} η → Heterogeneous.≡⇒∼ {_} {_} {_} {C = Functors A B} {proj₁ f} {proj₁ g} {f = record { 24 | η = 25 | λ X → 26 | Category._∘_ B (Functor.F₁ (proj₁ g) (Category.id A)) 27 | (NaturalTransformation.η (proj₁ η) X); 28 | commute = λ {a} {b} h → let open Category.HomReasoning B in begin 29 | B [ B [ Functor.F₁ (proj₁ g) (Category.id A) ∘ NaturalTransformation.η (proj₁ η) b ] ∘ {!.Categories.NaturalTransformation.Core.F.F₁ (proj₁ f) (proj₁ g) h!} ] 30 | ↓⟨ {!!} ⟩ 31 | B [ Functor.F₁ {!!} {!!} ∘ B [ Functor.F₁ (proj₁ g) (Category.id A) ∘ NaturalTransformation.η (proj₁ η) a ] ] 32 | ∎ } 33 | } {proj₁ {_} {_} {_} {_} η} (Category.Equiv.trans B (Category.∘-resp-≡ˡ B (Functor.identity (proj₁ g))) (Category.identityˡ B)) 34 | } 35 | where 36 | my-∘ : {A B C : Category o ℓ e} → Bifunctor (Functors B C) (Functors A B) (Functors A C) 37 | my-∘ {A} {B} {C} = record 38 | { F₀ = uncurry′ _∘F_ 39 | ; F₁ = uncurry′ _∘₀_ 40 | ; identity = λ {Fs} → Category.Equiv.trans C (Category.identityʳ C) (Functor.identity (proj₁ Fs)) 41 | ; homomorphism = λ {_ _ _ f g} → Category.Equiv.sym C (interchange {α = proj₁ g} {proj₂ g} {proj₁ f} {proj₂ f}) 42 | ; F-resp-≡ = λ {_ _ f g} → my-resp {f = f} {g} 43 | } 44 | where 45 | PF = Product (Functors B C) (Functors A B) 46 | 47 | .my-resp : ∀ {F G} {f g : PF [ F , G ]} → PF [ f ≡ g ] → Functors A C [ uncurry′ _∘₀_ f ≡ uncurry _∘₀_ g ] 48 | my-resp {f = f₁ , f₂} {g₁ , g₂} (f₁≡g₁ , f₂≡g₂) = ∘₀-resp-≡ {f = f₁} {g₁} {f₂} {g₂} f₁≡g₁ f₂≡g₂ -------------------------------------------------------------------------------- /Categories/Agda.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Agda where 3 | 4 | open import Level 5 | import Relation.Binary 6 | import Function.Equality 7 | 8 | import Categories.Support.Equivalence 9 | import Categories.Support.SetoidFunctions 10 | 11 | open import Categories.Support.PropositionalEquality 12 | open import Categories.Category 13 | open import Categories.Functor using (Functor) 14 | 15 | Sets : ∀ o → Category _ _ _ 16 | Sets o = record 17 | { Obj = Set o 18 | ; _⇒_ = λ d c → d → c 19 | ; _≡_ = λ f g → ∀ {x} → f x ≣ g x 20 | 21 | ; _∘_ = λ f g x → f (g x) 22 | ; id = λ x → x 23 | 24 | ; assoc = ≣-refl 25 | ; identityˡ = ≣-refl 26 | ; identityʳ = ≣-refl 27 | ; equiv = record { refl = ≣-refl; sym = s; trans = t } 28 | ; ∘-resp-≡ = ∘-resp-≡′ 29 | } 30 | where 31 | s : {A B : Set o} → {i j : A → B} → ({x : A} → i x ≣ j x) → {x : A} → j x ≣ i x 32 | s f {x} rewrite f {x} = ≣-refl 33 | 34 | t : {A B : Set o} {i j k : A → B} → ({x : A} → i x ≣ j x) → ({x : A} → j x ≣ k x) → {x : A} → i x ≣ k x 35 | t f g {x} rewrite f {x} | g {x} = ≣-refl 36 | 37 | ∘-resp-≡′ : {A B C : Set o} {f h : B → C} {g i : A → B} → 38 | (∀ {x} → f x ≣ h x) → 39 | (∀ {x} → g x ≣ i x) → 40 | (∀ {x} → f (g x) ≣ h (i x)) 41 | ∘-resp-≡′ {g = g} f≡h g≡i {x} rewrite f≡h {g x} | g≡i {x} = ≣-refl 42 | 43 | -- use standard library setoids here, not our special irrelevant ones 44 | Setoids : ∀ c ℓ → Category (suc (ℓ ⊔ c)) (ℓ ⊔ c) (ℓ ⊔ c) 45 | Setoids c ℓ = record 46 | { Obj = Setoid c ℓ 47 | ; _⇒_ = _⟶_ 48 | ; _≡_ = λ {X} {Y} → _≡′_ {X} {Y} 49 | ; _∘_ = _∘′_ 50 | ; id = id′ 51 | ; assoc = λ {_} {_} {_} {D} → Setoid.refl D 52 | ; identityˡ = λ {_} {B} → Setoid.refl B 53 | ; identityʳ = λ {_} {B} → Setoid.refl B 54 | ; equiv = λ {A} {B} → record 55 | { refl = Setoid.refl B 56 | ; sym = λ f → Setoid.sym B f 57 | ; trans = λ f g → Setoid.trans B f g 58 | } 59 | ; ∘-resp-≡ = λ {A} {B} {C} {f} {h} {g} {i} → ∘-resp-≡′ {A} {B} {C} {f} {h} {g} {i} 60 | } 61 | where 62 | infix 4 _≡′_ 63 | open Function.Equality using (_⟨$⟩_; cong; _⟶_) renaming (_∘_ to _∘′_; id to id′) 64 | open Relation.Binary using (Setoid; module Setoid; Rel) 65 | 66 | _≡′_ : ∀ {X Y} → Rel (X ⟶ Y) _ 67 | _≡′_ {X} {Y} f g = ∀ {x} → Setoid._≈_ Y (f ⟨$⟩ x) (g ⟨$⟩ x) 68 | 69 | .∘-resp-≡′ : ∀ {A B C} {f h : B ⟶ C} {g i : A ⟶ B} → f ≡′ h → g ≡′ i → f ∘′ g ≡′ h ∘′ i 70 | ∘-resp-≡′ {C = C} {h = h} f≡h g≡i {x} = Setoid.trans C f≡h (cong h g≡i) 71 | 72 | -- setoids with irrelevant equality 73 | ISetoids : ∀ c ℓ → Category (suc (ℓ ⊔ c)) (ℓ ⊔ c) (ℓ ⊔ c) 74 | ISetoids c ℓ = record 75 | { Obj = Setoid c ℓ 76 | ; _⇒_ = _⟶_ 77 | ; _≡_ = λ {A B} → Setoid._≈_ (A ⇨ B) 78 | ; _∘_ = _∙_ 79 | ; id = id′ 80 | ; assoc = λ {A B C D} {f g h} → 81 | cong (h ∙ g ∙ f) 82 | ; identityˡ = λ {A B f} → cong f 83 | ; identityʳ = λ {A B f} → cong f 84 | ; equiv = λ {A B} → Setoid.isEquivalence (A ⇨ B) 85 | ; ∘-resp-≡ = λ f≡h g≡i x≡y → f≡h (g≡i x≡y) 86 | } 87 | where 88 | open Relation.Binary using (Rel) 89 | open Categories.Support.Equivalence 90 | open Categories.Support.SetoidFunctions renaming (id to id′) 91 | 92 | 93 | Lift-IS : ∀ {c ℓ} a b → Functor (ISetoids c ℓ) (ISetoids (c ⊔ a) (ℓ ⊔ b)) 94 | Lift-IS {c} {ℓ} a b = record { 95 | F₀ = Lift-setoid {c} {ℓ} {a} {b}; 96 | F₁ = λ f → record { _⟨$⟩_ = λ x → lift (f ⟨$⟩ (lower x)); cong = λ eq → lift (cong f (lower eq)) }; 97 | identity = λ x → x; 98 | homomorphism = λ {_} {_} {_} {f} {g} eq → lift (cong g (cong f (lower eq))); 99 | F-resp-≡ = λ eq₀ eq₁ → lift (eq₀ (lower eq₁)) } 100 | where 101 | open import Categories.Support.Equivalence 102 | open Categories.Support.SetoidFunctions renaming (id to id′) 103 | 104 | -------------------------------------------------------------------------------- /Categories/Agda/ISetoids/Cocomplete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Agda.ISetoids.Cocomplete where 3 | 4 | open import Level 5 | open import Relation.Binary using (Setoid; module Setoid; Preorder; module Preorder; Rel; _=[_]⇒_) 6 | open import Data.Product using (Σ; _,_) 7 | -- import Relation.Binary.EqReasoning as EqReasoning 8 | 9 | open import Categories.Support.Equivalence using (module I→R-Wrapper; setoid-i→r) renaming (Setoid to ISetoid; module Setoid to ISetoid) 10 | open import Categories.Support.SetoidFunctions using (module _⟶_) 11 | open import Categories.Support.PropositionalEquality 12 | import Categories.Support.ZigZag as ZigZag 13 | 14 | open import Categories.Category 15 | open import Categories.Functor 16 | open import Categories.Agda 17 | open import Categories.Colimit 18 | open import Categories.Object.Initial 19 | open import Categories.Cocones 20 | open import Categories.Cocone 21 | open import Categories.Agda.ISetoids.Cocomplete.Helpers 22 | 23 | open I→R-Wrapper 24 | 25 | ISetoidsCocomplete : ∀ {o ℓ e c ℓ′} → Cocomplete o ℓ e (ISetoids (c ⊔ ding (o ⊔ ℓ ⊔ e)) (c ⊔ ℓ′ ⊔ ding (o ⊔ ℓ ⊔ e))) 26 | ISetoidsCocomplete {o} {ℓ} {e} {c} {cℓ} = record { colimit = colimit } 27 | where 28 | c′ = c ⊔ ding (o ⊔ ℓ ⊔ e) 29 | ℓ′ = c ⊔ cℓ ⊔ ding (o ⊔ ℓ ⊔ e) 30 | C = ISetoids c′ ℓ′ 31 | colimit : ∀ {J : Category o ℓ e} (F : Functor J C) → Colimit F 32 | colimit {J} F = record { initial = my-initial-cocone } 33 | where 34 | module J = Category J 35 | open Functor F 36 | open ISetoid 37 | open _⟶_ 38 | 39 | open ColimitCocone {c = c} {cℓ} {J} F 40 | 41 | .my-!-unique : (A : Cocone F) (φ : CoconeMorphism ⊥ A) (x : vertex-carrier) → (_≈_ (Cocone.N A) (CoconeMorphism.f (! {A}) ⟨$⟩ x) (CoconeMorphism.f φ ⟨$⟩ x)) 42 | my-!-unique A φ (X , x) = sym (Cocone.N A) (CoconeMorphism.commute φ (refl (F₀ X))) 43 | 44 | my-initial-cocone : Initial (Cocones F) 45 | my-initial-cocone = record 46 | { ⊥ = ⊥ 47 | ; ! = ! 48 | ; !-unique = λ {A} φ {x} {y} x≡y → trans (Cocone.N A) (my-!-unique A φ x) (cong (CoconeMorphism.f φ) x≡y) 49 | } -------------------------------------------------------------------------------- /Categories/Agda/ISetoids/Cocomplete/Helpers.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Agda.ISetoids.Cocomplete.Helpers where 3 | 4 | open import Level 5 | open import Relation.Binary using (Setoid; module Setoid; Preorder; module Preorder; Rel; _=[_]⇒_) 6 | open import Data.Product using (Σ; _,_; Σ-syntax) 7 | -- import Relation.Binary.EqReasoning as EqReasoning 8 | 9 | open import Categories.Support.Equivalence using (module I→R-Wrapper; setoid-i→r; squash) renaming (Setoid to ISetoid; module Setoid to ISetoid) 10 | open import Categories.Support.SetoidFunctions using (module _⟶_) 11 | open import Categories.Support.PropositionalEquality 12 | import Categories.Support.ZigZag as ZigZag 13 | 14 | open import Categories.Category 15 | open import Categories.Functor 16 | open import Categories.Agda 17 | open import Categories.Colimit 18 | open import Categories.Object.Initial 19 | open import Categories.Cocones 20 | open import Categories.Cocone 21 | 22 | open I→R-Wrapper 23 | 24 | ding : Level → Level 25 | ding ℓ = ℓ 26 | 27 | module ColimitCocone {o ℓ e c ℓ′} {J : Category o ℓ e} (F : Functor J (ISetoids (c ⊔ ding (o ⊔ ℓ ⊔ e)) (c ⊔ ℓ′ ⊔ ding (o ⊔ ℓ ⊔ e)))) where 28 | c′ = c ⊔ ding (o ⊔ ℓ ⊔ e) 29 | ℓ″ = c ⊔ ℓ′ ⊔ ding (o ⊔ ℓ ⊔ e) 30 | C = ISetoids c′ ℓ″ 31 | D = Cocones F 32 | module J = Category J 33 | open Functor F 34 | open ISetoid 35 | open _⟶_ 36 | 37 | vertex-carrier = Σ[ x ∈ J.Obj ] Carrier (F₀ x) 38 | 39 | -- this lets us build a preorder out of an irrelevant setoid, since there 40 | -- are no irrelevant preorders 41 | dot-setoid : (j : J.Obj) → Setoid c′ ℓ″ 42 | dot-setoid j = setoid-i→r (F₀ j) 43 | module DotSetoid (j : J.Obj) = Setoid (dot-setoid j) 44 | open DotSetoid using () renaming (_≈_ to _[_≈_]) 45 | 46 | -- _↝_ means that an arrow in the diagram directly constrains the two 47 | -- objects of vertex-carrier to be equal. completing this to an 48 | -- equivalence relation gives us the setoid we need 49 | _↝_ : (x y : vertex-carrier) → Set ℓ″ 50 | (X , x) ↝ (Y , y) = Σ[ f ∈ J [ X , Y ] ] (Y [ (F₁ f ⟨$⟩ x) ≈ y ]) 51 | 52 | ↝-preorder : Preorder c′ c′ ℓ″ 53 | ↝-preorder = record 54 | { Carrier = vertex-carrier 55 | ; _≈_ = _≣_ 56 | ; _∼_ = _↝_ 57 | ; isPreorder = record 58 | { isEquivalence = ≣-isEquivalence 59 | ; reflexive = ↝-reflexive 60 | ; trans = ↝-trans 61 | } 62 | } 63 | where 64 | ↝-refl : ∀ {i} → (i ↝ i) 65 | ↝-refl {X , x} = J.id , squash (identity (refl (F₀ X))) 66 | 67 | ↝-reflexive : ∀ {i j} → (i ≣ j) → (i ↝ j) 68 | ↝-reflexive ≣-refl = ↝-refl 69 | 70 | -- the nice proof. pretty proof! 71 | {- 72 | ↝-trans : ∀ {i j k} → (i ↝ j) → (j ↝ k) → (i ↝ k) 73 | ↝-trans {_ , i} {_ , j} {k₁ , k} (f , squash fi≈j) (g , gj≈k) = 74 | J [ g ∘ f ] , ( 75 | begin 76 | F₁ (J [ g ∘ f ]) ⟨$⟩ i 77 | ≈⟨ squash homomorphism ⟩ 78 | F₁ g ⟨$⟩ (F₁ f ⟨$⟩ i) 79 | ≈⟨ squash (cong (F₁ g) fi≈j) ⟩ 80 | F₁ g ⟨$⟩ j 81 | ≈⟨ gj≈k ⟩ 82 | k 83 | ∎ ) 84 | where 85 | open EqReasoning (dot-setoid k₁) 86 | -} 87 | 88 | -- the ugly proof that loads before the sun burns out. 89 | ↝-trans : ∀ {i j k} → (i ↝ j) → (j ↝ k) → (i ↝ k) 90 | ↝-trans {X , _} {k = Z , _} (f , squash fx≈y) (g , gy≈z) = 91 | J [ g ∘ f ] , DotSetoid.trans Z (squash (trans (F₀ Z) (homomorphism (refl (F₀ X))) (cong (F₁ g) fx≈y))) gy≈z 92 | 93 | vertex = ZigZag.isetoid ↝-preorder 94 | 95 | ⊥ : Cocone F 96 | ⊥ = record 97 | { N = vertex 98 | ; ψ = λ X → record 99 | { _⟨$⟩_ = _,_ X 100 | ; cong = λ pf → ZigZag.disorient (ZigZag.slish (J.id , squash (identity pf))) 101 | -- cong would be easier if i'd used the proper dotwise equality for 102 | -- ↝-preorder above ... since this is basically the proof ↝ respects it 103 | } 104 | ; commute = λ {X Y} f {x : _} {y : _} (x≈y : _) → ZigZag.disorient (ZigZag.slish (f , (squash (cong (F₁ f) x≈y)))) 105 | } 106 | 107 | ! : {A : Cocone F} → CoconeMorphism {F = F} ⊥ A 108 | ! {A} = record 109 | { f = record 110 | { _⟨$⟩_ = my-f 111 | ; cong = ZigZag.iminimal ↝-preorder A.N my-f my-precong 112 | } 113 | ; commute = λ {X x y} x≈y → cong (A.ψ X) x≈y 114 | } 115 | where 116 | module A = Cocone A 117 | my-f : vertex-carrier → Carrier A.N 118 | my-f (X , x) = A.ψ X ⟨$⟩ x 119 | 120 | .my-precong : _↝_ =[ my-f ]⇒ (_≈_ A.N) 121 | my-precong {X , x} {Y , y} (f , squash fx≈y) = trans A.N (A.commute f (refl (F₀ X))) (cong (A.ψ Y) (irr fx≈y)) 122 | -------------------------------------------------------------------------------- /Categories/Agda/ISetoids/Complete.agda: -------------------------------------------------------------------------------- 1 | module Categories.Agda.ISetoids.Complete where 2 | 3 | open import Level 4 | open import Relation.Binary using (Setoid; module Setoid; Preorder; module Preorder; Rel; _=[_]⇒_) 5 | open import Data.Product using (Σ; _,_; proj₁; proj₂) 6 | open import Categories.Support.IProduct 7 | -- import Relation.Binary.EqReasoning as EqReasoning 8 | 9 | open import Categories.Support.Equivalence using (module I→R-Wrapper; setoid-i→r; Fam-setoid; ∀[_]-setoid_) 10 | renaming (Setoid to ISetoid; module Setoid to ISetoid) 11 | open import Categories.Support.SetoidFunctions using (module _⟶_) 12 | open import Categories.Support.PropositionalEquality 13 | import Categories.Support.ZigZag as ZigZag 14 | 15 | open import Categories.Category 16 | open import Categories.Functor 17 | import Categories.NaturalTransformation as NT 18 | open import Categories.Agda 19 | open import Categories.Colimit 20 | open import Categories.Object.Initial 21 | open import Categories.Cocones 22 | open import Categories.Cocone 23 | 24 | ISetoidsComplete : ∀ {o ℓ e c ℓ′} → Cocomplete o ℓ e (Category.op (ISetoids (c ⊔ ℓ′ ⊔ (o ⊔ ℓ ⊔ e)) (c ⊔ (o ⊔ ℓ ⊔ e)))) 25 | ISetoidsComplete {o} {ℓ} {e} {c} {cℓ} = record { colimit = colimit } 26 | where 27 | c′ = c ⊔ cℓ ⊔ (o ⊔ ℓ ⊔ e) 28 | ℓ′ = c ⊔ (o ⊔ ℓ ⊔ e) 29 | C = Category.op (ISetoids c′ ℓ′) 30 | colimit : ∀ {J : Category o ℓ e} (F : Functor J C) → Colimit F 31 | colimit {J} F = record { initial = record 32 | { ⊥ = record 33 | { N = Fam-setoid (Σ′ (∀ j → Carrier (F₀ j)) (λ f → ∀ {a b} (h : a J.⇒ b) → (F₀ a ≈ F₁ h ⟨$⟩ f b) (f a))) 34 | (∀[ J.Obj ]-setoid F₀) proj₁′ 35 | ; ψ = λ X → record 36 | { _⟨$⟩_ = λ f → proj₁′ f X 37 | ; cong = λ eq → eq X 38 | } 39 | ; commute = λ {X} {Y} h {f} {g} f≈g → trans (F₀ X) (f≈g X) (sym (F₀ X) (proj₂′ g h)) 40 | } 41 | ; ! = λ {A} → record 42 | { f = record 43 | { _⟨$⟩_ = λ X → (λ j → Cocone.ψ A j ⟨$⟩ X) , (λ {a} {b} h → sym (F₀ a) (Cocone.commute A h (refl (Cocone.N A)))) 44 | ; cong = λ i≈j a → cong (Cocone.ψ A a) i≈j 45 | } 46 | ; commute = λ {X} x≈y → cong (Cocone.ψ A X) x≈y 47 | } 48 | ; !-unique = λ {A} f x≈y a → sym (F₀ a) (CoconeMorphism.commute f (sym (Cocone.N A) x≈y)) 49 | } } 50 | where 51 | module J = Category J 52 | open Functor F 53 | open ISetoid 54 | open _⟶_ 55 | -------------------------------------------------------------------------------- /Categories/Bicategory.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Bicategory where 3 | 4 | open import Level 5 | -- open import Data.Product using (curry; _,_) 6 | -- open import Function using () renaming (_∘_ to _·_) 7 | 8 | open import Categories.Support.PropositionalEquality 9 | 10 | open import Categories.Category 11 | open import Categories.Categories 12 | open import Categories.Object.Terminal 13 | open import Categories.Terminal 14 | open import Categories.Functor using (Functor) renaming (_∘_ to _∘F_; _≡_ to _≡F_; id to idF) 15 | open import Categories.Functor.Constant using (Constant) 16 | open import Categories.Bifunctor using (Bifunctor; reduce-×) 17 | open import Categories.Product using (assocʳ; πˡ; πʳ) 18 | open import Categories.NaturalIsomorphism 19 | open import Categories.Bicategory.Helpers using (module BicategoryHelperFunctors) 20 | 21 | record Bicategory (o ℓ t e : Level) : Set (suc (o ⊔ ℓ ⊔ t ⊔ e)) where 22 | open Terminal (One {ℓ} {t} {e}) 23 | field 24 | Obj : Set o 25 | _⇒_ : (A B : Obj) → Category ℓ t e 26 | id : {A : Obj} → Functor ⊤ (A ⇒ A) 27 | —∘— : {A B C : Obj} → Bifunctor (B ⇒ C) (A ⇒ B) (A ⇒ C) 28 | 29 | _∘_ : {A B C : Obj} {L R : Category ℓ t e} → Functor L (B ⇒ C) → Functor R (A ⇒ B) → Bifunctor L R (A ⇒ C) 30 | _∘_ {A} {B} {C} F G = reduce-× {D₁ = B ⇒ C} {D₂ = A ⇒ B} —∘— F G 31 | 32 | field 33 | λᵤ : {A B : Obj} → NaturalIsomorphism (id ∘ idF {C = A ⇒ B}) (πʳ {C = ⊤} {A ⇒ B}) 34 | ρᵤ : {A B : Obj} → NaturalIsomorphism (idF {C = A ⇒ B} ∘ id) (πˡ {C = A ⇒ B} {⊤}) 35 | α : {A B C D : Obj} → NaturalIsomorphism (idF ∘ —∘—) (((—∘— ∘ idF) ∘F assocʳ (C ⇒ D) (B ⇒ C) (A ⇒ B))) 36 | 37 | private module BHF = BicategoryHelperFunctors(Obj)(_⇒_)(—∘—)(id) 38 | private module EQ = BHF.Coherence(λᵤ)(ρᵤ)(α) 39 | 40 | private module _⇒_ (A B : Obj) = Category (A ⇒ B) 41 | open _⇒_ public using () renaming (Obj to _⇒₁_) 42 | 43 | field 44 | .triangle : {A B C : Obj} (f : A ⇒₁ B) (g : B ⇒₁ C) → EQ.Triangle f g 45 | .pentagon : {A B C D E : Obj} (f : A ⇒₁ B) (g : B ⇒₁ C) (h : C ⇒₁ D) (i : D ⇒₁ E) 46 | → EQ.Pentagon f g h i 47 | 48 | -- do note that most of the "convenience" definitions in the Helpers module should really 49 | -- be here (but usable in both places). Clean that up later. 50 | -------------------------------------------------------------------------------- /Categories/Bicategory/Helpers.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Bicategory.Helpers where 3 | 4 | -- quite a bit of the code below is taken from Categories.2-Category 5 | open import Data.Nat using (_+_) 6 | open import Function using (flip) 7 | open import Data.Product using (_,_; curry) 8 | 9 | open import Categories.Category 10 | open import Categories.Categories using (Categories) 11 | import Categories.Functor 12 | open import Categories.Terminal using (OneC; One; unit) 13 | open import Categories.Object.Terminal using (module Terminal) 14 | open import Categories.Bifunctor hiding (identityˡ; identityʳ; assoc) renaming (id to idF; _≡_ to _≡F_; _∘_ to _∘F_) 15 | open import Categories.NaturalIsomorphism 16 | open import Categories.NaturalTransformation using (NaturalTransformation) renaming (_≡_ to _≡ⁿ_; id to idⁿ) 17 | open import Categories.Product using (Product; assocʳ; πˡ; πʳ) 18 | 19 | module BicategoryHelperFunctors {o ℓ t e} (Obj : Set o) (_⇒_ : (A B : Obj) → Category ℓ t e) 20 | (—⊗— : {A B C : Obj} → Bifunctor (B ⇒ C) (A ⇒ B) (A ⇒ C)) 21 | (id : {A : Obj} → Functor {ℓ} {t} {e} OneC (A ⇒ A)) where 22 | 23 | open Terminal (One {ℓ} {t} {e}) 24 | 25 | _∘_ : {A B C : Obj} {L R : Category ℓ t e} → Functor L (B ⇒ C) → Functor R (A ⇒ B) → Bifunctor L R (A ⇒ C) 26 | _∘_ {A} {B} {C} F G = reduce-× {D₁ = B ⇒ C} {D₂ = A ⇒ B} —⊗— F G 27 | 28 | -- convenience! 29 | module _⇒_ (A B : Obj) = Category (A ⇒ B) 30 | open _⇒_ public using () renaming (Obj to _⇒₁_) 31 | 32 | private module imp⇒ {X Y : Obj} = Category (X ⇒ Y) 33 | open imp⇒ public using () renaming (_⇒_ to _⇒₂_; id to id₂; _∘_ to _•_; _≡_ to _≡₂_) 34 | 35 | id₁ : ∀ {A} → A ⇒₁ A 36 | id₁ {A} = Functor.F₀ (id {A}) unit 37 | 38 | id₁₂ : ∀ {A} → id₁ {A} ⇒₂ id₁ {A} 39 | id₁₂ {A} = id₂ {A = id₁ {A}} 40 | 41 | infixr 9 _∘₁_ 42 | _∘₁_ : ∀ {A B C} → B ⇒₁ C → A ⇒₁ B → A ⇒₁ C 43 | _∘₁_ = curry (Functor.F₀ —⊗—) 44 | 45 | -- horizontal composition 46 | infixr 9 _∘₂_ 47 | _∘₂_ : ∀ {A B C} {g g′ : B ⇒₁ C} {f f′ : A ⇒₁ B} (β : g ⇒₂ g′) (α : f ⇒₂ f′) → (g ∘₁ f) ⇒₂ (g′ ∘₁ f′) 48 | _∘₂_ = curry (Functor.F₁ —⊗—) 49 | 50 | -- left whiskering 51 | infixl 9 _◃_ 52 | _◃_ : ∀ {A B C} {g g′ : B ⇒₁ C} → g ⇒₂ g′ → (f : A ⇒₁ B) → (g ∘₁ f) ⇒₂ (g′ ∘₁ f) 53 | β ◃ f = β ∘₂ id₂ 54 | 55 | -- right whiskering 56 | infixr 9 _▹_ 57 | _▹_ : ∀ {A B C} (g : B ⇒₁ C) → {f f′ : A ⇒₁ B} → f ⇒₂ f′ → (g ∘₁ f) ⇒₂ (g ∘₁ f′) 58 | g ▹ α = id₂ ∘₂ α 59 | 60 | module Coherence 61 | (identityˡ : {A B : Obj} → NaturalIsomorphism (id ∘ idF {C = A ⇒ B}) (πʳ {C = ⊤} {A ⇒ B})) 62 | (identityʳ : {A B : Obj} → NaturalIsomorphism (idF {C = A ⇒ B} ∘ id) (πˡ {C = A ⇒ B})) 63 | (assoc : {A B C D : Obj} → NaturalIsomorphism (idF ∘ —⊗—) ((—⊗— ∘ idF) ∘F assocʳ (C ⇒ D) (B ⇒ C) (A ⇒ B)) ) where 64 | 65 | open Categories.NaturalTransformation.NaturalTransformation 66 | 67 | -- left/right are inverted between in certain situations! 68 | -- private so as to not clash with the ones in Bicategory itself 69 | private 70 | ρᵤ : {A B : Obj} (f : A ⇒₁ B) → (id₁ ∘₁ f) ⇒₂ f 71 | ρᵤ f = η (NaturalIsomorphism.F⇒G identityˡ) (unit , f) 72 | 73 | λᵤ : {A B : Obj} (f : A ⇒₁ B) → (f ∘₁ id₁) ⇒₂ f 74 | λᵤ f = η (NaturalIsomorphism.F⇒G identityʳ) (f , unit) 75 | 76 | α : {A B C D : Obj} (f : A ⇒₁ B) (g : B ⇒₁ C) (h : C ⇒₁ D) → 77 | (h ∘₁ (g ∘₁ f)) ⇒₂ ((h ∘₁ g) ∘₁ f) 78 | α f g h = η (NaturalIsomorphism.F⇒G assoc) (h , g , f) 79 | 80 | -- Triangle identity. Look how closely it matches with the one on 81 | -- http://ncatlab.org/nlab/show/bicategory 82 | Triangle : {A B C : Obj} (f : A ⇒₁ B) (g : B ⇒₁ C) → Set e 83 | Triangle f g = (g ▹ ρᵤ f) ≡₂ ((λᵤ g ◃ f) • (α f id₁ g)) 84 | 85 | -- Pentagon identity. Ditto. 86 | Pentagon : {A B C D E : Obj} (f : A ⇒₁ B) (g : B ⇒₁ C) (h : C ⇒₁ D) (i : D ⇒₁ E) → Set e 87 | Pentagon f g h i = ((α g h i ◃ f) • (α f (h ∘₁ g) i)) • (i ▹ α f g h) ≡₂ (α f g (i ∘₁ h) • α (g ∘₁ f) h i) 88 | -------------------------------------------------------------------------------- /Categories/Bifunctor.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Bifunctor where 3 | 4 | open import Level 5 | open import Data.Product using (_,_; swap) 6 | 7 | open import Categories.Category 8 | 9 | open import Categories.Functor public 10 | open import Categories.Product 11 | 12 | Bifunctor : ∀ {o ℓ e} {o′ ℓ′ e′} {o′′ ℓ′′ e′′} → Category o ℓ e → Category o′ ℓ′ e′ → Category o′′ ℓ′′ e′′ → Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′ ⊔ o′′ ⊔ ℓ′′ ⊔ e′′) 13 | Bifunctor C D E = Functor (Product C D) E 14 | 15 | overlap-× : ∀ {o ℓ e} {o′₁ ℓ′₁ e′₁} {o′₂ ℓ′₂ e′₂} {C : Category o ℓ e} {D₁ : Category o′₁ ℓ′₁ e′₁} {D₂ : Category o′₂ ℓ′₂ e′₂} (H : Bifunctor D₁ D₂ C) {o″ ℓ″ e″} {E : Category o″ ℓ″ e″} (F : Functor E D₁) (G : Functor E D₂) → Functor E C 16 | overlap-× H F G = H ∘ (F ※ G) 17 | 18 | reduce-× : ∀ {o ℓ e} {o′₁ ℓ′₁ e′₁} {o′₂ ℓ′₂ e′₂} {C : Category o ℓ e} {D₁ : Category o′₁ ℓ′₁ e′₁} {D₂ : Category o′₂ ℓ′₂ e′₂} (H : Bifunctor D₁ D₂ C) {o″₁ ℓ″₁ e″₁} {E₁ : Category o″₁ ℓ″₁ e″₁} (F : Functor E₁ D₁) {o″₂ ℓ″₂ e″₂} {E₂ : Category o″₂ ℓ″₂ e″₂} (G : Functor E₂ D₂) → Bifunctor E₁ E₂ C 19 | reduce-× H F G = H ∘ (F ⁂ G) 20 | 21 | flip-bifunctor : ∀ {o ℓ e} {o′ ℓ′ e′} {o′′ ℓ′′ e′′} {C : Category o ℓ e} → {D : Category o′ ℓ′ e′} → {E : Category o′′ ℓ′′ e′′} → Bifunctor C D E → Bifunctor D C E 22 | flip-bifunctor {C = C} {D = D} {E = E} b = _∘_ b (Swap {C = C} {D = D}) 23 | -------------------------------------------------------------------------------- /Categories/Bifunctor/NaturalTransformation.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Bifunctor.NaturalTransformation where 3 | 4 | open import Level 5 | 6 | open import Categories.Category 7 | open import Categories.Bifunctor 8 | open import Categories.Product 9 | 10 | open import Categories.NaturalTransformation public 11 | 12 | -- just for completeness ... 13 | BiNaturalTransformation : ∀ {o ℓ e} {o′ ℓ′ e′} {o′′ ℓ′′ e′′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} {E : Category o′′ ℓ′′ e′′} → Bifunctor C D E → Bifunctor C D E → Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′ ⊔ o′′ ⊔ ℓ′′ ⊔ e′′) 14 | BiNaturalTransformation F G = NaturalTransformation F G 15 | 16 | reduceN-× : ∀ {o ℓ e} {o′₁ ℓ′₁ e′₁} {o′₂ ℓ′₂ e′₂} {C : Category o ℓ e} {D₁ : Category o′₁ ℓ′₁ e′₁} {D₂ : Category o′₂ ℓ′₂ e′₂} (H : Bifunctor D₁ D₂ C) {o″₁ ℓ″₁ e″₁} {E₁ : Category o″₁ ℓ″₁ e″₁} {F F′ : Functor E₁ D₁} (φ : NaturalTransformation F F′) {o″₂ ℓ″₂ e″₂} {E₂ : Category o″₂ ℓ″₂ e″₂} {G G′ : Functor E₂ D₂} (γ : NaturalTransformation G G′) → NaturalTransformation (reduce-× {D₁ = D₁} {D₂ = D₂} H F G) (reduce-× {D₁ = D₁} {D₂ = D₂} H F′ G′) 17 | reduceN-× H φ γ = H ∘ˡ (φ ⁂ⁿ γ) 18 | 19 | overlapN-× : ∀ {o ℓ e} {o′₁ ℓ′₁ e′₁} {o′₂ ℓ′₂ e′₂} {C : Category o ℓ e} {D₁ : Category o′₁ ℓ′₁ e′₁} {D₂ : Category o′₂ ℓ′₂ e′₂} (H : Bifunctor D₁ D₂ C) {o″ ℓ″ e″} {E : Category o″ ℓ″ e″} {F F′ : Functor E D₁} (φ : NaturalTransformation F F′) {G G′ : Functor E D₂} (γ : NaturalTransformation G G′) → NaturalTransformation (overlap-× {D₁ = D₁} {D₂ = D₂} H F G) (overlap-× {D₁ = D₁} {D₂ = D₂} H F′ G′) 20 | overlapN-× H φ γ = H ∘ˡ (φ ※ⁿ γ) -------------------------------------------------------------------------------- /Categories/Categories.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Categories where 3 | 4 | open import Level 5 | 6 | open import Categories.Category 7 | open import Categories.Functor 8 | 9 | Categories : ∀ o ℓ e → Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) 10 | Categories o ℓ e = record 11 | { Obj = Category o ℓ e 12 | ; _⇒_ = Functor 13 | ; _≡_ = _≡_ 14 | ; _∘_ = _∘_ 15 | ; id = id 16 | ; assoc = λ {_} {_} {_} {_} {F} {G} {H} → assoc {F = F} {G} {H} 17 | ; identityˡ = λ {_} {_} {F} → identityˡ {F = F} 18 | ; identityʳ = λ {_} {_} {F} → identityʳ {F = F} 19 | ; equiv = λ {X} {Y} → equiv {C = X} {D = Y} 20 | ; ∘-resp-≡ = λ {_} {_} {_} {f} {h} {g} {i} → ∘-resp-≡ {F = f} {h} {g} {i} 21 | } 22 | 23 | {- 24 | module Agda where 25 | open import Category.Agda 26 | open import Category.Discrete 27 | 28 | D : ∀ o → Functor (Agda o) (Categories o o zero) 29 | D o = record 30 | { F₀ = Discrete 31 | ; F₁ = λ f → record 32 | { F₀ = λ x → f x 33 | ; F₁ = ≣-cong f 34 | ; identity = tt 35 | ; homomorphism = tt 36 | ; F-resp-≡ = λ _ → tt 37 | } 38 | ; identity = λ _ → refl tt 39 | ; homomorphism = λ _ → refl tt 40 | ; F-resp-≡ = F-resp-≡′ 41 | } 42 | where 43 | F-resp-≡′ : {A B : Set o} {F G : A → B} → (∀ x → F x ≣ G x) → ∀ {X Y : A} (f : X ≣ Y) → [ Discrete B ] ≣-cong F f ∼ ≣-cong G f 44 | F-resp-≡′ {A} {B} {F} {G} f {x} ≣-refl = helper {F = F} {G} (f x) 45 | where 46 | helper : {F G : A → B} → {p q : B} → p ≣ q → [ Discrete B ] ≣-refl {x = p} ∼ ≣-refl {x = q} 47 | helper ≣-refl = refl tt 48 | 49 | Ob : ∀ o → Functor (Categories o o zero) (Agda o) 50 | Ob o = record 51 | { F₀ = Category.Obj 52 | ; F₁ = Functor.F₀ 53 | ; identity = λ _ → ≣-refl 54 | ; homomorphism = λ _ → ≣-refl 55 | ; F-resp-≡ = λ {_} {_} {F} {G} → F-resp-≡′ {F = F} {G} 56 | } 57 | where 58 | F-resp-≡′ : {A B : Category o o zero} {F G : Functor A B} 59 | → ({X Y : Category.Obj A} (f : Category.Hom A X Y) → [ B ] Functor.F₁ F f ∼ Functor.F₁ G f) 60 | → ((x : Category.Obj A) → Functor.F₀ F x ≣ Functor.F₀ G x) 61 | F-resp-≡′ {A} {B} {F} {G} F∼G x = helper (F∼G (Category.id A {x})) 62 | where 63 | helper : ∀ {X Y} {p : Category.Hom B X X} {q : Category.Hom B Y Y} → [ B ] p ∼ q → X ≣ Y 64 | helper (refl q) = ≣-refl 65 | 66 | open import Category.Adjunction 67 | 68 | D⊣Ob : ∀ o → Adjunction (D o) (Ob o) 69 | D⊣Ob o = record 70 | { unit = record 71 | { η = λ _ x → x 72 | ; commute = λ _ _ → ≣-refl 73 | } 74 | ; counit = record 75 | { η = λ X → record 76 | { F₀ = λ x → x 77 | ; F₁ = counit-id′ {X} 78 | ; identity = IsEquivalence.refl (Category.equiv X) 79 | ; homomorphism = {!!} 80 | ; F-resp-≡ = {!!} 81 | } 82 | ; commute = {!!} 83 | } 84 | ; zig = {!!} 85 | ; zag = λ _ → ≣-refl 86 | } 87 | where 88 | counit-id′ : {X : Category o o zero} → {A B : Category.Obj X} → A ≣ B → Category.Hom X A B 89 | counit-id′ {X} ≣-refl = Category.id X 90 | -} -------------------------------------------------------------------------------- /Categories/Categories/Products.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Level 4 | 5 | module Categories.Categories.Products (o ℓ e : Level) where 6 | 7 | open import Categories.Category 8 | open import Categories.Categories 9 | 10 | import Categories.Object.Products 11 | open Categories.Object.Products (Categories o ℓ e) 12 | 13 | open import Categories.Product 14 | import Categories.Product.Projections 15 | import Categories.Product.Properties 16 | 17 | open import Categories.Terminal 18 | 19 | private 20 | import Categories.Object.Product 21 | open Categories.Object.Product (Categories o ℓ e) 22 | renaming (Product to ProductObject) 23 | 24 | product : 25 | {A B : Category o ℓ e} 26 | → ProductObject A B 27 | 28 | product {A}{B} = record 29 | { A×B = Product A B 30 | ; π₁ = ∏₁ 31 | ; π₂ = ∏₂ 32 | ; ⟨_,_⟩ = _※_ 33 | ; commute₁ = λ {C}{f}{g} → ※-commute₁ {A = C}{f}{g} 34 | ; commute₂ = λ {C}{f}{g} → ※-commute₂ {A = C}{f}{g} 35 | ; universal = λ {C}{f}{g}{i} → ※-universal {A = C}{f}{g}{i} 36 | } where 37 | open Categories.Product.Projections A B 38 | open Categories.Product.Properties A B 39 | 40 | products : Products 41 | products = record 42 | { terminal = One {o}{ℓ}{e} 43 | ; binary = record 44 | { product = λ {A}{B} → product {A}{B} 45 | } 46 | } 47 | -------------------------------------------------------------------------------- /Categories/Closed.agda: -------------------------------------------------------------------------------- 1 | module Categories.Closed where -------------------------------------------------------------------------------- /Categories/Cocone.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Cocone where 3 | 4 | open import Level 5 | 6 | open import Categories.Category 7 | open import Categories.Functor hiding (_≡_; _∘_) 8 | 9 | record Cocone {o ℓ e} {o′ ℓ′ e′} {C : Category o ℓ e} {J : Category o′ ℓ′ e′} (F : Functor J C) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′) where 10 | module J = Category J 11 | open Category C 12 | open Functor F 13 | field 14 | N : Obj 15 | ψ : ∀ X → ((F₀ X) ⇒ N) 16 | .commute : ∀ {X Y} (f : J [ X , Y ]) → ψ X ≡ ψ Y ∘ F₁ f 17 | -------------------------------------------------------------------------------- /Categories/Cocones.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Cocones where 3 | 4 | open import Level 5 | 6 | open import Categories.Category 7 | open import Categories.Functor hiding (_∘_; _≡_; equiv; id; assoc; identityˡ; identityʳ; ∘-resp-≡) 8 | open import Categories.Cocone 9 | 10 | record CoconeMorphism {o ℓ e} {o′ ℓ′ e′} {C : Category o ℓ e} {J : Category o′ ℓ′ e′} {F : Functor J C} (c₁ c₂ : Cocone F) : Set (ℓ ⊔ e ⊔ o′ ⊔ ℓ′) where 11 | module c₁ = Cocone c₁ 12 | module c₂ = Cocone c₂ 13 | module C = Category C 14 | module J = Category J 15 | open C 16 | field 17 | f : C [ c₁.N , c₂.N ] 18 | .commute : ∀ {X} → f ∘ c₁.ψ X ≡ c₂.ψ X 19 | 20 | Cocones : ∀ {o ℓ e} {o′ ℓ′ e′} {C : Category o ℓ e} {J : Category o′ ℓ′ e′} (F : Functor J C) → Category (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′) (ℓ ⊔ e ⊔ o′ ⊔ ℓ′) e 21 | Cocones {C = C} F = record 22 | { Obj = Obj′ 23 | ; _⇒_ = Hom′ 24 | ; _≡_ = _≡′_ 25 | ; _∘_ = _∘′_ 26 | ; id = record { f = id; commute = identityˡ } 27 | ; assoc = assoc 28 | ; identityˡ = identityˡ 29 | ; identityʳ = identityʳ 30 | ; equiv = record 31 | { refl = Equiv.refl 32 | ; sym = Equiv.sym 33 | ; trans = Equiv.trans 34 | } 35 | ; ∘-resp-≡ = ∘-resp-≡ 36 | } 37 | where 38 | open Category C 39 | open Cocone 40 | open CoconeMorphism 41 | open Functor F 42 | 43 | infixr 9 _∘′_ 44 | infix 4 _≡′_ 45 | 46 | Obj′ = Cocone F 47 | 48 | Hom′ : Obj′ → Obj′ → Set _ 49 | Hom′ = CoconeMorphism 50 | 51 | _≡′_ : ∀ {A B} → Hom′ A B → Hom′ A B → Set _ 52 | F ≡′ G = f F ≡ f G 53 | 54 | _∘′_ : ∀ {A B C} → Hom′ B C → Hom′ A B → Hom′ A C 55 | _∘′_ {A} {B} {C} F G = record 56 | { f = f F ∘ f G 57 | ; commute = commute′ 58 | } 59 | where 60 | .commute′ : ∀ {X} → (f F ∘ f G) ∘ ψ A X ≡ ψ C X 61 | commute′ {X} = 62 | begin 63 | (f F ∘ f G) ∘ ψ A X 64 | ↓⟨ assoc ⟩ 65 | f F ∘ (f G ∘ ψ A X) 66 | ↓⟨ ∘-resp-≡ʳ (CoconeMorphism.commute G) ⟩ 67 | f F ∘ ψ B X 68 | ↓⟨ CoconeMorphism.commute F ⟩ 69 | ψ C X 70 | ∎ 71 | where 72 | open HomReasoning 73 | 74 | -------------------------------------------------------------------------------- /Categories/Coend.agda: -------------------------------------------------------------------------------- 1 | module Categories.Coend where 2 | -------------------------------------------------------------------------------- /Categories/Coequalizer.agda: -------------------------------------------------------------------------------- 1 | module Categories.Coequalizer where 2 | -------------------------------------------------------------------------------- /Categories/Colimit.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Colimit where 3 | 4 | open import Level 5 | 6 | open import Categories.Category 7 | open import Categories.Functor 8 | open import Categories.NaturalTransformation 9 | open import Categories.Functor.Constant 10 | open import Categories.Cocones 11 | open import Categories.Cocone 12 | open import Categories.Object.Initial 13 | 14 | -- Isomorphic to an terminal object, but worth keeping distinct in case we change its definition 15 | record Colimit {o ℓ e} {o′ ℓ′ e′} {C : Category o ℓ e} {J : Category o′ ℓ′ e′} (F : Functor J C) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) where 16 | field 17 | initial : Initial (Cocones F) 18 | 19 | module I = Initial initial 20 | module Ic = Cocone I.⊥ 21 | private module C = Category C 22 | 23 | ∃F : C.Obj 24 | ∃F = Ic.N 25 | 26 | ι : NaturalTransformation F (Constant ∃F) 27 | ι = record { η = Ic.ψ; commute = λ f → C.Equiv.trans (C.Equiv.sym (Ic.commute f)) (C.Equiv.sym C.identityˡ) } 28 | 29 | <_> : ∀ {Q} → NaturalTransformation F (Constant Q) → ∃F C.⇒ Q 30 | <_> {Q} η = CoconeMorphism.f (I.! {record 31 | { N = Q 32 | ; ψ = η.η 33 | ; commute = λ f → C.Equiv.trans (C.Equiv.sym C.identityˡ) (C.Equiv.sym (η.commute f)) }}) 34 | where 35 | module η = NaturalTransformation η 36 | 37 | record Cocomplete (o ℓ e : Level) {o′ ℓ′ e′} (C : Category o′ ℓ′ e′) : Set (suc o ⊔ suc ℓ ⊔ suc e ⊔ o′ ⊔ ℓ′ ⊔ e′) where 38 | field 39 | colimit : ∀ {J : Category o ℓ e} (F : Functor J C) → Colimit F 40 | -------------------------------------------------------------------------------- /Categories/Comma/Functors.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Comma.Functors where 3 | 4 | open import Categories.Category 5 | open import Categories.Functor renaming (_∘_ to _∘F_) 6 | open import Categories.FunctorCategory renaming (Functors to [_⇒_]) 7 | open import Categories.NaturalTransformation using (module NaturalTransformation) 8 | open import Data.Product using (∃; _,_; proj₁; proj₂; zip; map) 9 | open import Level 10 | open import Relation.Binary using (Rel) 11 | open import Categories.Support.EqReasoning 12 | open import Categories.Square 13 | 14 | open import Categories.Comma 15 | open import Categories.Product renaming (Product to _×_) 16 | 17 | πComma : 18 | ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂ o₃ ℓ₃ e₃} 19 | → {A : Category o₁ ℓ₁ e₁} 20 | → {B : Category o₂ ℓ₂ e₂} 21 | → {C : Category o₃ ℓ₃ e₃} 22 | → (S : Functor B C) (T : Functor A C) 23 | → Functor (S ↓ T) (B × A) 24 | πComma {A = A} {B} {C} S T = 25 | record 26 | { F₀ = F₀′ 27 | ; F₁ = F₁′ 28 | ; identity = B×A.Equiv.refl 29 | ; homomorphism = B×A.Equiv.refl 30 | ; F-resp-≡ = λ pf → pf 31 | } 32 | where 33 | module S↓T = Category (S ↓ T) 34 | module B×A = Category (B × A) 35 | open Comma S T using (_,_[_]) renaming (_,_,_ to ⟨_,_,_⟩) 36 | 37 | F₀′ : S↓T.Obj → B×A.Obj 38 | F₀′ ⟨ α , β , f ⟩ = α , β 39 | 40 | F₁′ : ∀ {X Y} → (S ↓ T) [ X , Y ] → (B × A) [ F₀′ X , F₀′ Y ] 41 | F₁′ (g , h [ commutes ]) = g , h 42 | 43 | InducedFunctor : ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂ o₃ ℓ₃ e₃} 44 | {A : Category o₁ ℓ₁ e₁} 45 | {B : Category o₂ ℓ₂ e₂} 46 | {C : Category o₃ ℓ₃ e₃} 47 | {s₁ s₂ d₁ d₂} 48 | (m : (Category.op [ A ⇒ C ] × [ B ⇒ C ]) [ (s₁ , s₂) , (d₁ , d₂) ]) 49 | → Functor (s₁ ↓ s₂) (d₁ ↓ d₂) 50 | InducedFunctor {A = A} {B} {C} {s₁} {s₂} {d₁} {d₂} m = record 51 | { F₀ = F₀′ 52 | ; F₁ = F₁′ 53 | ; identity = λ {X} → d₁↓d₂.Equiv.refl {_} {_} {d₁↓d₂.id {F₀′ X}} 54 | ; homomorphism = λ {_ _ _ f g} → d₁↓d₂.Equiv.refl {_} {_} {F₁′ ((s₁ ↓ s₂) [ g ∘ f ])} 55 | ; F-resp-≡ = λ pf → pf 56 | } 57 | where 58 | module s₁↓s₂ = Category (s₁ ↓ s₂) 59 | module d₁↓d₂ = Category (d₁ ↓ d₂) 60 | module s₁ = Functor s₁ 61 | module s₂ = Functor s₂ 62 | module d₁ = Functor d₁ 63 | module d₂ = Functor d₂ 64 | module m₁ = NaturalTransformation (proj₁ m) 65 | module m₂ = NaturalTransformation (proj₂ m) 66 | 67 | open Category C 68 | open Comma renaming (_,_,_ to ⟨_,_,_⟩) 69 | 70 | F₀′ : s₁↓s₂.Obj → d₁↓d₂.Obj 71 | F₀′ ⟨ α , β , f ⟩ = ⟨ α , β , m₂.η β ∘ f ∘ m₁.η α ⟩ 72 | 73 | F₁′ : ∀ {X Y} → (s₁ ↓ s₂) [ X , Y ] → (d₁ ↓ d₂) [ F₀′ X , F₀′ Y ] 74 | F₁′ {⟨ α₁ , β₁ , f₁ ⟩} {⟨ α₂ , β₂ , f₂ ⟩} (g , h [ commutes ]) = g , h [ ( 75 | begin 76 | d₂.F₁ h ∘ m₂.η β₁ ∘ f₁ ∘ m₁.η α₁ 77 | ↑⟨ pushˡ (m₂.commute h) ⟩ 78 | (m₂.η β₂ ∘ s₂.F₁ h) ∘ f₁ ∘ m₁.η α₁ 79 | ↓⟨ pullˡ (pullʳ commutes) ⟩ 80 | (m₂.η β₂ ∘ f₂ ∘ s₁.F₁ g) ∘ m₁.η α₁ 81 | ↑⟨ extendˡ (extendˡ (m₁.commute g)) ⟩ 82 | (m₂.η β₂ ∘ f₂ ∘ m₁.η α₂) ∘ d₁.F₁ g 83 | ∎ ) ] 84 | where 85 | open HomReasoning 86 | open GlueSquares C 87 | -------------------------------------------------------------------------------- /Categories/Comonad.agda: -------------------------------------------------------------------------------- 1 | module Categories.Comonad where 2 | 3 | open import Level using (_⊔_) 4 | 5 | open import Categories.Category using (Category) 6 | open import Categories.Functor hiding (_≡_; assoc; identityˡ; identityʳ) 7 | open import Categories.NaturalTransformation renaming (id to idN) 8 | 9 | record Comonad {o ℓ e} (C : Category o ℓ e) : Set (o ⊔ ℓ ⊔ e) where 10 | field 11 | F : Endofunctor C 12 | ε : NaturalTransformation F id 13 | δ : NaturalTransformation F (F ∘ F) 14 | 15 | open Functor F 16 | 17 | field 18 | .assoc : (δ ∘ʳ F) ∘₁ δ ≡ (F ∘ˡ δ) ∘₁ δ 19 | .identityˡ : (F ∘ˡ ε) ∘₁ δ ≡ idN 20 | .identityʳ : (ε ∘ʳ F) ∘₁ δ ≡ idN 21 | -------------------------------------------------------------------------------- /Categories/Comonad/Cofree.agda: -------------------------------------------------------------------------------- 1 | module Categories.Comonad.Cofree where 2 | -------------------------------------------------------------------------------- /Categories/Cone.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Cone where 3 | 4 | open import Level 5 | open import Relation.Binary using (IsEquivalence; Setoid; module Setoid) 6 | 7 | open import Categories.Support.PropositionalEquality 8 | 9 | open import Categories.Category 10 | open import Categories.Functor hiding (_≡_; _∘_) 11 | 12 | module ConeOver {o ℓ e} {o′ ℓ′ e′} {C : Category o ℓ e} {J : Category o′ ℓ′ e′} (F : Functor J C) where 13 | module J = Category J 14 | open Category C 15 | open Functor F 16 | record Cone : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′) where 17 | field 18 | N : Obj 19 | ψ : ∀ X → (N ⇒ F₀ X) 20 | .commute : ∀ {X Y} (f : J [ X , Y ]) → F₁ f ∘ ψ X ≡ ψ Y 21 | 22 | Cone′ = Cone 23 | 24 | record _≜_ (X Y : Cone) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′) where 25 | module X = Cone X 26 | module Y = Cone Y 27 | field 28 | N-≣ : X.N ≣ Y.N 29 | ψ-≡ : ∀ j → C [ X.ψ j ∼ Y.ψ j ] 30 | 31 | ≜-is-equivalence : IsEquivalence _≜_ 32 | ≜-is-equivalence = record 33 | { refl = record { N-≣ = ≣-refl; ψ-≡ = λ j → refl } 34 | ; sym = λ X≜Y → let module X≜Y = _≜_ X≜Y in record 35 | { N-≣ = ≣-sym X≜Y.N-≣ 36 | ; ψ-≡ = λ j → sym (X≜Y.ψ-≡ j) 37 | } 38 | ; trans = λ X≜Y Y≜Z → 39 | let module X≜Y = _≜_ X≜Y in 40 | let module Y≜Z = _≜_ Y≜Z in 41 | record 42 | { N-≣ = ≣-trans X≜Y.N-≣ Y≜Z.N-≣ 43 | ; ψ-≡ = λ j → trans (X≜Y.ψ-≡ j) (Y≜Z.ψ-≡ j) 44 | } 45 | } 46 | where open Heterogeneous C 47 | 48 | cone-setoid : Setoid _ _ 49 | cone-setoid = record 50 | { Carrier = Cone 51 | ; _≈_ = _≜_ 52 | ; isEquivalence = ≜-is-equivalence 53 | } 54 | 55 | record ConeUnder (N : Obj) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′) where 56 | field 57 | ψ′ : ∀ X → (N ⇒ F₀ X) 58 | .commute : ∀ {X Y} (f : J [ X , Y ]) → F₁ f ∘ ψ′ X ≡ ψ′ Y 59 | 60 | ConeUnder′ = ConeUnder 61 | 62 | untether : ∀ {N} → ConeUnder N → Cone 63 | untether {N} K = record { N = N; ψ = ψ′; commute = commute } 64 | where open ConeUnder K 65 | 66 | -- this Henry Ford bit is a little weird but it turns out to be handy 67 | tether : ∀ {N} → (K : Cone) → (N ≣ Cone.N K) → ConeUnder N 68 | tether K ≣-refl = record { ψ′ = Cone.ψ K; commute = Cone.commute K } 69 | 70 | record _≜′_ {N} (X Y : ConeUnder N) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′) where 71 | module X = ConeUnder X 72 | module Y = ConeUnder Y 73 | field 74 | .ψ′-≡ : ∀ j → C [ X.ψ′ j ≡ Y.ψ′ j ] 75 | 76 | ≜′-is-equivalence : (N : Obj) → IsEquivalence (_≜′_ {N}) 77 | ≜′-is-equivalence N = record 78 | { refl = record { ψ′-≡ = λ j → Equiv.refl } 79 | ; sym = λ X≜′Y → let module X≜′Y = _≜′_ X≜′Y in record 80 | { ψ′-≡ = λ j → Equiv.sym (X≜′Y.ψ′-≡ j) 81 | } 82 | ; trans = λ X≜′Y Y≜′Z → 83 | let module X≜′Y = _≜′_ X≜′Y in 84 | let module Y≜′Z = _≜′_ Y≜′Z in 85 | record { ψ′-≡ = λ j → Equiv.trans (X≜′Y.ψ′-≡ j) (Y≜′Z.ψ′-≡ j) } 86 | } 87 | where open Heterogeneous C 88 | 89 | cone-under-setoid : (N : Obj) → Setoid _ _ 90 | cone-under-setoid N = record 91 | { Carrier = ConeUnder N 92 | ; _≈_ = _≜′_ 93 | ; isEquivalence = ≜′-is-equivalence N 94 | } 95 | 96 | private 97 | .lemma₁ : ∀ {N j} {f : N ⇒ F₀ j} K → C [ f ∼ Cone.ψ K j ] → (pf : N ≣ Cone.N K) → f ≡ ConeUnder.ψ′ (tether K pf) j 98 | lemma₁ K eq ≣-refl = ∼⇒≡ eq 99 | where open Heterogeneous C 100 | 101 | -- allow replacing the N-≣ because sometimes refl is available at the place 102 | -- of use. in fact that's the main reason to even have ConeUnders. 103 | -- coincidentally, it makes this easier to write! 104 | homogenize : ∀ {K₁ K₂} → K₁ ≜ K₂ → (pf : Cone.N K₁ ≣ Cone.N K₂) → tether K₁ ≣-refl ≜′ tether K₂ pf 105 | homogenize {K₂ = K₂} K₁≜K₂ pf = record { ψ′-≡ = λ j → lemma₁ K₂ (ψ-≡ j) pf } 106 | where 107 | open _≜_ K₁≜K₂ 108 | open Heterogeneous C 109 | 110 | heterogenize : ∀ {N} {K₁ K₂ : ConeUnder N} → K₁ ≜′ K₂ → untether K₁ ≜ untether K₂ 111 | heterogenize {N} {K₁} {K₂} pf = record 112 | { N-≣ = ≣-refl 113 | ; ψ-≡ = λ j → Heterogeneous.≡⇒∼ (ψ′-≡ j) 114 | } 115 | where open _≜′_ pf 116 | 117 | open Setoid cone-setoid public renaming (refl to ≜-refl; sym to ≜-sym; trans to ≜-trans; reflexive to ≜-reflexive) 118 | module ≜′ {N : Obj} = Setoid (cone-under-setoid N) 119 | 120 | open ConeOver public using (cone-setoid) renaming (_≜_ to _[_≜_]; Cone′ to Cone; _≜′_ to _[_≜′_]; ConeUnder′ to ConeUnder) 121 | 122 | module Cone {o ℓ e} {o′ ℓ′ e′} {C : Category o ℓ e} {J : Category o′ ℓ′ e′} {F : Functor J C} (K : ConeOver.Cone F) = ConeOver.Cone K 123 | -------------------------------------------------------------------------------- /Categories/Discrete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Discrete where 3 | 4 | open import Level 5 | open import Data.Unit using (⊤) 6 | open import Function using (flip) 7 | 8 | open import Categories.Category 9 | open import Categories.Support.PropositionalEquality 10 | 11 | Discrete : ∀ {o} (A : Set o) → Category o o zero 12 | Discrete A = record 13 | { Obj = A 14 | ; _⇒_ = _≣_ 15 | ; _≡_ = λ _ _ → ⊤ 16 | ; _∘_ = flip ≣-trans 17 | ; id = ≣-refl 18 | ; assoc = _ 19 | ; identityˡ = _ 20 | ; identityʳ = _ 21 | ; equiv = _ 22 | ; ∘-resp-≡ = _ 23 | } -------------------------------------------------------------------------------- /Categories/Enriched.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Enriched where 3 | 4 | open import Categories.Category 5 | open import Categories.Monoidal 6 | 7 | -- moar -------------------------------------------------------------------------------- /Categories/Equalizer.agda: -------------------------------------------------------------------------------- 1 | module Categories.Equalizer where 2 | -------------------------------------------------------------------------------- /Categories/Equivalence/Strong.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Equivalence.Strong where 3 | 4 | -- Strong equivalence of categories. Same as ordinary equivalence in Cat. 5 | -- May not include everything we'd like to think of as equivalences, namely 6 | -- the full, faithful functors that are essentially surjective on objects. 7 | 8 | open import Level 9 | open import Relation.Binary using (IsEquivalence; module IsEquivalence) 10 | open import Function using () renaming (_∘_ to _∙_) 11 | 12 | open import Categories.Category 13 | open import Categories.Functor hiding (equiv) 14 | open import Categories.NaturalIsomorphism as NI hiding (equiv) 15 | open import Categories.NaturalTransformation as NT hiding (id; equiv) 16 | open import Categories.Morphisms using (Iso; module Iso) 17 | 18 | record WeakInverse {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} (F : Functor C D) (G : Functor D C) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) where 19 | field 20 | F∘G≅id : NaturalIsomorphism (F ∘ G) id 21 | G∘F≅id : NaturalIsomorphism (G ∘ F) id 22 | F∘G⇒id = NaturalIsomorphism.F⇒G F∘G≅id 23 | id⇒F∘G = NaturalIsomorphism.F⇐G F∘G≅id 24 | G∘F⇒id = NaturalIsomorphism.F⇒G G∘F≅id 25 | id⇒G∘F = NaturalIsomorphism.F⇐G G∘F≅id 26 | .F∘G-iso : _ 27 | F∘G-iso = NaturalIsomorphism.iso F∘G≅id 28 | .F∘G-isoˡ : _ 29 | F∘G-isoˡ = λ x → Iso.isoˡ {C = D} (F∘G-iso x) 30 | .F∘G-isoʳ : _ 31 | F∘G-isoʳ = λ x → Iso.isoʳ {C = D} (F∘G-iso x) 32 | .G∘F-iso : _ 33 | G∘F-iso = NaturalIsomorphism.iso G∘F≅id 34 | .G∘F-isoˡ : _ 35 | G∘F-isoˡ = λ x → Iso.isoˡ {C = C} (G∘F-iso x) 36 | .G∘F-isoʳ : _ 37 | G∘F-isoʳ = λ x → Iso.isoʳ {C = C} (G∘F-iso x) 38 | 39 | record StrongEquivalence {o ℓ e o′ ℓ′ e′} (C : Category o ℓ e) (D : Category o′ ℓ′ e′) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) where 40 | field 41 | F : Functor C D 42 | G : Functor D C 43 | weak-inverse : WeakInverse F G 44 | open WeakInverse weak-inverse public 45 | 46 | module Equiv where 47 | refl : ∀ {o ℓ e} {C : Category o ℓ e} → StrongEquivalence C C 48 | refl = record 49 | { F = id 50 | ; G = id 51 | ; weak-inverse = record 52 | { F∘G≅id = IsEquivalence.refl NI.equiv 53 | ; G∘F≅id = IsEquivalence.refl NI.equiv 54 | } 55 | } 56 | 57 | sym : ∀ {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} → StrongEquivalence C D → StrongEquivalence D C 58 | sym Inv = record 59 | { F = Inv.G 60 | ; G = Inv.F 61 | ; weak-inverse = record 62 | { F∘G≅id = Inv.G∘F≅id 63 | ; G∘F≅id = Inv.F∘G≅id 64 | } 65 | } 66 | where 67 | module Inv = StrongEquivalence Inv 68 | 69 | trans : ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂ o₃ ℓ₃ e₃} {C₁ : Category o₁ ℓ₁ e₁} {C₂ : Category o₂ ℓ₂ e₂} {C₃ : Category o₃ ℓ₃ e₃} → StrongEquivalence C₁ C₂ → StrongEquivalence C₂ C₃ → StrongEquivalence C₁ C₃ 70 | trans {C₁ = C₁} {C₂} {C₃} A B = record 71 | { F = B.F ∘ A.F 72 | ; G = A.G ∘ B.G 73 | ; weak-inverse = record 74 | { F∘G≅id = IsEquivalence.trans NI.equiv ((B.F ⓘˡ A.F∘G≅id) ⓘʳ B.G) B.F∘G≅id 75 | ; G∘F≅id = IsEquivalence.trans NI.equiv ((A.G ⓘˡ B.G∘F≅id) ⓘʳ A.F) A.G∘F≅id 76 | } 77 | } 78 | where 79 | module A = StrongEquivalence A 80 | module B = StrongEquivalence B 81 | 82 | equiv : ∀ {o ℓ e} → IsEquivalence (StrongEquivalence {o} {ℓ} {e}) 83 | equiv = record { refl = Equiv.refl; sym = Equiv.sym; trans = Equiv.trans } 84 | -------------------------------------------------------------------------------- /Categories/Fam.agda: -------------------------------------------------------------------------------- 1 | module Categories.Fam where 2 | 3 | open import Level 4 | open import Relation.Binary using (Rel) 5 | import Relation.Binary.HeterogeneousEquality as Het 6 | open Het using (_≅_) renaming (refl to ≣-refl) 7 | 8 | open import Categories.Support.PropositionalEquality 9 | 10 | open import Categories.Category 11 | 12 | module Fam {a b : Level} where 13 | record Fam : Set (suc a ⊔ suc b) where 14 | constructor _,_ 15 | field 16 | U : Set a 17 | T : U → Set b 18 | open Fam public 19 | 20 | record Hom (A B : Fam) : Set (a ⊔ b) where 21 | constructor _,_ 22 | field 23 | f : U A → U B 24 | φ : (x : U A) → T A x → T B (f x) 25 | 26 | record _≡Fam_ {X Y} (f g : (Hom X Y)) : Set (a ⊔ b) where 27 | constructor _,_ 28 | field 29 | f≡g : {x : _} → Hom.f f x ≣ Hom.f g x 30 | φ≡γ : {x : _} {bx : _} → Hom.φ f x bx ≅ Hom.φ g x bx 31 | 32 | module Eq = _≡Fam_ 33 | 34 | Cat : Category (suc a ⊔ suc b) (a ⊔ b) (a ⊔ b) 35 | Cat = record { 36 | Obj = Fam; 37 | _⇒_ = Hom; 38 | _≡_ = _≡Fam_; 39 | id = id′; 40 | _∘_ = _∘′_; 41 | assoc = ≣-refl , ≣-refl; 42 | identityˡ = ≣-refl , ≣-refl; 43 | identityʳ = ≣-refl , ≣-refl; 44 | equiv = record { 45 | refl = ≣-refl , ≣-refl; 46 | sym = \ { (f≡g , φ≡γ) → ≣-sym f≡g , Het.sym φ≡γ }; 47 | trans = λ {(f≡g , φ≡γ) (g≡h , γ≡η) → ≣-trans f≡g g≡h , Het.trans φ≡γ γ≡η} }; 48 | ∘-resp-≡ = ∘-resp-≡′ } 49 | where 50 | id′ : {A : Fam} → Hom A A 51 | id′ = (\ x → x) , (\ x bx → bx) 52 | 53 | _∘′_ : {A B C : Fam} → Hom B C → Hom A B → Hom A C 54 | _∘′_ (f , φ) (g , γ) = (λ x → f (g x)) , (λ x bx → φ (g x) (γ x bx)) 55 | 56 | sym′ : ∀ {X Y} → Relation.Binary.Symmetric (_≡Fam_ {X} {Y}) 57 | sym′ {Ax , Bx} {Ay , By} {f , φ} {g , γ} (f≡g , φ≡γ) = ≣-sym f≡g , Het.sym φ≡γ 58 | 59 | ∘-resp-≡′ : {A B C : Fam} {f h : Hom B C} {g i : Hom A B} → f ≡Fam h → g ≡Fam i → (f ∘′ g) ≡Fam (h ∘′ i) 60 | ∘-resp-≡′ {f = (f , φ)} {g , γ} {h , η} {i , ι} (f≡g , φ≡γ) (h≡i , η≡ι) = 61 | ≣-trans f≡g (≣-cong g h≡i) , Het.trans φ≡γ (Het.cong₂ γ (Het.≡-to-≅ h≡i) η≡ι) 62 | 63 | open Category Cat public 64 | 65 | 66 | Fam : ∀ a b → Category (suc a ⊔ suc b) (a ⊔ b) (a ⊔ b) 67 | Fam a b = Fam.Cat {a} {b} 68 | -------------------------------------------------------------------------------- /Categories/Fibration.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Fibration where 3 | 4 | open import Level hiding (lift) 5 | open import Data.Product 6 | 7 | open import Categories.Category 8 | open import Categories.Functor hiding (_∘_; _≡_) 9 | open import Categories.Morphism.Cartesian 10 | import Categories.Morphisms as Morphisms 11 | 12 | record CartesianLifting {o₀ ℓ₀ e₀} {o₁ ℓ₁ e₁} {E : Category o₀ ℓ₀ e₀} {B : Category o₁ ℓ₁ e₁} 13 | (p : Functor E B) {a e} (f : B [ a , Functor.F₀ p e ]) : Set (o₀ ⊔ ℓ₀ ⊔ e₀ ⊔ o₁ ⊔ ℓ₁ ⊔ e₁) where 14 | private module E = Category E 15 | private module B = Category B 16 | open B using (_∘_; _≡_) 17 | 18 | open Functor p renaming (F₀ to p₀; F₁ to p₁) 19 | open Morphisms B 20 | 21 | field 22 | e′ : E.Obj 23 | 24 | φ : E [ e′ , e ] 25 | proof : ∃ \ (h : a ≅ p₀ e′) → f ∘ _≅_.g h ≡ p₁ φ 26 | φ-cartesian : Cartesian p φ 27 | 28 | 29 | record Fibration {o₀ ℓ₀ e₀} {o₁ ℓ₁ e₁} (E : Category o₀ ℓ₀ e₀) (B : Category o₁ ℓ₁ e₁) : Set (o₀ ⊔ ℓ₀ ⊔ e₀ ⊔ o₁ ⊔ ℓ₁ ⊔ e₁) where 30 | private module E = Category E 31 | private module B = Category B 32 | open B using (_∘_; _≡_) 33 | 34 | field 35 | p : Functor E B 36 | 37 | open Functor p renaming (F₀ to p₀; F₁ to p₁) 38 | 39 | open Morphisms B 40 | 41 | field 42 | lift : ∀ {a} e → (f : B [ a , p₀ e ]) → CartesianLifting p f 43 | -------------------------------------------------------------------------------- /Categories/Free.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | module Categories.Free where 4 | 5 | open import Categories.Category 6 | open import Categories.Free.Core 7 | open import Categories.Free.Functor 8 | open import Categories.Graphs.Underlying 9 | open import Categories.Functor 10 | using (Functor) 11 | open import Graphs.Graph 12 | open import Graphs.GraphMorphism 13 | using (GraphMorphism; module GraphMorphism) 14 | open import Data.Star 15 | 16 | -- Exports from other modules: 17 | -- Free₀, Free₁ and Free 18 | open Categories.Free.Core public 19 | using (Free₀) 20 | open Categories.Free.Functor public 21 | using (Free) 22 | 23 | -- TODO: 24 | -- Prove Free⊣Underlying : Adjunction Free Underlying 25 | -- Define Adjunction.left and Adjunction.right as conveniences 26 | -- (or whatever other names make sense for the hom-set maps 27 | -- C [ F _ , _ ] → D [ _ , G _ ] and inverse, respectively) 28 | -- Let Cata = Adjunction.left Free⊣Underlying 29 | Cata : ∀{o₁ ℓ₁ e₁}{G : Graph o₁ ℓ₁ e₁} 30 | {o₂ ℓ₂ e₂}{C : Category o₂ ℓ₂ e₂} 31 | → (F : GraphMorphism G (Underlying₀ C)) 32 | → Functor (Free₀ G) C 33 | Cata {G = G} {C = C} F = record 34 | { F₀ = F₀ 35 | ; F₁ = F₁* 36 | ; identity = refl 37 | ; homomorphism = λ{X}{Y}{Z}{f}{g} → homomorphism {X}{Y}{Z}{f}{g} 38 | ; F-resp-≡ = F₁*-resp-≡ 39 | } 40 | where 41 | open Category C 42 | open GraphMorphism F 43 | open Equiv 44 | open HomReasoning 45 | open PathEquality using (ε-cong; _◅-cong_) 46 | 47 | F₁* : ∀ {A B} → Free₀ G [ A , B ] → C [ F₀ A , F₀ B ] 48 | F₁* ε = id 49 | F₁* (f ◅ fs) = F₁* fs ∘ F₁ f 50 | 51 | .homomorphism : ∀ {X Y Z} {f : Free₀ G [ X , Y ]} {g : Free₀ G [ Y , Z ]} 52 | → C [ F₁* (Free₀ G [ g ∘ f ]) ≡ C [ F₁* g ∘ F₁* f ] ] 53 | homomorphism {f = ε} = sym identityʳ 54 | homomorphism {f = f ◅ fs}{gs} = 55 | begin 56 | F₁* (fs ◅◅ gs) ∘ F₁ f 57 | ↓⟨ homomorphism {f = fs}{gs} ⟩∘⟨ refl ⟩ 58 | (F₁* gs ∘ F₁* fs) ∘ F₁ f 59 | ↓⟨ assoc ⟩ 60 | F₁* gs ∘ F₁* fs ∘ F₁ f 61 | ∎ 62 | 63 | .F₁*-resp-≡ : ∀ {A B} {f g : Free₀ G [ A , B ]} → Free₀ G [ f ≡ g ] → C [ F₁* f ≡ F₁* g ] 64 | F₁*-resp-≡ {f = ε}{.ε} ε-cong = refl 65 | F₁*-resp-≡ {f = f ◅ fs}{g ◅ gs} (f≈g ◅-cong fs≈gs) = 66 | begin 67 | F₁* fs ∘ F₁ f 68 | ↓⟨ F₁*-resp-≡ fs≈gs ⟩∘⟨ F-resp-≈ f≈g ⟩ 69 | F₁* gs ∘ F₁ g 70 | ∎ 71 | F₁*-resp-≡ {f = f ◅ fs}{ε} () 72 | -------------------------------------------------------------------------------- /Categories/Free/Core.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | -- Free category generated by a directed graph. 4 | -- The graph is given by a type (Obj) of nodes and a (Obj × Obj)-indexed 5 | -- Setoid of (directed) edges. Every inhabitant of the node type is 6 | -- considered to be a (distinct) node of the graph, and (equivalence 7 | -- classes of) inhabitants of the edge type (for a given pair of 8 | -- nodes) are considered edges of the graph. 9 | module Categories.Free.Core where 10 | 11 | open import Categories.Category 12 | open import Categories.Functor 13 | using (Functor) 14 | open import Categories.Support.Equivalence 15 | open import Graphs.Graph 16 | open import Graphs.GraphMorphism 17 | open import Relation.Binary 18 | using (module IsEquivalence) 19 | open import Relation.Binary.PropositionalEquality 20 | using () 21 | renaming (_≡_ to _≣_; refl to ≣-refl) 22 | open import Data.Star 23 | open import Categories.Support.StarEquality 24 | open import Data.Star.Properties 25 | using (gmap-◅◅) 26 | open import Level using (_⊔_) 27 | 28 | EdgeSetoid : ∀ {o ℓ e} → (G : Graph o ℓ e) (A B : Graph.Obj G) → Setoid ℓ e 29 | EdgeSetoid G A B = Graph.edge-setoid G {A}{B} 30 | 31 | module PathEquality {o ℓ e} (G : Graph o ℓ e) = StarEquality (EdgeSetoid G) 32 | 33 | Free₀ : ∀ {o ℓ e} → Graph o ℓ e → Category o (o ⊔ ℓ) (o ⊔ ℓ ⊔ e) 34 | Free₀ {o}{ℓ}{e} G = record 35 | { Obj = Obj 36 | ; _⇒_ = Star _↝_ 37 | ; id = ε 38 | ; _∘_ = _▻▻_ 39 | ; _≡_ = _≡_ 40 | ; equiv = equiv 41 | ; assoc = λ{A}{B}{C}{D} {f}{g}{h} → ▻▻-assoc {A}{B}{C}{D} f g h 42 | ; identityˡ = IsEquivalence.reflexive equiv f◅◅ε≣f 43 | ; identityʳ = IsEquivalence.reflexive equiv ≣-refl 44 | ; ∘-resp-≡ = _▻▻-cong_ 45 | } 46 | where 47 | open Graph G 48 | hiding (equiv) 49 | 50 | open PathEquality G 51 | renaming (_≈_ to _≡_; isEquivalence to equiv) 52 | 53 | f◅◅ε≣f : ∀ {A B}{f : Star _↝_ A B} -> (f ◅◅ ε) ≣ f 54 | f◅◅ε≣f {f = ε} = ≣-refl 55 | f◅◅ε≣f {f = x ◅ xs} rewrite f◅◅ε≣f {f = xs} = ≣-refl 56 | 57 | Free₁ : ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} 58 | {A : Graph o₁ ℓ₁ e₁} 59 | {B : Graph o₂ ℓ₂ e₂} 60 | → GraphMorphism A B → Functor (Free₀ A) (Free₀ B) 61 | Free₁ {o₁}{ℓ₁}{e₁}{o₂}{ℓ₂}{e₂}{A}{B} G = record 62 | { F₀ = G₀ 63 | ; F₁ = gmap G₀ G₁ 64 | ; identity = refl 65 | ; homomorphism = λ {X}{Y}{Z}{f}{g} → homomorphism {X}{Y}{Z}{f}{g} 66 | ; F-resp-≡ = F-resp-≡ 67 | } 68 | where 69 | open GraphMorphism G 70 | renaming (F₀ to G₀; F₁ to G₁; F-resp-≈ to G-resp-≈) 71 | open Category.Equiv (Free₀ B) 72 | 73 | .homomorphism : ∀ {X Y Z} 74 | {f : Free₀ A [ X , Y ]} 75 | {g : Free₀ A [ Y , Z ]} 76 | → Free₀ B [ gmap G₀ G₁ (f ◅◅ g) ≡ gmap G₀ G₁ f ◅◅ gmap G₀ G₁ g ] 77 | homomorphism {f = f}{g} = reflexive (gmap-◅◅ G₀ G₁ f g) 78 | 79 | .F-resp-≡ : ∀ {X Y} 80 | {f g : Free₀ A [ X , Y ]} 81 | → Free₀ A [ f ≡ g ] 82 | → Free₀ B [ gmap G₀ G₁ f ≡ gmap G₀ G₁ g ] 83 | F-resp-≡ {X}{Y}{f}{g} f≡g 84 | = gmap-cong G₀ G₁ G₁ (λ x y x≈y → G-resp-≈ x≈y) f g f≡g 85 | where 86 | open StarCong₂ (EdgeSetoid A) (EdgeSetoid B) 87 | -------------------------------------------------------------------------------- /Categories/Free/Functor.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | -- Proof that 'Free' is a functor 4 | module Categories.Free.Functor where 5 | 6 | open import Categories.Support.PropositionalEquality 7 | 8 | open import Categories.Categories 9 | open import Categories.Category 10 | renaming (_[_∼_] to _[_~C_]) 11 | open import Categories.Free.Core 12 | open import Categories.Functor 13 | using (Functor) 14 | renaming (_≡_ to _≡F_; _∘_ to _∘F_) 15 | open import Categories.Graphs 16 | open import Data.Product 17 | open import Graphs.Graph 18 | renaming (_[_~_] to _[_~G_]; module Heterogeneous to HeterogeneousG) 19 | open import Graphs.GraphMorphism 20 | renaming (_∘_ to _∘G_) 21 | open import Data.Star 22 | open import Data.Star.Properties 23 | using (gmap-◅◅; gmap-id) 24 | open import Categories.Support.StarEquality 25 | open import Level using (_⊔_) 26 | 27 | ε∼ε : ∀{o₁ ℓ₁ e₁}{X : Graph o₁ ℓ₁ e₁} 28 | {o₂ ℓ₂ e₂}{Y : Graph o₂ ℓ₂ e₂} 29 | → (F G : GraphMorphism X Y) 30 | → F ≈ G 31 | → {x : Graph.Obj X} 32 | → Free₀ Y [ ε {x = GraphMorphism.F₀ F x} ~C ε {x = GraphMorphism.F₀ G x} ] 33 | ε∼ε {Y = Y} F G (F≈G₀ , F≈G₁) {x} = ≣-subst (λ z → Free₀ Y [ ε {x = GraphMorphism.F₀ F x} ~C ε {x = z} ]) (F≈G₀ x) (Heterogeneous.refl (Free₀ Y)) 34 | -- the below should probably work, but there's an agda bug 35 | -- XXX bug id? mokus? anybody? bueller? 36 | {- 37 | ε∼ε {Y = Y} F G F≈G {x} rewrite proj₁ F≈G x = refl 38 | where open Heterogeneous (Free₀ Y) 39 | -} 40 | 41 | _◅~◅_ : 42 | ∀ {o ℓ e}{G : Graph o ℓ e} 43 | {a₀ a₁ b₀ b₁ c₀ c₁ : Graph.Obj G} 44 | {f : G [ a₀ ↝ b₀ ]} 45 | {g : G [ a₁ ↝ b₁ ]} 46 | {fs : Free₀ G [ b₀ , c₀ ]} 47 | {gs : Free₀ G [ b₁ , c₁ ]} 48 | → G [ f ~G g ] 49 | → Free₀ G [ fs ~C gs ] 50 | → Free₀ G [ (f ◅ fs) ~C (g ◅ gs) ] 51 | _◅~◅_ {G = G} (HeterogeneousG.≈⇒~ hd) (Heterogeneous.≡⇒∼ tl) 52 | = ≡⇒∼ (hd ◅-cong tl) 53 | where 54 | open Heterogeneous (Free₀ G) 55 | open PathEquality G 56 | 57 | Free : ∀ {o ℓ e} → Functor (Graphs o ℓ e) (Categories o (o ⊔ ℓ) (o ⊔ ℓ ⊔ e)) 58 | Free {o}{ℓ}{e} = record 59 | { F₀ = Free₀ 60 | ; F₁ = Free₁ 61 | ; identity = λ {A} f → Heterogeneous.reflexive (Free₀ A) (gmap-id f) 62 | ; homomorphism = λ {X}{Y}{Z}{f}{g} → homomorphism {X}{Y}{Z}{f}{g} 63 | ; F-resp-≡ = λ {X}{Y}{F G : GraphMorphism X Y} → Free₁-resp-≡ {X}{Y}{F}{G} 64 | } 65 | where 66 | module Graphs = Category (Graphs o ℓ e) 67 | module Categories = Category (Categories o (o ⊔ ℓ) (o ⊔ ℓ ⊔ e)) 68 | 69 | .homomorphism : ∀ {X Y Z} {f : GraphMorphism X Y} {g : GraphMorphism Y Z} 70 | → Free₁ (g ∘G f) ≡F (Free₁ g ∘F Free₁ f) 71 | homomorphism ε = Heterogeneous.refl _ 72 | homomorphism {X}{Y}{Z}{f}{g}{S}{U} (_◅_ {.S}{T}{.U} h hs) = 73 | HeterogeneousG.refl Z ◅~◅ homomorphism {X}{Y}{Z}{f}{g}{T}{U} hs 74 | 75 | .Free₁-resp-≡ : ∀ {X Y} {F G : GraphMorphism X Y} 76 | → F ≈ G 77 | → Free₁ F ≡F Free₁ G 78 | Free₁-resp-≡ {X}{Y}{F}{G} F≈G {S}{.S} ε = ε∼ε F G F≈G 79 | Free₁-resp-≡ {X}{Y}{F}{G} F≈G {S}{U} (_◅_ {.S}{T}{.U} h hs) 80 | = proj₂ F≈G h ◅~◅ Free₁-resp-≡ {X}{Y}{F}{G} F≈G {T}{U} hs 81 | -------------------------------------------------------------------------------- /Categories/Functor/Algebra.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Functor.Algebra where 3 | 4 | open import Level hiding (lift) 5 | 6 | open import Categories.Category 7 | open import Categories.Functor 8 | 9 | record F-Algebra {o ℓ e} {C : Category o ℓ e} (F : Endofunctor C) : Set (o ⊔ ℓ) where 10 | constructor _,_ 11 | open Category C 12 | open Functor F 13 | field 14 | A : Obj 15 | α : F₀ A ⇒ A 16 | 17 | lift : ∀ {o ℓ e} {C : Category o ℓ e} {F : Endofunctor C} → F-Algebra F → F-Algebra F 18 | lift {F = F} (A , α) = record { A = F₀ A; α = F₁ α } 19 | where 20 | open Functor F -------------------------------------------------------------------------------- /Categories/Functor/Algebras.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Functor.Algebras where 3 | 4 | open import Level hiding (lift) 5 | 6 | open import Categories.Category 7 | open import Categories.Functor hiding (_≡_; id; _∘_; equiv; assoc; identityˡ; identityʳ; ∘-resp-≡) 8 | open import Categories.Functor.Algebra 9 | 10 | record F-Algebra-Morphism {o ℓ e} {C : Category o ℓ e} {F : Endofunctor C} (X Y : F-Algebra F) : Set (ℓ ⊔ e) where 11 | constructor _,_ 12 | open Category C 13 | module X = F-Algebra X 14 | module Y = F-Algebra Y 15 | open Functor F 16 | field 17 | f : X.A ⇒ Y.A 18 | .commutes : f ∘ X.α ≡ Y.α ∘ F₁ f 19 | 20 | F-Algebras : ∀ {o ℓ e} {C : Category o ℓ e} → Endofunctor C → Category (ℓ ⊔ o) (e ⊔ ℓ) e 21 | F-Algebras {C = C} F = record 22 | { Obj = Obj′ 23 | ; _⇒_ = Hom′ 24 | ; _≡_ = _≡′_ 25 | ; _∘_ = _∘′_ 26 | ; id = id′ 27 | ; assoc = assoc 28 | ; identityˡ = identityˡ 29 | ; identityʳ = identityʳ 30 | ; equiv = record 31 | { refl = refl 32 | ; sym = sym 33 | ; trans = trans 34 | } 35 | ; ∘-resp-≡ = ∘-resp-≡ 36 | } 37 | where 38 | open Category C 39 | open Equiv 40 | open Functor F 41 | 42 | Obj′ = F-Algebra F 43 | 44 | Hom′ : (A B : Obj′) → Set _ 45 | Hom′ = F-Algebra-Morphism 46 | 47 | _≡′_ : ∀ {A B} (f g : Hom′ A B) → Set _ 48 | (f , _) ≡′ (g , _) = f ≡ g 49 | 50 | _∘′_ : ∀ {A B C} → Hom′ B C → Hom′ A B → Hom′ A C 51 | _∘′_ {A} {B} {C} (f , pf₁) (g , pf₂) = _,_ {X = A} {C} (f ∘ g) pf -- TODO: find out why the hell I need to provide these implicits 52 | where 53 | module A = F-Algebra A 54 | module B = F-Algebra B 55 | module C = F-Algebra C 56 | 57 | .pf : (f ∘ g) ∘ A.α ≡ C.α ∘ (F₁ (f ∘ g)) 58 | pf = begin 59 | (f ∘ g) ∘ A.α 60 | ↓⟨ assoc ⟩ 61 | f ∘ (g ∘ A.α) 62 | ↓⟨ ∘-resp-≡ʳ pf₂ ⟩ 63 | f ∘ (B.α ∘ F₁ g) 64 | ↑⟨ assoc ⟩ 65 | (f ∘ B.α) ∘ F₁ g 66 | ↓⟨ ∘-resp-≡ˡ pf₁ ⟩ 67 | (C.α ∘ F₁ f) ∘ F₁ g 68 | ↓⟨ assoc ⟩ 69 | C.α ∘ (F₁ f ∘ F₁ g) 70 | ↑⟨ ∘-resp-≡ʳ homomorphism ⟩ 71 | C.α ∘ (F₁ (f ∘ g)) 72 | ∎ 73 | where 74 | open HomReasoning 75 | 76 | id′ : ∀ {A} → Hom′ A A 77 | id′ {A} = _,_ {X = A} {A} id pf 78 | where 79 | module A = F-Algebra A 80 | 81 | .pf : id ∘ A.α ≡ A.α ∘ F₁ id 82 | pf = begin 83 | id ∘ A.α 84 | ↓⟨ identityˡ ⟩ 85 | A.α 86 | ↑⟨ identityʳ ⟩ 87 | A.α ∘ id 88 | ↑⟨ ∘-resp-≡ʳ identity ⟩ 89 | A.α ∘ F₁ id 90 | ∎ 91 | where 92 | open HomReasoning 93 | 94 | 95 | open import Categories.Object.Initial 96 | 97 | module Lambek {o ℓ e} {C : Category o ℓ e} {F : Endofunctor C} (I : Initial (F-Algebras F)) where 98 | open Category C 99 | open Equiv 100 | module FA = Category (F-Algebras F) renaming (_∘_ to _∘FA_; _≡_ to _≡FA_) 101 | open Functor F 102 | import Categories.Morphisms as Morphisms 103 | open Morphisms C 104 | open Initial I 105 | open F-Algebra ⊥ 106 | 107 | lambek : A ≅ F₀ A 108 | lambek = record 109 | { f = f′ 110 | ; g = g′ 111 | ; iso = iso′ 112 | } 113 | where 114 | f′ : A ⇒ F₀ A 115 | f′ = F-Algebra-Morphism.f (! {lift ⊥}) 116 | 117 | g′ : F₀ A ⇒ A 118 | g′ = α 119 | 120 | .iso′ : Iso f′ g′ 121 | iso′ = record 122 | { isoˡ = isoˡ′ 123 | ; isoʳ = begin 124 | f′ ∘ g′ 125 | ↓⟨ F-Algebra-Morphism.commutes (! {lift ⊥}) ⟩ 126 | F₁ g′ ∘ F₁ f′ 127 | ↑⟨ homomorphism ⟩ 128 | F₁ (g′ ∘ f′) 129 | ↓⟨ F-resp-≡ isoˡ′ ⟩ 130 | F₁ id 131 | ↓⟨ identity ⟩ 132 | id 133 | ∎ 134 | } 135 | where 136 | open FA hiding (id; module HomReasoning) 137 | open HomReasoning 138 | 139 | isoˡ′ = ⊥-id ((_,_ {C = C} {F} g′ refl) ∘FA !) 140 | 141 | open Lambek public 142 | -------------------------------------------------------------------------------- /Categories/Functor/Coalgebra.agda: -------------------------------------------------------------------------------- 1 | module Categories.Functor.Coalgebra where 2 | -------------------------------------------------------------------------------- /Categories/Functor/Coalgebras.agda: -------------------------------------------------------------------------------- 1 | module Categories.Functor.Coalgebras where 2 | -------------------------------------------------------------------------------- /Categories/Functor/Constant.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Functor.Constant where 3 | 4 | open import Categories.Category 5 | open import Categories.Functor 6 | 7 | Constant : ∀ {o′ ℓ′ e′} {D : Category o′ ℓ′ e′} (x : Category.Obj D) → ∀ {o ℓ e} {C : Category o ℓ e} → Functor C D 8 | Constant {D = D} x = record 9 | { F₀ = λ _ → x 10 | ; F₁ = λ _ → D.id 11 | ; identity = D.Equiv.refl 12 | ; homomorphism = D.Equiv.sym D.identityˡ 13 | ; F-resp-≡ = λ _ → D.Equiv.refl 14 | } 15 | where 16 | module D = Category D 17 | -------------------------------------------------------------------------------- /Categories/Functor/Core.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Functor.Core where 6 | 7 | open import Level 8 | open import Categories.Support.EqReasoning 9 | 10 | record Functor {o ℓ e o′ ℓ′ e′} (C : Category o ℓ e) (D : Category o′ ℓ′ e′) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) where 11 | eta-equality 12 | private module C = Category C 13 | private module D = Category D 14 | 15 | 16 | field 17 | F₀ : C.Obj → D.Obj 18 | {F₁} : ∀ {A B} → C [ A , B ] → D [ F₀ A , F₀ B ] 19 | 20 | .{identity} : ∀ {A} → D [ F₁ (C.id {A}) ≡ D.id ] 21 | .{homomorphism} : ∀ {X Y Z} {f : C [ X , Y ]} {g : C [ Y , Z ]} 22 | → D [ F₁ (C [ g ∘ f ]) ≡ D [ F₁ g ∘ F₁ f ] ] 23 | .{F-resp-≡} : ∀ {A B} {F G : C [ A , B ]} → C [ F ≡ G ] → D [ F₁ F ≡ F₁ G ] 24 | 25 | op : Functor C.op D.op 26 | op = record 27 | { F₀ = F₀ 28 | ; F₁ = F₁ 29 | ; identity = identity 30 | ; homomorphism = homomorphism 31 | ; F-resp-≡ = F-resp-≡ 32 | } 33 | 34 | 35 | 36 | Endofunctor : ∀ {o ℓ e} → Category o ℓ e → Set _ 37 | Endofunctor C = Functor C C 38 | 39 | Contravariant : ∀ {o ℓ e o′ ℓ′ e′} (C : Category o ℓ e) (D : Category o′ ℓ′ e′) → Set _ 40 | Contravariant C D = Functor C.op D 41 | where module C = Category C 42 | 43 | id : ∀ {o ℓ e} {C : Category o ℓ e} → Endofunctor C 44 | id {C = C} = record 45 | { F₀ = λ x → x 46 | ; F₁ = λ x → x 47 | ; identity = refl 48 | ; homomorphism = refl 49 | ; F-resp-≡ = λ x → x 50 | } 51 | where open Category.Equiv C 52 | 53 | infixr 9 _∘_ 54 | 55 | _∘_ : ∀ {o ℓ e} {o′ ℓ′ e′} {o′′ ℓ′′ e′′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} {E : Category o′′ ℓ′′ e′′} 56 | → Functor D E → Functor C D → Functor C E 57 | _∘_ {C = C} {D = D} {E = E} F G = record 58 | { F₀ = λ x → F₀ (G₀ x) 59 | ; F₁ = λ f → F₁ (G₁ f) 60 | ; identity = identity′ 61 | ; homomorphism = homomorphism′ 62 | ; F-resp-≡ = ∘-resp-≡′ 63 | } 64 | where 65 | module C = Category C 66 | module D = Category D 67 | module E = Category E 68 | module F = Functor F 69 | module G = Functor G 70 | open F 71 | open G renaming (F₀ to G₀; F₁ to G₁; F-resp-≡ to G-resp-≡) 72 | 73 | .identity′ : ∀ {A} → E [ F₁ (G₁ (C.id {A})) ≡ E.id ] 74 | identity′ = begin 75 | F₁ (G₁ C.id) 76 | ≈⟨ F-resp-≡ G.identity ⟩ 77 | F₁ D.id 78 | ≈⟨ F.identity ⟩ 79 | E.id 80 | ∎ 81 | where 82 | open SetoidReasoning E.hom-setoid 83 | 84 | .homomorphism′ : ∀ {X Y Z} {f : C [ X , Y ]} {g : C [ Y , Z ]} 85 | → E [ F₁ (G₁ (C [ g ∘ f ])) ≡ E [ F₁ (G₁ g) ∘ F₁ (G₁ f) ] ] 86 | homomorphism′ {f = f} {g = g} = begin 87 | F₁ (G₁ (C [ g ∘ f ])) 88 | ≈⟨ F-resp-≡ G.homomorphism ⟩ 89 | F₁ (D [ G₁ g ∘ G₁ f ]) 90 | ≈⟨ F.homomorphism ⟩ 91 | (E [ F₁ (G₁ g) ∘ F₁ (G₁ f) ]) 92 | ∎ 93 | where 94 | open SetoidReasoning E.hom-setoid 95 | 96 | .∘-resp-≡′ : ∀ {A B} {F G : C [ A , B ]} 97 | → C [ F ≡ G ] → E [ F₁ (G₁ F) ≡ F₁ (G₁ G) ] 98 | ∘-resp-≡′ = λ x → F-resp-≡ (G-resp-≡ x) 99 | -------------------------------------------------------------------------------- /Categories/Functor/Diagonal.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Functor.Diagonal where 3 | 4 | open import Data.Product 5 | 6 | open import Categories.Category 7 | open import Categories.Functor 8 | open import Categories.Product 9 | open import Categories.FunctorCategory 10 | open import Categories.Functor.Constant 11 | 12 | import Categories.Power as Power 13 | 14 | Δ : ∀ {o ℓ e} → (C : Category o ℓ e) → Functor C (Product C C) 15 | Δ C = record 16 | { F₀ = λ x → x , x 17 | ; F₁ = λ f → f , f 18 | ; identity = refl , refl 19 | ; homomorphism = refl , refl 20 | ; F-resp-≡ = λ x → x , x 21 | } 22 | where 23 | open Category C 24 | open Equiv 25 | 26 | Δ′ : ∀ {o ℓ e} → (I : Set) → (C : Category o ℓ e) → Functor C (Power.Exp C I) 27 | Δ′ I C = record 28 | { F₀ = λ x _ → x 29 | ; F₁ = λ f _ → f 30 | ; identity = λ _ → refl 31 | ; homomorphism = λ _ → refl 32 | ; F-resp-≡ = λ x _ → x 33 | } 34 | where 35 | open Power C 36 | open Category C 37 | open Equiv 38 | 39 | ΔF : ∀ {o ℓ e o₁ ℓ₁ e₁} {C : Category o ℓ e} (I : Category o₁ ℓ₁ e₁) → Functor C (Functors I C) 40 | ΔF {C = C} I = record 41 | { F₀ = λ c → Constant c 42 | ; F₁ = λ f → record { η = λ X → f; commute = λ g → trans C.identityʳ (sym C.identityˡ) } 43 | ; identity = refl 44 | ; homomorphism = refl 45 | ; F-resp-≡ = λ x → x 46 | } 47 | where 48 | module C = Category C 49 | open C.Equiv 50 | module I = Category I 51 | -------------------------------------------------------------------------------- /Categories/Functor/Discrete.agda: -------------------------------------------------------------------------------- 1 | module Categories.Functor.Discrete where 2 | 3 | open import Categories.Category 4 | open import Categories.Functor 5 | open import Categories.Agda 6 | open import Categories.Categories 7 | open import Categories.Support.PropositionalEquality 8 | import Categories.Discrete as D 9 | 10 | Discrete : ∀ {o} -> Functor (Sets o) (Categories o o _) 11 | Discrete {o} = record { 12 | F₀ = D.Discrete; 13 | F₁ = F₁; 14 | identity = λ f → Heterogeneous.≡⇒∼ _; 15 | homomorphism = λ f → Heterogeneous.≡⇒∼ _; 16 | F-resp-≡ = F-resp-≡} 17 | where 18 | F₁ : {A B : Category.Obj (Sets o)} → Sets o [ A , B ] → 19 | Categories o o _ [ D.Discrete A , D.Discrete B ] 20 | F₁ f = record { 21 | F₀ = f; 22 | F₁ = ≣-cong f; 23 | identity = _; 24 | homomorphism = _; 25 | F-resp-≡ = _ } 26 | 27 | F-resp-≡ : {A B : Set o} {F G : Sets o [ A , B ]} → 28 | Sets o [ F ≡ G ] → Categories o o _ [ F₁ F ≡ F₁ G ] 29 | F-resp-≡ F≡G {a} ≣-refl rewrite F≡G {a} = Heterogeneous.≡⇒∼ _ 30 | -------------------------------------------------------------------------------- /Categories/Functor/Equivalence/Strong.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Functor.Equivalence.Strong where 3 | 4 | open import Level 5 | open import Relation.Binary using (IsEquivalence) 6 | open import Relation.Nullary using (¬_) 7 | open import Data.Product using (Σ; ∃; _,_; proj₁; proj₂) 8 | open import Function using () renaming (_∘_ to _∙_) 9 | 10 | open import Categories.Category 11 | open import Categories.Functor hiding (equiv) 12 | open import Categories.Functor.Properties 13 | import Categories.Morphisms as Morphisms 14 | 15 | record StrongEquivalence {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} (F : Functor C D) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) where 16 | field 17 | .full : Full F 18 | .faithful : Faithful F 19 | eso : EssentiallySurjective F 20 | 21 | record StronglyEquivalent {o ℓ e o′ ℓ′ e′} (C : Category o ℓ e) (D : Category o′ ℓ′ e′) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) where 22 | field 23 | C⇒D : Functor C D 24 | strong-equivalence : StrongEquivalence C⇒D 25 | open StrongEquivalence strong-equivalence public 26 | 27 | refl : ∀ {o ℓ e} {C : Category o ℓ e} → StronglyEquivalent C C 28 | refl {C = C} = record 29 | { C⇒D = id 30 | ; strong-equivalence = record 31 | { full = λ {X Y} → my-full X Y 32 | ; faithful = λ _ _ eq → eq 33 | ; eso = λ d → d , (record { f = C.id; g = C.id; iso = record { isoˡ = C.identityˡ; isoʳ = C.identityˡ } }) 34 | } 35 | } 36 | where 37 | module C = Category C 38 | open C hiding (id) 39 | .my-full : ∀ X Y → ¬ Σ (X ⇒ Y) (λ f → ¬ ∃ (λ g → C [ g ≡ f ])) 40 | my-full X Y (f , elim) = elim (f , Equiv.refl) 41 | 42 | sym : ∀ {o ℓ e o′ ℓ′ e′} → {C : Category o ℓ e} {D : Category o′ ℓ′ e′} → StronglyEquivalent C D → StronglyEquivalent D C 43 | sym {C = C} {D = D} Op = record 44 | { C⇒D = record 45 | { F₀ = λ d → proj₁ (Op.eso d) 46 | ; F₁ = λ f → {!rev (proj₂ (Op.eso _)) ∘ f ∘ ?!} 47 | ; identity = {!!} 48 | ; homomorphism = {!!} 49 | ; F-resp-≡ = {!!} 50 | } 51 | ; strong-equivalence = record 52 | { full = {!!} 53 | ; faithful = {!!} 54 | ; eso = {!!} 55 | } 56 | } 57 | where 58 | module Op = StronglyEquivalent Op 59 | open Op using () renaming (C⇒D to F) 60 | open Functor F 61 | open Morphisms._≅_ C renaming (f to fwd; g to rev) 62 | 63 | trans : ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂ o₃ ℓ₃ e₃} {C₁ : Category o₁ ℓ₁ e₁} {C₂ : Category o₂ ℓ₂ e₂} {C₃ : Category o₃ ℓ₃ e₃} → StronglyEquivalent C₁ C₂ → StronglyEquivalent C₂ C₃ → StronglyEquivalent C₁ C₃ 64 | trans {C₁ = C₁} {C₂} {C₃} C₁⇒C₂ C₂⇒C₃ = record 65 | { C⇒D = G ∘ F 66 | ; strong-equivalence = record 67 | { full = λ {X Y} → my-full X Y 68 | ; faithful = λ f g → C₁⇒C₂.faithful _ _ ∙ C₂⇒C₃.faithful _ _ 69 | ; eso = my-eso } 70 | } 71 | where 72 | module C₁⇒C₂ = StronglyEquivalent C₁⇒C₂ 73 | module C₂⇒C₃ = StronglyEquivalent C₂⇒C₃ 74 | open C₁⇒C₂ using () renaming (C⇒D to F) 75 | open C₂⇒C₃ using () renaming (C⇒D to G) 76 | module F = Functor F 77 | open F using (F₀; F₁) 78 | module G = Functor G 79 | open G using () renaming (F₀ to G₀; F₁ to G₁) 80 | module E = Category.Equiv C₃ 81 | open Morphisms C₃ 82 | .my-full : ∀ X Y → ¬ ∃ (λ f → ¬ ∃ (λ (g : C₁ [ X , Y ]) → C₃ [ G₁ (F₁ g) ≡ f ])) 83 | my-full X Y (f , elim) = C₂⇒C₃.full (f , lemma) 84 | where 85 | lemma : ¬ ∃ (λ h → C₃ [ G₁ h ≡ f ]) 86 | lemma (h , G₁h≡f) = C₁⇒C₂.full (h , lemma₂) 87 | where 88 | lemma₂ : ¬ ∃ (λ g → C₂ [ F₁ g ≡ h ]) 89 | lemma₂ (g , F₁g≡h) = elim (g , E.trans (G.F-resp-≡ F₁g≡h) G₁h≡f) 90 | my-eso : ∀ c₃ → ∃ (λ c₁ → G₀ (F₀ c₁) ≅ c₃) 91 | my-eso c₃ with C₂⇒C₃.eso c₃ 92 | my-eso c₃ | c₂ , ff₃ with C₁⇒C₂.eso c₂ 93 | my-eso c₃ | c₂ , ff₃ | c₁ , ff₂ = c₁ , (ff₃ ⓘ resp-≅ ff₂) 94 | where open FunctorsAlways G 95 | 96 | equiv : ∀ {o ℓ e} → IsEquivalence (StronglyEquivalent {o} {ℓ} {e}) 97 | equiv = record 98 | { refl = refl 99 | ; sym = sym 100 | ; trans = trans 101 | } -------------------------------------------------------------------------------- /Categories/Functor/Monoidal.agda: -------------------------------------------------------------------------------- 1 | module Categories.Functor.Monoidal where 2 | -------------------------------------------------------------------------------- /Categories/Functor/Product.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Functor.Product where 3 | 4 | open import Categories.Category 5 | open import Categories.Functor using (Functor) 6 | import Categories.Object.Product as Product 7 | import Categories.Object.BinaryProducts as BinaryProducts 8 | 9 | -- Ugh, we should start bundling things (categories with binary products, in this case) up consistently 10 | _[_][_×-] : ∀ {o ℓ e} → (C : Category o ℓ e) → BinaryProducts.BinaryProducts C → Category.Obj C → Functor C C 11 | C [ P ][ O ×-] = record 12 | { F₀ = λ x → Product.A×B (product {O} {x}) 13 | ; F₁ = λ f → ⟨ π₁ , f ∘ π₂ ⟩ 14 | ; identity = λ {x} → identity′ {x} 15 | ; homomorphism = λ {x} {y} {z} {f} {g} → homomorphism′ {x} {y} {z} {f} {g} 16 | ; F-resp-≡ = λ f≡g → ⟨⟩-cong₂ refl (∘-resp-≡ˡ f≡g) 17 | } 18 | where 19 | open Category C 20 | open Equiv 21 | open Product C 22 | open BinaryProducts.BinaryProducts P 23 | 24 | .identity′ : {A : Obj} → ⟨ π₁ , id ∘ π₂ ⟩ ≡ id 25 | identity′ = 26 | begin 27 | ⟨ π₁ , id ∘ π₂ ⟩ 28 | ≈⟨ ⟨⟩-cong₂ refl identityˡ ⟩ 29 | ⟨ π₁ , π₂ ⟩ 30 | ≈⟨ η ⟩ 31 | id 32 | ∎ 33 | where open HomReasoning 34 | 35 | .homomorphism′ : {X Y Z : Obj} {f : X ⇒ Y} {g : Y ⇒ Z} → ⟨ π₁ , (g ∘ f) ∘ π₂ ⟩ ≡ ⟨ π₁ , g ∘ π₂ ⟩ ∘ ⟨ π₁ , f ∘ π₂ ⟩ 36 | homomorphism′ {f = f} {g} = 37 | begin 38 | ⟨ π₁ , (g ∘ f) ∘ π₂ ⟩ 39 | ↓⟨ ⟨⟩-cong₂ refl assoc ⟩ 40 | ⟨ π₁ , g ∘ (f ∘ π₂) ⟩ 41 | ↑⟨ ⟨⟩-cong₂ refl (∘-resp-≡ʳ commute₂) ⟩ 42 | ⟨ π₁ , g ∘ (π₂ ∘ ⟨ π₁ , f ∘ π₂ ⟩) ⟩ 43 | ↑⟨ ⟨⟩-cong₂ commute₁ assoc ⟩ 44 | ⟨ π₁ ∘ ⟨ π₁ , f ∘ π₂ ⟩ , (g ∘ π₂) ∘ ⟨ π₁ , f ∘ π₂ ⟩ ⟩ 45 | ↑⟨ ⟨⟩∘ ⟩ 46 | ⟨ π₁ , g ∘ π₂ ⟩ ∘ ⟨ π₁ , f ∘ π₂ ⟩ 47 | ∎ 48 | where open HomReasoning 49 | 50 | _[_][-×_] : ∀ {o ℓ e} → (C : Category o ℓ e) → BinaryProducts.BinaryProducts C → Category.Obj C → Functor C C 51 | C [ P ][-× O ] = record 52 | { F₀ = λ x → Product.A×B (product {x} {O}) 53 | ; F₁ = λ f → ⟨ f ∘ π₁ , π₂ ⟩ 54 | ; identity = λ {x} → identity′ {x} 55 | ; homomorphism = λ {x} {y} {z} {f} {g} → homomorphism′ {x} {y} {z} {f} {g} 56 | ; F-resp-≡ = λ f≡g → ⟨⟩-cong₂ (∘-resp-≡ˡ f≡g) refl 57 | } 58 | where 59 | open Category C 60 | open Equiv 61 | open Product C 62 | open BinaryProducts.BinaryProducts P 63 | 64 | .identity′ : {A : Obj} → ⟨ id ∘ π₁ , π₂ ⟩ ≡ id 65 | identity′ = 66 | begin 67 | ⟨ id ∘ π₁ , π₂ ⟩ 68 | ≈⟨ ⟨⟩-cong₂ identityˡ refl ⟩ 69 | ⟨ π₁ , π₂ ⟩ 70 | ≈⟨ η ⟩ 71 | id 72 | ∎ 73 | where open HomReasoning 74 | 75 | .homomorphism′ : {X Y Z : Obj} {f : X ⇒ Y} {g : Y ⇒ Z} → ⟨ (g ∘ f) ∘ π₁ , π₂ ⟩ ≡ ⟨ g ∘ π₁ , π₂ ⟩ ∘ ⟨ f ∘ π₁ , π₂ ⟩ 76 | homomorphism′ {f = f} {g} = 77 | begin 78 | ⟨ (g ∘ f) ∘ π₁ , π₂ ⟩ 79 | ↓⟨ ⟨⟩-cong₂ assoc refl ⟩ 80 | ⟨ g ∘ (f ∘ π₁) , π₂ ⟩ 81 | ↑⟨ ⟨⟩-cong₂ (∘-resp-≡ʳ commute₁) refl ⟩ 82 | ⟨ g ∘ (π₁ ∘ ⟨ f ∘ π₁ , π₂ ⟩) , π₂ ⟩ 83 | ↑⟨ ⟨⟩-cong₂ assoc commute₂ ⟩ 84 | ⟨ (g ∘ π₁) ∘ ⟨ f ∘ π₁ , π₂ ⟩ , π₂ ∘ ⟨ f ∘ π₁ , π₂ ⟩ ⟩ 85 | ↑⟨ ⟨⟩∘ ⟩ 86 | ⟨ g ∘ π₁ , π₂ ⟩ ∘ ⟨ f ∘ π₁ , π₂ ⟩ 87 | ∎ 88 | where open HomReasoning 89 | -------------------------------------------------------------------------------- /Categories/Functor/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Functor.Properties where 3 | 4 | open import Relation.Binary using (_Preserves_⟶_) 5 | 6 | open import Categories.Category 7 | open import Categories.Functor 8 | import Categories.Morphisms as Morphisms 9 | 10 | module FunctorsAlways {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} (F : Functor C D) where 11 | open Functor F 12 | module C = Category C 13 | module D = Category D 14 | 15 | resp-Iso : ∀ {A B} {f : C [ A , B ]} {g : C [ B , A ]} 16 | → Morphisms.Iso C f g → Morphisms.Iso D (F₁ f) (F₁ g) 17 | resp-Iso {f = f} {g} I = record 18 | { isoˡ = begin 19 | D [ F₁ g ∘ F₁ f ] 20 | ↑⟨ homomorphism ⟩ 21 | F₁ (C [ g ∘ f ]) 22 | ↓⟨ F-resp-≡ I.isoˡ ⟩ 23 | F₁ (C.id) 24 | ↓⟨ identity ⟩ 25 | D.id 26 | ∎ 27 | ; isoʳ = begin 28 | D [ F₁ f ∘ F₁ g ] 29 | ↑⟨ homomorphism ⟩ 30 | F₁ (C [ f ∘ g ]) 31 | ↓⟨ F-resp-≡ I.isoʳ ⟩ 32 | F₁ (C.id) 33 | ↓⟨ identity ⟩ 34 | D.id 35 | ∎ 36 | } 37 | where 38 | module I = Morphisms.Iso I 39 | open D.HomReasoning 40 | 41 | resp-≅ : F₀ Preserves Morphisms._≅_ C ⟶ Morphisms._≅_ D 42 | resp-≅ I = record 43 | { f = F₁ I.f 44 | ; g = F₁ I.g 45 | ; iso = resp-Iso I.iso 46 | } 47 | where 48 | module I = Morphisms._≅_ I 49 | -------------------------------------------------------------------------------- /Categories/Functor/Representable.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Functor.Representable where 3 | 4 | open import Level 5 | 6 | open import Categories.Category 7 | open import Categories.Agda 8 | open import Categories.Functor using (Functor) 9 | open import Categories.Functor.Hom 10 | open import Categories.NaturalIsomorphism 11 | 12 | record Representable {o ℓ e} (C : Category o ℓ e) : Set (o ⊔ suc ℓ ⊔ suc e) where 13 | open Category C 14 | open Hom C 15 | field 16 | F : Functor C (ISetoids ℓ e) -- should this be a parameter to the record? 17 | A : Obj 18 | Iso : NaturalIsomorphism F Hom[ A ,-] 19 | -------------------------------------------------------------------------------- /Categories/FunctorCategory.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.FunctorCategory where 3 | 4 | open import Data.Product 5 | 6 | open import Categories.Category 7 | import Categories.Functor as Cat 8 | open import Categories.Functor hiding (equiv; id; _∘_; _≡_) 9 | open import Categories.NaturalTransformation 10 | open import Categories.Product 11 | open import Categories.Square 12 | 13 | Functors : ∀ {o ℓ e} {o′ ℓ′ e′} → Category o ℓ e → Category o′ ℓ′ e′ → Category _ _ _ 14 | Functors C D = record 15 | { Obj = Functor C D 16 | ; _⇒_ = NaturalTransformation 17 | ; _≡_ = _≡_ 18 | ; _∘_ = _∘₁_ 19 | ; id = id 20 | ; assoc = λ {_} {_} {_} {_} {f} {g} {h} → assoc₁ {X = f} {g} {h} 21 | ; identityˡ = λ {_} {_} {f} → identity₁ˡ {X = f} 22 | ; identityʳ = λ {_} {_} {f} → identity₁ʳ {X = f} 23 | ; equiv = λ {F} {G} → equiv {F = F} {G = G} 24 | ; ∘-resp-≡ = λ {_} {_} {_} {f} {h} {g} {i} → ∘₁-resp-≡ {f = f} {h} {g} {i} 25 | } 26 | 27 | 28 | eval : ∀ {o ℓ e} {o′ ℓ′ e′} {C : Category o ℓ e}{D : Category o′ ℓ′ e′} → Functor (Product (Functors C D) C) D 29 | eval {C = C} {D = D} = 30 | record { 31 | F₀ = λ x → let F , c = x in F₀ F c; 32 | F₁ = λ { {F , c₁} {G , c₂} (ε , f) → F₁ G f D.∘ η ε _}; 33 | identity = λ { {F , c} → begin 34 | F₁ F C.id D.∘ D.id ↓⟨ D.identityʳ ⟩ 35 | F₁ F C.id ↓⟨ identity F ⟩ 36 | D.id ∎}; 37 | homomorphism = λ { {F , c₁} {G , c₂} {H , c₃} {ε₁ , f₁} {ε₂ , f₂} → 38 | begin 39 | F₁ H (f₂ C.∘ f₁) D.∘ η ε₂ c₁ D.∘ η ε₁ c₁ ↑⟨ D.assoc ⟩ 40 | (F₁ H (f₂ C.∘ f₁) D.∘ η ε₂ c₁) D.∘ η ε₁ c₁ 41 | ↓⟨ D.∘-resp-≡ˡ (begin 42 | F₁ H (C [ f₂ ∘ f₁ ]) D.∘ η ε₂ c₁ ↓⟨ D.∘-resp-≡ˡ (homomorphism H) ⟩ 43 | (F₁ H f₂ D.∘ F₁ H f₁) D.∘ η ε₂ c₁ ↓⟨ D.assoc ⟩ 44 | F₁ H f₂ D.∘ F₁ H f₁ D.∘ η ε₂ c₁ ↑⟨ D.∘-resp-≡ʳ (commute ε₂ f₁) ⟩ 45 | F₁ H f₂ D.∘ η ε₂ c₂ D.∘ F₁ G f₁ ↑⟨ D.assoc ⟩ 46 | (F₁ H f₂ D.∘ η ε₂ c₂) D.∘ F₁ G f₁ ∎) 47 | ⟩ 48 | ((F₁ H f₂ D.∘ η ε₂ c₂) D.∘ F₁ G f₁) D.∘ η ε₁ c₁ ↓⟨ D.assoc ⟩ 49 | (F₁ H f₂ D.∘ η ε₂ c₂) D.∘ F₁ G f₁ D.∘ η ε₁ c₁ ∎ }; 50 | F-resp-≡ = λ { {F , c₁} {G , c₂} {ε₁ , f₁} {ε₂ , f₂} (ε₁≡ε₂ , f₁≡f₂) 51 | → D.∘-resp-≡ (F-resp-≡ G f₁≡f₂) ε₁≡ε₂} } 52 | where 53 | module C = Category C 54 | module D = Category D 55 | 56 | open Functor 57 | open NaturalTransformation 58 | open D.HomReasoning 59 | 60 | Cat[-∘_] : ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂ o₃ ℓ₃ e₃} {A : Category o₁ ℓ₁ e₁} {B : Category o₂ ℓ₂ e₂} 61 | {C : Category o₃ ℓ₃ e₃} -> Functor A B -> Functor (Functors B C) (Functors A C) 62 | Cat[-∘_] {C = C} r = record 63 | { F₀ = λ X → X Cat.∘ r 64 | ; F₁ = λ η → η ∘ʳ r 65 | ; identity = C.Equiv.refl 66 | ; homomorphism = C.Equiv.refl 67 | ; F-resp-≡ = λ x → x 68 | } 69 | where 70 | module C = Category C 71 | -------------------------------------------------------------------------------- /Categories/GlobularSet.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.GlobularSet where 3 | 4 | open import Level 5 | open import Data.Unit 6 | 7 | open import Categories.Support.PropositionalEquality 8 | open import Categories.Category 9 | open import Categories.Globe 10 | open import Categories.Functor 11 | open import Categories.Presheaf 12 | open import Categories.Agda 13 | 14 | record GlobularSet (o : Level) : Set (suc o) where 15 | field 16 | presheaf : Presheaf Globe (Sets o) 17 | 18 | Trivial : GlobularSet zero 19 | Trivial = record 20 | { presheaf = record 21 | { F₀ = λ x → ⊤ 22 | ; F₁ = λ x y → y 23 | ; identity = ≣-refl 24 | ; homomorphism = ≣-refl 25 | ; F-resp-≡ = λ x → ≣-refl 26 | } 27 | } 28 | 29 | zeroCell : ∀ {ℓ} → GlobularSet ℓ → Set ℓ 30 | zeroCell G = Functor.F₀ (GlobularSet.presheaf G) 0 31 | 32 | {- 33 | lift : ∀ {ℓ} → (G : GlobularSet ℓ) → (A : Set ℓ) → (f g : zeroCell G → A) → GlobularSet ℓ 34 | lift {ℓ} G A f g = record 35 | { presheaf = record 36 | { F₀ = F₀′ 37 | ; F₁ = F₁′ 38 | ; identity = ≣-refl 39 | ; homomorphism = {!!} 40 | ; F-resp-≡ = F-resp-≡′ 41 | } 42 | } 43 | where 44 | open GlobularSet G 45 | open Functor presheaf 46 | 47 | F₀′ : ℕ → Set ℓ 48 | F₀′ zero = A 49 | F₀′ (suc n) = F₀ n 50 | 51 | F₁′ : ∀ {A B} → GlobeHom B A → F₀′ A → F₀′ B 52 | F₁′ I a = a 53 | F₁′ {suc zero} (σ Z Obj 44 | i₁ : {A B : Obj} -> A ⇒ (A + B) 45 | i₂ : {A B : Obj} -> B ⇒ (A + B) 46 | [_,_] : {A B : Obj} -> ∀ {C} → (A ⇒ C) → (B ⇒ C) → ((A + B) ⇒ C) 47 | 48 | .commute₁ : {A B : Obj} -> ∀ {C} {f : A ⇒ C} {g : B ⇒ C} → [ f , g ] ∘ i₁ ≡ f 49 | .commute₂ : {A B : Obj} -> ∀ {C} {f : A ⇒ C} {g : B ⇒ C} → [ f , g ] ∘ i₂ ≡ g 50 | .universal : {A B : Obj} -> ∀ {C} {f : A ⇒ C} {g : B ⇒ C} {h : (A + B) ⇒ C} 51 | → h ∘ i₁ ≡ f 52 | → h ∘ i₂ ≡ g 53 | → [ f , g ] ≡ h 54 | 55 | coproduct→product : ∀ {A B} → Coproduct A B → Op×.Product A B 56 | coproduct→product A+B = record 57 | { A×B = A+B.A+B 58 | ; π₁ = A+B.i₁ 59 | ; π₂ = A+B.i₂ 60 | ; ⟨_,_⟩ = A+B.[_,_] 61 | ; commute₁ = A+B.commute₁ 62 | ; commute₂ = A+B.commute₂ 63 | ; universal = A+B.universal 64 | } 65 | where 66 | module A+B = Coproduct A+B 67 | 68 | product→coproduct : ∀ {A B} → Op×.Product A B → Coproduct A B 69 | product→coproduct A×B = record 70 | { A+B = A×B.A×B 71 | ; i₁ = A×B.π₁ 72 | ; i₂ = A×B.π₂ 73 | ; [_,_] = A×B.⟨_,_⟩ 74 | ; commute₁ = A×B.commute₁ 75 | ; commute₂ = A×B.commute₂ 76 | ; universal = A×B.universal 77 | } 78 | where 79 | module A×B = Op×.Product A×B 80 | 81 | open import Categories.Morphisms 82 | 83 | Commutative : ∀ {A B} → (p₁ : Coproduct A B) (p₂ : Coproduct B A) → _≅_ C (Coproduct.A+B p₁) (Coproduct.A+B p₂) 84 | Commutative p₁ p₂ = opⁱ (Op×.Commutative (coproduct→product p₂) (coproduct→product p₁)) 85 | 86 | Associative : ∀ {X Y Z} (p₁ : Coproduct X Y) (p₂ : Coproduct Y Z) (p₃ : Coproduct X (Coproduct.A+B p₂)) (p₄ : Coproduct (Coproduct.A+B p₁) Z) → _≅_ C (Coproduct.A+B p₃) (Coproduct.A+B p₄) 87 | Associative p₁ p₂ p₃ p₄ = reverseⁱ C (opⁱ (Op×.Associative (coproduct→product p₁) (coproduct→product p₂) (coproduct→product p₃) (coproduct→product p₄))) 88 | -------------------------------------------------------------------------------- /Categories/Object/Coproducts.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Object.Coproducts {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open Category C 8 | 9 | open import Level 10 | 11 | import Categories.Object.Initial as Initial 12 | import Categories.Object.BinaryCoproducts as BinaryCoproducts 13 | 14 | open Initial C 15 | open BinaryCoproducts C 16 | 17 | -- this should really be 'FiniteCoproducts', no? 18 | record Coproducts : Set (o ⊔ ℓ ⊔ e) where 19 | field 20 | initial : Initial 21 | binary : BinaryCoproducts -------------------------------------------------------------------------------- /Categories/Object/Exponentiating/Adjunction.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | open import Categories.Category 3 | open import Categories.Object.BinaryProducts 4 | open import Categories.Object.Exponentiating 5 | 6 | module Categories.Object.Exponentiating.Adjunction {o ℓ e} 7 | (C : Category o ℓ e) 8 | (binary : BinaryProducts C) 9 | (Σ : Category.Obj C) 10 | (exponentiating : Exponentiating C binary Σ) where 11 | 12 | open Category C 13 | open BinaryProducts binary 14 | open Exponentiating exponentiating 15 | 16 | import Categories.Object.Product 17 | open Categories.Object.Product C 18 | 19 | import Categories.Object.Product.Morphisms 20 | open Categories.Object.Product.Morphisms C 21 | 22 | open Equiv 23 | open HomReasoning 24 | 25 | import Categories.Object.Exponentiating.Functor 26 | open Categories.Object.Exponentiating.Functor C binary Σ exponentiating 27 | 28 | open import Categories.Functor 29 | using (Functor; Contravariant) 30 | renaming (id to idF; _≡_ to _≡F_; _∘_ to _∘F_) 31 | 32 | open import Categories.Adjunction hiding (_≡_; id) 33 | open import Categories.NaturalTransformation 34 | using (NaturalTransformation; module NaturalTransformation) 35 | open import Categories.Monad 36 | using (Monad) 37 | 38 | Σ↑-Self-Adjunction : Adjunction (Functor.op Σ↑-Functor) Σ↑-Functor 39 | Σ↑-Self-Adjunction = record 40 | { unit = record 41 | { η = λ _ → flip id 42 | ; commute = unit-commute 43 | } 44 | ; counit = record 45 | { η = λ _ → flip id 46 | ; commute = counit-commute 47 | } 48 | ; zig = zig-zag 49 | ; zag = zig-zag 50 | } where 51 | .lem₁ : ∀{A B C D}{f : (B × C) ⇒ D}{g : A ⇒ (C × B)} 52 | → (f ∘ swap ∘ second id) ∘ g 53 | ≡ f ∘ swap ∘ g 54 | lem₁ {A}{B}{C}{D}{f}{g} = 55 | begin 56 | (f ∘ swap ∘ second id) ∘ g 57 | ↓⟨ (refl ⟩∘⟨ refl ⟩∘⟨ second-id product) ⟩∘⟨ refl ⟩ 58 | (f ∘ swap ∘ id) ∘ g 59 | ↓⟨ (refl ⟩∘⟨ identityʳ) ⟩∘⟨ refl ⟩ 60 | (f ∘ swap) ∘ g 61 | ↓⟨ assoc ⟩ 62 | f ∘ swap ∘ g 63 | ∎ 64 | 65 | .lem₂ : ∀ {X Y}{f : X ⇒ Y} 66 | → eval {Σ↑ Y} ∘ first (flip id ∘ f) 67 | ≡ eval {X} ∘ swap ∘ second [Σ↑ f ] 68 | lem₂ {X}{Y}{f} = 69 | begin 70 | eval {Σ↑ Y} ∘ first (flip id ∘ f) 71 | ↑⟨ refl ⟩∘⟨ first∘first ⟩ 72 | eval {Σ↑ Y} ∘ first (flip id) ∘ first f 73 | ↑⟨ assoc ⟩ 74 | (eval {Σ↑ Y} ∘ first (flip id)) ∘ first f 75 | ↓⟨ β ⟩∘⟨ refl ⟩ 76 | (eval {Y} ∘ swap ∘ second id) ∘ first f 77 | ↓⟨ lem₁ ⟩ 78 | eval {Y} ∘ swap ∘ first f 79 | ↓⟨ refl ⟩∘⟨ swap∘⁂ ⟩ 80 | eval {Y} ∘ second f ∘ swap 81 | ↑⟨ assoc ⟩ 82 | (eval {Y} ∘ second f) ∘ swap 83 | ↑⟨ β ⟩∘⟨ refl ⟩ 84 | (eval {X} ∘ first (λ-abs X (eval {Y} ∘ second f))) ∘ swap 85 | ↓⟨ assoc ⟩ 86 | eval {X} ∘ first (λ-abs X (eval {Y} ∘ second f)) ∘ swap 87 | ↑⟨ refl ⟩∘⟨ swap∘⁂ ⟩ 88 | eval {X} ∘ swap ∘ second (λ-abs X (eval ∘ second f)) 89 | ∎ 90 | 91 | .unit-commute : ∀ {X Y} (f : X ⇒ Y) 92 | → flip id ∘ f ≡ [Σ² f ] ∘ flip id 93 | unit-commute {X}{Y} f = 94 | begin 95 | flip id ∘ f 96 | ↓⟨ λ-unique lem₂ ⟩ 97 | flip [Σ↑ f ] 98 | ↑⟨ λ-cong lem₁ ⟩ 99 | λ-abs (Σ↑ Y) ((eval {X} ∘ swap ∘ second id) ∘ second [Σ↑ f ]) 100 | ↓⟨ λ-distrib ⟩ 101 | [Σ↑ [Σ↑ f ] ] ∘ flip id 102 | ∎ 103 | 104 | .counit-commute : ∀ {X Y} (f : X ⇒ Y) 105 | → [Σ² f ] ∘ flip id ≡ flip id ∘ f 106 | counit-commute f = sym (unit-commute f) 107 | 108 | .zig-zag : ∀{X} 109 | → id ≡ [Σ↑ flip id ] ∘ flip id 110 | zig-zag {X} = 111 | begin 112 | id 113 | ↑⟨ flip² ⟩ 114 | flip (flip id) 115 | ↑⟨ λ-cong lem₁ ⟩ 116 | λ-abs X ((eval ∘ swap ∘ second id) ∘ second (flip id)) 117 | ↓⟨ λ-distrib ⟩ 118 | [Σ↑ flip id ] ∘ flip id 119 | ∎ 120 | -------------------------------------------------------------------------------- /Categories/Object/Exponentiating/Functor.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | open import Categories.Category 3 | open import Categories.Object.BinaryProducts 4 | open import Categories.Object.Exponentiating 5 | 6 | module Categories.Object.Exponentiating.Functor {o ℓ e} 7 | (C : Category o ℓ e) 8 | (binary : BinaryProducts C) 9 | (Σ : Category.Obj C) 10 | (exponentiating : Exponentiating C binary Σ) where 11 | 12 | open Category C 13 | open BinaryProducts binary 14 | open Exponentiating exponentiating 15 | 16 | open Equiv 17 | open HomReasoning 18 | 19 | import Categories.Object.Product.Morphisms 20 | open Categories.Object.Product.Morphisms C 21 | 22 | open import Categories.Functor 23 | using (Functor; Contravariant) 24 | renaming (_∘_ to _∘F_) 25 | 26 | Σ↑-Functor : Contravariant C C 27 | Σ↑-Functor = record 28 | { F₀ = Σ↑_ 29 | ; F₁ = [Σ↑_] 30 | ; identity = identity 31 | ; homomorphism = homomorphism 32 | ; F-resp-≡ = F-resp-≡ 33 | } 34 | where 35 | .identity : ∀ {A} → [Σ↑ id {A} ] ≡ id 36 | identity {A} = 37 | begin 38 | λ-abs A (eval ∘ second id) 39 | ↓⟨ λ-cong (∘-resp-≡ refl (id⁂id product)) ⟩ 40 | λ-abs A (eval ∘ id) 41 | ↓⟨ λ-cong identityʳ ⟩ 42 | λ-abs A eval 43 | ↓⟨ λ-η-id ⟩ 44 | id 45 | ∎ 46 | 47 | .homomorphism : ∀ {X Y Z} 48 | {f : X ⇒ Y} {g : Y ⇒ Z} 49 | → [Σ↑ (g ∘ f) ] ≡ [Σ↑ f ] ∘ [Σ↑ g ] 50 | homomorphism {X}{Y}{Z}{f}{g} = 51 | begin 52 | λ-abs X (eval ∘ second (g ∘ f)) 53 | ↑⟨ λ-cong (refl ⟩∘⟨ second∘second) ⟩ 54 | λ-abs X (eval ∘ second g ∘ second f) 55 | ↑⟨ λ-cong assoc ⟩ 56 | λ-abs X ((eval ∘ second g) ∘ second f) 57 | ↓⟨ λ-distrib ⟩ 58 | λ-abs X (eval ∘ second f) 59 | ∘ 60 | λ-abs Y (eval ∘ second g) 61 | ∎ 62 | 63 | .F-resp-≡ : ∀ {A B}{f g : A ⇒ B } 64 | → f ≡ g → [Σ↑ f ] ≡ [Σ↑ g ] 65 | F-resp-≡ {A}{B}{f}{g} f≡g = 66 | begin 67 | λ-abs A (eval ∘ second f) 68 | ↓⟨ λ-cong (refl ⟩∘⟨ ⟨⟩-cong₂ refl (f≡g ⟩∘⟨ refl)) ⟩ 69 | λ-abs A (eval ∘ second g) 70 | ∎ 71 | 72 | Σ²-Functor : Functor C C 73 | Σ²-Functor = Σ↑-Functor ∘F Functor.op Σ↑-Functor 74 | -------------------------------------------------------------------------------- /Categories/Object/Indexed.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | open import Categories.Category 3 | open import Categories.Support.Equivalence 4 | 5 | module Categories.Object.Indexed {o ℓ e c q} (C : Category o ℓ e) (B : Setoid c q) where 6 | 7 | open import Categories.Support.SetoidFunctions 8 | open Category C 9 | open _⟶_ public using () renaming (cong to cong₀; _⟨$⟩_ to _!_) 10 | 11 | Objoid = set→setoid Obj 12 | 13 | Dust = B ⟶ Objoid 14 | 15 | dust-setoid = B ⇨ Objoid -------------------------------------------------------------------------------- /Categories/Object/Initial.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Object.Initial {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open Category C 8 | 9 | open import Level 10 | 11 | record Initial : Set (o ⊔ ℓ ⊔ e) where 12 | field 13 | ⊥ : Obj 14 | ! : ∀ {A} → (⊥ ⇒ A) 15 | .!-unique : ∀ {A} → (f : ⊥ ⇒ A) → ! ≡ f 16 | 17 | .!-unique₂ : ∀ {A} → (f g : ⊥ ⇒ A) → f ≡ g 18 | !-unique₂ f g = 19 | begin 20 | f 21 | ↑⟨ !-unique f ⟩ 22 | ! 23 | ↓⟨ !-unique g ⟩ 24 | g 25 | ∎ 26 | where 27 | open HomReasoning 28 | 29 | .⊥-id : (f : ⊥ ⇒ ⊥) → f ≡ id 30 | ⊥-id f = !-unique₂ f id 31 | 32 | import Categories.Morphisms 33 | open Categories.Morphisms C 34 | open Initial 35 | 36 | .to-⊥-is-Epi : ∀ {A : Obj} {i : Initial} → (f : A ⇒ ⊥ i) → Epi f 37 | to-⊥-is-Epi {_} {i} f = helper 38 | where 39 | helper : ∀ {C : Obj} → (g h : ⊥ i ⇒ C) → g ∘ f ≡ h ∘ f → g ≡ h 40 | helper g h _ = !-unique₂ i g h 41 | 42 | -------------------------------------------------------------------------------- /Categories/Object/Products.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Object.Products {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open Category C 8 | 9 | open import Level 10 | 11 | import Categories.Object.Terminal as Terminal 12 | import Categories.Object.BinaryProducts as BinaryProducts 13 | 14 | open Terminal C 15 | open BinaryProducts C 16 | 17 | -- this should really be 'FiniteProducts', no? 18 | record Products : Set (o ⊔ ℓ ⊔ e) where 19 | field 20 | terminal : Terminal 21 | binary : BinaryProducts 22 | -------------------------------------------------------------------------------- /Categories/Object/Products/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Object.Products.Properties {o ℓ e} 6 | (C : Category o ℓ e) where 7 | 8 | open Category C 9 | 10 | open import Level 11 | 12 | import Categories.Object.Terminal 13 | open Categories.Object.Terminal C 14 | 15 | import Categories.Object.BinaryProducts 16 | open Categories.Object.BinaryProducts C 17 | 18 | import Categories.Object.Products 19 | open Categories.Object.Products C 20 | 21 | import Categories.Morphisms 22 | open Categories.Morphisms C 23 | 24 | module Properties (P : Products) where 25 | open Products P 26 | 27 | open Terminal terminal 28 | open BinaryProducts binary 29 | open HomReasoning 30 | open Equiv 31 | 32 | unitˡ : ∀ {X} → (⊤ × X) ≅ X 33 | unitˡ {X} = record 34 | { f = π₂ 35 | ; g = ⟨ ! , id {X} ⟩ 36 | ; iso = record 37 | { isoˡ = 38 | begin 39 | ⟨ ! , id {X} ⟩ ∘ π₂ 40 | ↓⟨ ⟨⟩∘ ⟩ 41 | ⟨ ! ∘ π₂ , id ∘ π₂ ⟩ 42 | ↓⟨ ⟨⟩-cong₂ (!-unique₂ (! ∘ π₂) π₁) (identityˡ) ⟩ 43 | ⟨ π₁ , π₂ ⟩ 44 | ↓⟨ η ⟩ 45 | id 46 | ∎ 47 | ; isoʳ = commute₂ 48 | } 49 | } 50 | 51 | .unitˡ-natural : ∀ {X Y} {f : X ⇒ Y} → ⟨ ! , id {Y} ⟩ ∘ f ≡ second f ∘ ⟨ ! , id {X} ⟩ 52 | unitˡ-natural {f = f} = 53 | begin 54 | ⟨ ! , id ⟩ ∘ f 55 | ↓⟨ ⟨⟩∘ ⟩ 56 | ⟨ ! ∘ f , id ∘ f ⟩ 57 | ↑⟨ ⟨⟩-cong₂ (!-unique (! ∘ f)) (id-comm {f = f}) ⟩ 58 | ⟨ ! , f ∘ id ⟩ 59 | ↑⟨ second∘⟨⟩ ⟩ 60 | second f ∘ ⟨ ! , id ⟩ 61 | ∎ 62 | 63 | unitʳ : ∀ {X} → (X × ⊤) ≅ X 64 | unitʳ {X} = record 65 | { f = π₁ 66 | ; g = ⟨ id {X} , ! ⟩ 67 | ; iso = record 68 | { isoˡ = 69 | begin 70 | ⟨ id {X} , ! ⟩ ∘ π₁ 71 | ↓⟨ ⟨⟩∘ ⟩ 72 | ⟨ id ∘ π₁ , ! ∘ π₁ ⟩ 73 | ↓⟨ ⟨⟩-cong₂ (identityˡ) (!-unique₂ (! ∘ π₁) π₂) ⟩ 74 | ⟨ π₁ , π₂ ⟩ 75 | ↓⟨ η ⟩ 76 | id 77 | ∎ 78 | ; isoʳ = commute₁ 79 | } 80 | } 81 | 82 | .unitʳ-natural : ∀ {X Y} {f : X ⇒ Y} → ⟨ id , ! ⟩ ∘ f ≡ first f ∘ ⟨ id , ! ⟩ 83 | unitʳ-natural {f = f} = 84 | begin 85 | ⟨ id , ! ⟩ ∘ f 86 | ↓⟨ ⟨⟩∘ ⟩ 87 | ⟨ id ∘ f , ! ∘ f ⟩ 88 | ↑⟨ ⟨⟩-cong₂ (id-comm {f = f}) (!-unique (! ∘ f)) ⟩ 89 | ⟨ f ∘ id , ! ⟩ 90 | ↑⟨ first∘⟨⟩ ⟩ 91 | first f ∘ ⟨ id , ! ⟩ 92 | ∎ 93 | -------------------------------------------------------------------------------- /Categories/Object/SubobjectClassifier.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Object.SubobjectClassifier {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open Category C 8 | 9 | open import Level 10 | 11 | open import Categories.Object.Terminal 12 | open import Categories.Morphisms 13 | open import Categories.Pullback 14 | 15 | record SubobjectClassifier : Set (o ⊔ ℓ ⊔ e) where 16 | field 17 | Ω : Obj 18 | χ : ∀ {U X} → (j : U ⇒ X) → (X ⇒ Ω) 19 | terminal : Terminal C 20 | 21 | open Terminal terminal 22 | 23 | field 24 | ⊤⇒Ω : ⊤ ⇒ Ω 25 | .j-pullback : ∀ {U X} → (j : U ⇒ X) → Mono C j → Pullback C ⊤⇒Ω (χ j) 26 | .χ-unique : ∀ {U X} → (j : U ⇒ X) → (χ′ : X ⇒ Ω) → Mono C j → Pullback C ⊤⇒Ω χ′ → χ′ ≡ χ j 27 | 28 | -------------------------------------------------------------------------------- /Categories/Object/Terminal.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Object.Terminal {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Level 8 | open import Relation.Binary using (IsEquivalence; Setoid) 9 | open import Categories.Support.PropositionalEquality 10 | 11 | open Category C 12 | 13 | record Terminal : Set (o ⊔ ℓ ⊔ e) where 14 | field 15 | ⊤ : Obj 16 | ! : ∀ {A} → (A ⇒ ⊤) 17 | .!-unique : ∀ {A} → (f : A ⇒ ⊤) → ! ≡ f 18 | 19 | .!-unique₂ : ∀ {A} → (f g : A ⇒ ⊤) → f ≡ g 20 | !-unique₂ f g = 21 | begin 22 | f 23 | ↑⟨ !-unique f ⟩ 24 | ! 25 | ↓⟨ !-unique g ⟩ 26 | g 27 | ∎ 28 | where open HomReasoning 29 | 30 | .⊤-id : (f : ⊤ ⇒ ⊤) → f ≡ id 31 | ⊤-id f = !-unique₂ f id 32 | 33 | open Heterogeneous C 34 | 35 | !-unique∼ : ∀ {A A′} → (f : A′ ⇒ ⊤) → (A ≣ A′) → ! {A} ∼ f 36 | !-unique∼ f ≣-refl = Heterogeneous.≡⇒∼ (!-unique f) 37 | 38 | import Categories.Morphisms 39 | open Categories.Morphisms C 40 | open Terminal 41 | 42 | .from-⊤-is-Mono : ∀ {A : Obj} {t : Terminal} → (f : ⊤ t ⇒ A) → Mono f 43 | from-⊤-is-Mono {_} {t} f = helper 44 | where 45 | helper : ∀ {C : Obj} -> (g h : C ⇒ ⊤ t) → f ∘ g ≡ f ∘ h → g ≡ h 46 | helper g h _ = !-unique₂ t g h 47 | 48 | open import Categories.Square 49 | open GlueSquares C 50 | open Heterogeneous C 51 | 52 | !-is-propertylike : ∀ t₁ t₂ → ⊤ t₁ ≣ ⊤ t₂ → ∀ {A} → ! t₁ {A} ∼ ! t₂ {A} 53 | !-is-propertylike t₁ t₂ eq = helper t₁ eq (! t₂) 54 | where 55 | helper : ∀ t {T} → (⊤ t ≣ T) → ∀ {A} (f : A ⇒ T) → ! t {A} ∼ f 56 | helper t ≣-refl f = ≡⇒∼ (!-unique t f) 57 | 58 | up-to-iso : (t₁ t₂ : Terminal) → ⊤ t₁ ≅ ⊤ t₂ 59 | up-to-iso t₁ t₂ = record 60 | { f = ! t₂ 61 | ; g = ! t₁ 62 | ; iso = record { isoˡ = ⊤-id t₁ _; isoʳ = ⊤-id t₂ _ } 63 | } 64 | 65 | up-to-iso-cong₂ : {t₁ t₁′ t₂ t₂′ : Terminal} → ⊤ t₁ ≣ ⊤ t₁′ → ⊤ t₂ ≣ ⊤ t₂′ → up-to-iso t₁ t₂ ∼ⁱ up-to-iso t₁′ t₂′ 66 | up-to-iso-cong₂ {t₁} {t₁′} {t₂} {t₂′} eq₁ eq₂ = heqⁱ (up-to-iso t₁ t₂) (up-to-iso t₁′ t₂′) (helper t₂ t₂′ eq₂ eq₁) (helper t₁ t₁′ eq₁ eq₂) 67 | where 68 | helper : ∀ t t′ → (⊤ t ≣ ⊤ t′) → ∀ {A A′} → A ≣ A′ → ! t {A} ∼ ! t′ {A′} 69 | helper t t′ eq-t ≣-refl = !-is-propertylike t t′ eq-t 70 | 71 | transport-by-iso : (t : Terminal) → ∀ {X} → ⊤ t ≅ X → Terminal 72 | transport-by-iso t {X} t≅X = record 73 | { ⊤ = X 74 | ; ! = f ∘ ! t 75 | ; !-unique = λ h → let open HomReasoning in 76 | begin 77 | f ∘ ! t 78 | ↓⟨ ∘-resp-≡ʳ (!-unique t (g ∘ h)) ⟩ 79 | f ∘ (g ∘ h) 80 | ↓⟨ cancelLeft isoʳ ⟩ 81 | h 82 | ∎ 83 | } 84 | where open _≅_ t≅X 85 | 86 | transport-by-iso-cong₂ : ∀ {t t′} → (⊤ t ≣ ⊤ t′) → ∀ {X} {i : ⊤ t ≅ X} {X′} {i′ : ⊤ t′ ≅ X′} → (i ∼ⁱ i′) → ⊤ (transport-by-iso t i) ≣ ⊤ (transport-by-iso t′ i′) 87 | transport-by-iso-cong₂ eq-t eq-i = codomain-≣ⁱ eq-i 88 | 89 | .up-to-iso-unique : ∀ t t′ → (i : ⊤ t ≅ ⊤ t′) → up-to-iso t t′ ≡ⁱ i 90 | up-to-iso-unique t t′ i = record { f-≡ = !-unique t′ _; g-≡ = !-unique t _ } 91 | 92 | .up-to-iso-invˡ : ∀ {t X} {i : ⊤ t ≅ X} → up-to-iso t (transport-by-iso t i) ≡ⁱ i 93 | up-to-iso-invˡ {t} {i = i} = up-to-iso-unique t (transport-by-iso t i) i 94 | 95 | up-to-iso-invʳ : ∀ {t t′} → ⊤ (transport-by-iso t (up-to-iso t t′)) ≣ ⊤ t′ 96 | up-to-iso-invʳ {t} {t′} = ≣-refl 97 | -------------------------------------------------------------------------------- /Categories/Object/Terminal/Exponentials.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Object.Terminal 5 | open import Level 6 | 7 | module Categories.Object.Terminal.Exponentials {o ℓ e : Level} 8 | (C : Category o ℓ e) 9 | (T : Terminal C) where 10 | 11 | open Category C 12 | open Terminal T 13 | 14 | import Categories.Object.Exponential 15 | import Categories.Object.Product 16 | import Categories.Object.Product.Morphisms 17 | import Categories.Object.Terminal.Products 18 | 19 | open Categories.Object.Exponential C hiding (repack) 20 | open Categories.Object.Product C 21 | open Categories.Object.Product.Morphisms C 22 | open Categories.Object.Terminal.Products C T 23 | 24 | [_↑⊤] : Obj → Obj 25 | [ B ↑⊤] = B 26 | 27 | [_↑⊤]-exponential : (B : Obj) → Exponential ⊤ B 28 | [ B ↑⊤]-exponential = record 29 | { B^A = [ B ↑⊤] 30 | ; product = [ B ×⊤]-product 31 | ; eval = id 32 | ; λg = λ {X} p g → g ∘ repack [ X ×⊤]-product p 33 | ; β = λ {X} p {g} → 34 | begin 35 | id ∘ (g ∘ [ p ]⟨ id , ! ⟩) ∘ [ p ]π₁ 36 | ↓⟨ identityˡ ⟩ 37 | (g ∘ [ p ]⟨ id , ! ⟩) ∘ [ p ]π₁ 38 | ↓⟨ assoc ⟩ 39 | g ∘ [ p ]⟨ id , ! ⟩ ∘ [ p ]π₁ 40 | ↓⟨ refl ⟩∘⟨ Product.⟨⟩∘ p ⟩ 41 | g ∘ [ p ]⟨ id ∘ [ p ]π₁ , ! ∘ [ p ]π₁ ⟩ 42 | ↓⟨ refl ⟩∘⟨ Product.⟨⟩-cong₂ p identityˡ (!-unique₂ (! ∘ [ p ]π₁) [ p ]π₂) ⟩ 43 | g ∘ [ p ]⟨ [ p ]π₁ , [ p ]π₂ ⟩ 44 | ↓⟨ refl ⟩∘⟨ Product.η p ⟩ 45 | g ∘ id 46 | ↓⟨ identityʳ ⟩ 47 | g 48 | ∎ 49 | ; λ-unique = λ{X} p {g}{h} h-commutes → 50 | begin 51 | h 52 | ↑⟨ identityʳ ⟩ 53 | h ∘ id 54 | ↑⟨ refl ⟩∘⟨ Product.commute₁ p ⟩ 55 | h ∘ [ p ]π₁ ∘ [ p ]⟨ id , ! ⟩ 56 | ↑⟨ assoc ⟩ 57 | (h ∘ [ p ]π₁) ∘ [ p ]⟨ id , ! ⟩ 58 | ↑⟨ identityˡ ⟩∘⟨ refl ⟩ 59 | (id ∘ h ∘ [ p ]π₁) ∘ [ p ]⟨ id , ! ⟩ 60 | ↓⟨ h-commutes ⟩∘⟨ refl ⟩ 61 | g ∘ repack [ X ×⊤]-product p 62 | ∎ 63 | } 64 | where 65 | open HomReasoning 66 | open Equiv 67 | 68 | [⊤↑_] : Obj → Obj 69 | [⊤↑ B ] = ⊤ 70 | 71 | [⊤↑_]-exponential : (A : Obj) → Exponential A ⊤ 72 | [⊤↑ A ]-exponential = record 73 | { B^A = [⊤↑ A ] 74 | ; product = [⊤× A ]-product 75 | ; eval = ! 76 | ; λg = λ _ _ → ! 77 | ; β = λ _ → !-unique₂ _ _ 78 | ; λ-unique = λ _ _ → !-unique₂ _ _ 79 | } 80 | -------------------------------------------------------------------------------- /Categories/Object/Terminal/Exponentiating.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Object.Products 5 | open import Categories.Object.Terminal 6 | open import Level 7 | 8 | module Categories.Object.Terminal.Exponentiating {o ℓ e : Level} 9 | (C : Category o ℓ e) 10 | (P : Products C) where 11 | 12 | open Category C 13 | open Products P 14 | open Terminal terminal 15 | 16 | open import Categories.Object.Exponentiating 17 | import Categories.Object.Terminal.Exponentials 18 | 19 | open Categories.Object.Terminal.Exponentials C terminal 20 | 21 | ⊤-exponentiating : Exponentiating C binary ⊤ 22 | ⊤-exponentiating = record 23 | { exponential = λ {X : Obj} → [⊤↑ X ]-exponential 24 | } 25 | -------------------------------------------------------------------------------- /Categories/Object/Terminal/Products.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Object.Terminal 5 | open import Level 6 | 7 | module Categories.Object.Terminal.Products {o ℓ e : Level} 8 | (C : Category o ℓ e) 9 | (T : Terminal C) where 10 | 11 | open import Categories.Object.Product 12 | 13 | open Category C 14 | open Terminal T 15 | 16 | [⊤×⊤] : Obj 17 | [⊤×⊤] = ⊤ 18 | 19 | [⊤×⊤]-product : Product C ⊤ ⊤ 20 | [⊤×⊤]-product = record 21 | { A×B = [⊤×⊤] 22 | ; π₁ = ! 23 | ; π₂ = ! 24 | ; ⟨_,_⟩ = λ _ _ → ! 25 | ; commute₁ = !-unique₂ _ _ 26 | ; commute₂ = !-unique₂ _ _ 27 | ; universal = λ _ _ → !-unique _ 28 | } 29 | 30 | [⊤×_] : Obj → Obj 31 | [⊤× X ] = X 32 | 33 | [⊤×_]-product : (X : Obj) → Product C ⊤ X 34 | [⊤×_]-product X = record 35 | { A×B = [⊤× X ] 36 | ; π₁ = ! 37 | ; π₂ = id 38 | ; ⟨_,_⟩ = λ _ f → f 39 | ; commute₁ = !-unique₂ _ _ 40 | ; commute₂ = identityˡ 41 | ; universal = λ {A}{f}{g}{i} _ id∘i≡g → 42 | begin 43 | g 44 | ↑⟨ id∘i≡g ⟩ 45 | id ∘ i 46 | ↓⟨ identityˡ ⟩ 47 | i 48 | ∎ 49 | } where open HomReasoning 50 | 51 | [_×⊤] : Obj → Obj 52 | [ X ×⊤] = X 53 | 54 | [_×⊤]-product : (X : Obj) → Product C X ⊤ 55 | [_×⊤]-product X = record 56 | { A×B = [ X ×⊤] 57 | ; π₁ = id 58 | ; π₂ = ! 59 | ; ⟨_,_⟩ = λ f _ → f 60 | ; commute₁ = identityˡ 61 | ; commute₂ = !-unique₂ _ _ 62 | ; universal = λ {A}{f}{g}{i} id∘i≡f _ → 63 | begin 64 | f 65 | ↑⟨ id∘i≡f ⟩ 66 | id ∘ i 67 | ↓⟨ identityˡ ⟩ 68 | i 69 | ∎ 70 | } where open HomReasoning 71 | 72 | -------------------------------------------------------------------------------- /Categories/Object/Zero.agda: -------------------------------------------------------------------------------- 1 | module Categories.Object.Zero where 2 | -------------------------------------------------------------------------------- /Categories/Opposite.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Opposite where 3 | 4 | -- some properties of the opposite category. 5 | -- XXX these should probably go somewhere else, but everywhere i think of 6 | -- has other problems. ☹ 7 | 8 | open import Categories.Category 9 | open import Categories.Functor 10 | open import Categories.FunctorCategory 11 | open import Categories.NaturalTransformation 12 | open import Categories.Morphisms renaming (_≅_ to _[_≅_]) 13 | 14 | opⁱ : ∀ {o ℓ e} {C : Category o ℓ e} {A B} → C [ A ≅ B ] → Category.op C [ B ≅ A ] 15 | opⁱ {C = C} A≅B = record 16 | { f = A≅B.f 17 | ; g = A≅B.g 18 | ; iso = record { isoˡ = A≅B.isoʳ; isoʳ = A≅B.isoˡ } 19 | } 20 | where 21 | module A≅B = Categories.Morphisms._≅_ A≅B 22 | 23 | opF : ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} {A : Category o₁ ℓ₁ e₁} {B : Category o₂ ℓ₂ e₂} -> 24 | (Functor (Category.op (Functors (Category.op A) (Category.op B))) (Functors A B)) 25 | opF {A = A} {B} = record { 26 | F₀ = Functor.op; 27 | F₁ = NaturalTransformation.op; 28 | identity = Category.Equiv.refl B; 29 | homomorphism = Category.Equiv.refl B; 30 | F-resp-≡ = λ x → x } 31 | -------------------------------------------------------------------------------- /Categories/Power/NaturalTransformation.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | open import Level 3 | open import Categories.Category 4 | module Categories.Power.NaturalTransformation {o ℓ e : Level} (C : Category o ℓ e) where 5 | 6 | open import Function using () renaming (_∘_ to _∙_) 7 | open import Data.Fin using (Fin; inject+; raise) 8 | open import Data.Sum using (_⊎_; [_,_]′; inj₁; inj₂) 9 | open import Data.Product using (_,_) 10 | 11 | import Categories.Power 12 | module Pow = Categories.Power C 13 | open Pow public 14 | open import Categories.Bifunctor using (Bifunctor) 15 | 16 | open import Categories.Bifunctor.NaturalTransformation renaming (id to idⁿ; _≡_ to _≡ⁿ_) 17 | open import Categories.Functor using (module Functor) renaming (_∘_ to _∘F_) 18 | 19 | flattenPⁿ : ∀ {D : Category o ℓ e} {n m} {F G : Powerfunctor′ D (Fin n ⊎ Fin m)} (η : NaturalTransformation F G) → NaturalTransformation (flattenP F) (flattenP G) 20 | flattenPⁿ {n = n} {m} η = record 21 | { η = λ Xs → η.η (Xs ∙ pack) 22 | ; commute = λ fs → η.commute (fs ∙ pack) 23 | } 24 | where 25 | private module η = NaturalTransformation η 26 | pack = [ inject+ m , raise n ]′ 27 | 28 | reduceN′ : ∀ (H : Bifunctor C C C) {I} {F F′ : Powerendo′ I} (φ : NaturalTransformation F F′) {J} {G G′ : Powerendo′ J} (γ : NaturalTransformation G G′) → NaturalTransformation (reduce′ H F G) (reduce′ H F′ G′) 29 | reduceN′ H {I} {F} {F′} φ {J} {G} {G′} γ = record 30 | { η = my-η 31 | ; commute = λ {Xs Ys} → my-commute Xs Ys 32 | } 33 | where 34 | module F = Functor F 35 | module F′ = Functor F′ 36 | module G = Functor G 37 | module G′ = Functor G′ 38 | module φ = NaturalTransformation φ 39 | module γ = NaturalTransformation γ 40 | module H = Functor H 41 | module L = Functor (reduce′ H F G) 42 | module R = Functor (reduce′ H F′ G′) 43 | my-η : ∀ Xs → C [ L.F₀ Xs , R.F₀ Xs ] 44 | my-η Xs = H.F₁ ((φ.η (Xs ∙ inj₁)) , (γ.η (Xs ∙ inj₂))) 45 | .my-commute : ∀ Xs Ys fs → C.CommutativeSquare (L.F₁ fs) (my-η Xs) (my-η Ys) (R.F₁ fs) 46 | my-commute Xs Ys fs = begin 47 | my-η Ys ∘ L.F₁ fs 48 | ↑⟨ H.homomorphism ⟩ 49 | H.F₁ ((φ.η (Ys ∙ inj₁) ∘ F.F₁ (fs ∙ inj₁)) , (γ.η (Ys ∙ inj₂) ∘ G.F₁ (fs ∙ inj₂))) 50 | ↓⟨ H.F-resp-≡ ((φ.commute (fs ∙ inj₁)) , (γ.commute (fs ∙ inj₂))) ⟩ 51 | H.F₁ ((F′.F₁ (fs ∙ inj₁) ∘ φ.η (Xs ∙ inj₁)) , (G′.F₁ (fs ∙ inj₂) ∘ γ.η (Xs ∙ inj₂))) 52 | ↓⟨ H.homomorphism ⟩ 53 | R.F₁ fs ∘ my-η Xs 54 | ∎ 55 | where 56 | open C using (_∘_; _≡_) 57 | open C.HomReasoning 58 | 59 | reduceN : ∀ (H : Bifunctor C C C) {n} {F F′ : Powerendo n} (φ : NaturalTransformation F F′) {m} {G G′ : Powerendo m} (γ : NaturalTransformation G G′) → NaturalTransformation (reduce H F G) (reduce H F′ G′) 60 | reduceN H {n} {F = f} {f′} F {m} {g} {g′} G = flattenPⁿ {F = a} {G = b} (reduceN′ H F G) 61 | where 62 | a : Categories.Bifunctor.Functor (Categories.Power.Exp C (Fin n ⊎ Fin m)) C 63 | a = reduce′ H f g 64 | b : Categories.Bifunctor.Functor (Categories.Power.Exp C (Fin n ⊎ Fin m)) C 65 | b = reduce′ H f′ g′ 66 | 67 | 68 | overlapN : ∀ (H : Bifunctor C C C) {n} {F F′ : Powerendo n} (φ : NaturalTransformation F F′) {G G′ : Powerendo n} (γ : NaturalTransformation G G′) → NaturalTransformation (overlaps {C} {C} H F G) (overlaps {C} {C} H F′ G′) 69 | overlapN H F G = overlapN-× {D₁ = C} {D₂ = C} H F G 70 | -------------------------------------------------------------------------------- /Categories/Preorder.agda: -------------------------------------------------------------------------------- 1 | module Categories.Preorder where 2 | -------------------------------------------------------------------------------- /Categories/Presheaf.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Presheaf where 3 | 4 | open import Categories.Category 5 | open import Categories.Functor 6 | 7 | Presheaf : ∀ {o ℓ e} {o′ ℓ′ e′} (C : Category o ℓ e) (V : Category o′ ℓ′ e′) → Set _ 8 | Presheaf C V = Functor C.op V 9 | where module C = Category C -------------------------------------------------------------------------------- /Categories/Presheaves.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Presheaves where 3 | 4 | open import Level 5 | 6 | open import Categories.Category 7 | open import Categories.Agda 8 | open import Categories.FunctorCategory 9 | 10 | Presheaves : ∀ {o ℓ e : Level} → Category o ℓ e → Category _ _ _ 11 | Presheaves {o} {ℓ} {e} C = Functors (Category.op C) (ISetoids ℓ e) -------------------------------------------------------------------------------- /Categories/Product/Projections.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Level 4 | open import Categories.Category 5 | open import Categories.Product 6 | 7 | module Categories.Product.Projections 8 | {o ℓ e o′ ℓ′ e′} 9 | (C : Category o ℓ e) 10 | (D : Category o′ ℓ′ e′) 11 | where 12 | 13 | open import Categories.Functor 14 | open import Data.Product using (_×_; Σ; _,_; proj₁; proj₂; zip; map; <_,_>; swap) 15 | 16 | ∏₁ : Functor (Product C D) C 17 | ∏₁ = record 18 | { F₀ = proj₁ 19 | ; F₁ = proj₁ 20 | ; identity = refl 21 | ; homomorphism = refl 22 | ; F-resp-≡ = proj₁ 23 | } where 24 | open Category.Equiv C 25 | 26 | 27 | ∏₂ : Functor (Product C D) D 28 | ∏₂ = record 29 | { F₀ = proj₂ 30 | ; F₁ = proj₂ 31 | ; identity = refl 32 | ; homomorphism = refl 33 | ; F-resp-≡ = proj₂ 34 | } where 35 | open Category.Equiv D 36 | -------------------------------------------------------------------------------- /Categories/Product/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | {-# OPTIONS --universe-polymorphism #-} 3 | open import Level 4 | open import Categories.Category 5 | open import Categories.Product 6 | 7 | -- Parameterize over the categories in whose product we are working 8 | module Categories.Product.Properties 9 | {o ℓ e o′ ℓ′ e′} 10 | (C : Category o ℓ e) 11 | (D : Category o′ ℓ′ e′) 12 | where 13 | 14 | C×D : Category _ _ _ 15 | C×D = Product C D 16 | module C×D = Category C×D 17 | 18 | import Categories.Product.Projections 19 | open Categories.Product.Projections C D 20 | 21 | open import Categories.Functor 22 | 23 | open import Relation.Binary using (module IsEquivalence) 24 | open import Relation.Binary.PropositionalEquality as PropEq 25 | using () 26 | renaming (_≡_ to _≣_) 27 | 28 | ∏₁※∏₂≣id : (∏₁ ※ ∏₂) ≣ id 29 | ∏₁※∏₂≣id = PropEq.refl 30 | 31 | ∏₁※∏₂-identity : (∏₁ ※ ∏₂) ≡ id 32 | ∏₁※∏₂-identity h = refl 33 | where open Heterogeneous C×D 34 | 35 | ※-distrib' : 36 | ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} 37 | → {A : Category o₁ ℓ₁ e₁} 38 | → {B : Category o₂ ℓ₂ e₂} 39 | → {F : Functor B C} 40 | → {G : Functor B D} 41 | → {H : Functor A B} 42 | → ((F ∘ H) ※ (G ∘ H)) ≣ ((F ※ G) ∘ H) 43 | ※-distrib' = PropEq.refl 44 | 45 | ※-distrib : 46 | ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} 47 | → {A : Category o₁ ℓ₁ e₁} 48 | → {B : Category o₂ ℓ₂ e₂} 49 | → (F : Functor B C) 50 | → (G : Functor B D) 51 | → (H : Functor A B) 52 | → ((F ∘ H) ※ (G ∘ H)) ≡ ((F ※ G) ∘ H) 53 | ※-distrib F G H h = refl 54 | where open Heterogeneous C×D 55 | 56 | ∏₁※∏₂-distrib : 57 | ∀ {o₁ ℓ₁ e₁}{A : Category o₁ ℓ₁ e₁} 58 | → (F : Functor A C×D) 59 | → ((∏₁ ∘ F) ※ (∏₂ ∘ F)) ≡ F 60 | ∏₁※∏₂-distrib F h = refl 61 | where open Heterogeneous C×D 62 | 63 | module Lemmas where 64 | open Heterogeneous C×D 65 | open import Data.Product 66 | 67 | lem₁ : {x₁ y₁ x₂ y₂ : Category.Obj C} 68 | → {x₃ y₃ : Category.Obj D} 69 | → (f₁ : C [ x₁ , y₁ ]) 70 | → (f₂ : C [ x₂ , y₂ ]) 71 | → (g : D [ x₃ , y₃ ]) 72 | → C [ f₁ ∼ f₂ ] 73 | → C×D [ (f₁ , g) ∼ (f₂ , g) ] 74 | lem₁ f₁ f₂ g (≡⇒∼ f₁≡f₂) = ≡⇒∼ (f₁≡f₂ , Category.Equiv.refl D) 75 | 76 | lem₂ : {x₁ y₁ x₂ y₂ : Category.Obj D} 77 | → {x₃ y₃ : Category.Obj C} 78 | → (f : C [ x₃ , y₃ ]) 79 | → (g₁ : D [ x₁ , y₁ ]) 80 | → (g₂ : D [ x₂ , y₂ ]) 81 | → D [ g₁ ∼ g₂ ] 82 | → C×D [ (f , g₁) ∼ (f , g₂) ] 83 | lem₂ f g₁ g₂ (≡⇒∼ g₁≡g₂) = ≡⇒∼ (Category.Equiv.refl C , g₁≡g₂) 84 | open Lemmas 85 | 86 | ※-preserves-≡ˡ : 87 | ∀ {o₁ ℓ₁ e₁} 88 | → {A : Category o₁ ℓ₁ e₁} 89 | → (F₁ : Functor A C) 90 | → (F₂ : Functor A C) 91 | → (G : Functor A D) 92 | → (F₁ ≡ F₂) → ((F₁ ※ G) ≡ (F₂ ※ G)) 93 | ※-preserves-≡ˡ F₁ F₂ G F₁≡F₂ h = 94 | lem₁ (Functor.F₁ F₁ h) (Functor.F₁ F₂ h) (Functor.F₁ G h) (F₁≡F₂ h) 95 | 96 | ※-preserves-≡ʳ : 97 | ∀ {o₁ ℓ₁ e₁} 98 | → {A : Category o₁ ℓ₁ e₁} 99 | → (F : Functor A C) 100 | → (G₁ : Functor A D) 101 | → (G₂ : Functor A D) 102 | → (G₁ ≡ G₂) → ((F ※ G₁) ≡ (F ※ G₂)) 103 | ※-preserves-≡ʳ F G₁ G₂ G₁≡G₂ h = 104 | lem₂ (Functor.F₁ F h) (Functor.F₁ G₁ h) (Functor.F₁ G₂ h) (G₁≡G₂ h) 105 | where open Heterogeneous C×D 106 | 107 | .※-preserves-≡ : 108 | ∀ {o₁ ℓ₁ e₁} 109 | → {A : Category o₁ ℓ₁ e₁} 110 | → (F : Functor A C) 111 | → (G : Functor A C) 112 | → (H : Functor A D) 113 | → (I : Functor A D) 114 | → (F ≡ G) → (H ≡ I) → ((F ※ H) ≡ (G ※ I)) 115 | ※-preserves-≡ {A = A} F G H I F≡G H≡I 116 | = trans {i = F ※ H}{j = G ※ H}{k = G ※ I} 117 | (※-preserves-≡ˡ F G H F≡G) 118 | (※-preserves-≡ʳ G H I H≡I) 119 | where open IsEquivalence (equiv {C = A}{D = C×D}) 120 | 121 | .※-universal : 122 | ∀ {o₁ ℓ₁ e₁}{A : Category o₁ ℓ₁ e₁} 123 | → {F : Functor A C} 124 | → {G : Functor A D} 125 | → {I : Functor A C×D} 126 | → (∏₁ ∘ I ≡ F) 127 | → (∏₂ ∘ I ≡ G) 128 | → ((F ※ G) ≡ I) 129 | ※-universal {_}{_}{_}{A}{F}{G}{I} p₁ p₂ = 130 | trans {i = F ※ G}{j = (∏₁ ∘ I) ※ (∏₂ ∘ I)}{k = I} 131 | (sym {i = (∏₁ ∘ I) ※ (∏₂ ∘ I)}{j = F ※ G} 132 | (※-preserves-≡ (∏₁ ∘ I) F (∏₂ ∘ I) G p₁ p₂)) 133 | (∏₁※∏₂-distrib I) 134 | where open IsEquivalence (equiv {C = A}{D = C×D}) 135 | 136 | ※-commute₁ : 137 | ∀ {o₁ ℓ₁ e₁}{A : Category o₁ ℓ₁ e₁} 138 | → {F : Functor A C} 139 | → {G : Functor A D} 140 | → (∏₁ ∘ (F ※ G) ≡ F) 141 | ※-commute₁ h = refl 142 | where open Heterogeneous C 143 | 144 | ※-commute₂ : 145 | ∀ {o₁ ℓ₁ e₁}{A : Category o₁ ℓ₁ e₁} 146 | → {F : Functor A C} 147 | → {G : Functor A D} 148 | → (∏₂ ∘ (F ※ G) ≡ G) 149 | ※-commute₂ h = refl 150 | where open Heterogeneous D 151 | 152 | -------------------------------------------------------------------------------- /Categories/Profunctor.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Profunctor where 3 | 4 | open import Level hiding (lift) 5 | 6 | open import Categories.Category 7 | open import Categories.Agda 8 | open import Categories.Bifunctor using (Functor; Bifunctor; _∘_) 9 | open import Categories.Functor.Hom 10 | open import Categories.Lan 11 | open import Categories.Yoneda 12 | 13 | Profunctor : ∀ {o ℓ e} {o′ ℓ′ e′} → Category o ℓ e → Category o′ ℓ′ e′ → Set _ 14 | Profunctor {ℓ = ℓ} {e} {ℓ′ = ℓ′} {e′} C D = Bifunctor (Category.op D) C (ISetoids (ℓ ⊔ ℓ′) (e ⊔ e′)) 15 | 16 | id : ∀ {o ℓ e} → {C : Category o ℓ e} → Profunctor C C 17 | id {C = C} = Hom[ C ][-,-] 18 | 19 | {- 20 | _∘_ : ∀ {o ℓ e} {o′ ℓ′ e′} {o′′ ℓ′′ e′′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} {E : Category o′′ ℓ′′ e′′} 21 | → Profunctor D E → Profunctor C D → Profunctor C E 22 | F ∘ G = {!!} 23 | -} -------------------------------------------------------------------------------- /Categories/Pullback.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | open import Categories.Category 3 | 4 | module Categories.Pullback {o ℓ e} (C : Category o ℓ e) where 5 | 6 | open Category C 7 | 8 | open import Level 9 | 10 | record Pullback {X Y Z}(f : X ⇒ Z)(g : Y ⇒ Z) : Set (o ⊔ ℓ ⊔ e) where 11 | field 12 | P : Obj 13 | p₁ : P ⇒ X 14 | p₂ : P ⇒ Y 15 | 16 | field 17 | .commutes : f ∘ p₁ ≡ g ∘ p₂ 18 | 19 | universal : ∀ {Q}(q₁ : Q ⇒ X)(q₂ : Q ⇒ Y) 20 | → .(commutes : (f ∘ q₁ ≡ g ∘ q₂)) → (Q ⇒ P) 21 | 22 | .universal-unique : ∀ {Q} {q₁ : Q ⇒ X} {q₂ : Q ⇒ Y} 23 | .{commutes : f ∘ q₁ ≡ g ∘ q₂} 24 | (u : Q ⇒ P) 25 | .(p₁∘u≡q₁ : p₁ ∘ u ≡ q₁) 26 | .(p₂∘u≡q₂ : p₂ ∘ u ≡ q₂) 27 | → u ≡ universal q₁ q₂ commutes 28 | 29 | .p₁∘universal≡q₁ : ∀ {Q} {q₁ : Q ⇒ X} {q₂ : Q ⇒ Y} 30 | .{commutes : f ∘ q₁ ≡ g ∘ q₂} 31 | → (p₁ ∘ universal q₁ q₂ commutes ≡ q₁) 32 | 33 | .p₂∘universal≡q₂ : ∀ {Q} {q₁ : Q ⇒ X} {q₂ : Q ⇒ Y} 34 | .{commutes : f ∘ q₁ ≡ g ∘ q₂} 35 | → (p₂ ∘ universal q₁ q₂ commutes ≡ q₂) 36 | 37 | 38 | glue : ∀ {W X Y Z : Obj} {f0 : W ⇒ X}{f : X ⇒ Z}{g : Y ⇒ Z} 39 | → (p : Pullback f g) → Pullback f0 (Pullback.p₁ p) → Pullback (f ∘ f0) g 40 | glue {f0 = f0} {f} {g} R L = record { 41 | P = L.P; 42 | p₁ = L.p₁; 43 | p₂ = R.p₂ ∘ L.p₂; 44 | commutes = begin (f ∘ f0) ∘ L.p₁ ↓⟨ assoc ⟩ 45 | f ∘ f0 ∘ L.p₁ ↓⟨ ∘-resp-≡ʳ L.commutes ⟩ 46 | f ∘ R.p₁ ∘ L.p₂ ↑⟨ assoc ⟩ 47 | (f ∘ R.p₁) ∘ L.p₂ ↓⟨ ∘-resp-≡ˡ R.commutes ⟩ 48 | (g ∘ R.p₂) ∘ L.p₂ ↓⟨ assoc ⟩ 49 | g ∘ R.p₂ ∘ L.p₂ ∎; 50 | universal = λ q₁ q₂ commutes → L.universal q₁ (R.universal (f0 ∘ q₁) q₂ (Equiv.trans (Equiv.sym assoc) commutes)) 51 | (Equiv.sym R.p₁∘universal≡q₁); 52 | universal-unique = λ {Q} {q₁} u p₁∘u≡q₁ p₂∘u≡q₂ → 53 | L.universal-unique u p₁∘u≡q₁ 54 | (R.universal-unique (L.p₂ ∘ u) (begin 55 | R.p₁ ∘ L.p₂ ∘ u ↑⟨ assoc ⟩ 56 | (R.p₁ ∘ L.p₂) ∘ u ↑⟨ ∘-resp-≡ˡ L.commutes ⟩ 57 | (f0 ∘ L.p₁) ∘ u ↓⟨ assoc ⟩ 58 | f0 ∘ L.p₁ ∘ u ↓⟨ ∘-resp-≡ʳ p₁∘u≡q₁ ⟩ 59 | f0 ∘ q₁ ∎) 60 | (Equiv.trans (Equiv.sym assoc) p₂∘u≡q₂)); 61 | p₁∘universal≡q₁ = L.p₁∘universal≡q₁; 62 | p₂∘universal≡q₂ = Equiv.trans assoc (Equiv.trans (∘-resp-≡ʳ L.p₂∘universal≡q₂) R.p₂∘universal≡q₂) } 63 | where 64 | module L = Pullback L 65 | module R = Pullback R 66 | open HomReasoning 67 | -------------------------------------------------------------------------------- /Categories/Pushout.agda: -------------------------------------------------------------------------------- 1 | module Categories.Pushout where 2 | -------------------------------------------------------------------------------- /Categories/Ran.agda: -------------------------------------------------------------------------------- 1 | module Categories.Ran where 2 | 3 | open import Level 4 | open import Categories.Category 5 | open import Categories.Functor hiding (_≡_) 6 | open import Categories.NaturalTransformation 7 | 8 | record Ran {o₀ ℓ₀ e₀} {o₁ ℓ₁ e₁} {o₂ ℓ₂ e₂} 9 | {A : Category o₀ ℓ₀ e₀} {B : Category o₁ ℓ₁ e₁} {C : Category o₂ ℓ₂ e₂} 10 | (F : Functor A B) (X : Functor A C) : Set (o₀ ⊔ ℓ₀ ⊔ e₀ ⊔ o₁ ⊔ ℓ₁ ⊔ e₁ ⊔ o₂ ⊔ ℓ₂ ⊔ e₂) where 11 | field 12 | R : Functor B C 13 | η : NaturalTransformation (R ∘ F) X 14 | 15 | δ : (M : Functor B C) → (α : NaturalTransformation (M ∘ F) X) → NaturalTransformation M R 16 | 17 | .δ-unique : {M : Functor B C} → {α : NaturalTransformation (M ∘ F) X} → 18 | (δ′ : NaturalTransformation M R) → α ≡ η ∘₁ (δ′ ∘ʳ F) → δ′ ≡ δ M α 19 | .commutes : (M : Functor B C) → (α : NaturalTransformation (M ∘ F) X) → α ≡ η ∘₁ (δ M α ∘ʳ F) 20 | -------------------------------------------------------------------------------- /Categories/Slice.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Slice {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open Category C 8 | open Equiv 9 | 10 | open import Level 11 | open import Relation.Binary using (Rel) 12 | 13 | record SliceObj (X : Obj) : Set (o ⊔ ℓ) where 14 | constructor sliceobj 15 | field 16 | {Y} : Obj 17 | arr : Y ⇒ X 18 | 19 | record Slice⇒ {A : Obj} (X Y : SliceObj A) : Set (ℓ ⊔ e) where 20 | constructor slicearr 21 | module X = SliceObj X 22 | module Y = SliceObj Y 23 | field 24 | {h} : X.Y ⇒ Y.Y 25 | .triangle : Y.arr ∘ h ≡ X.arr 26 | 27 | slice : Obj → Category _ _ _ 28 | slice x = record 29 | { Obj = Obj′ 30 | ; _⇒_ = _⇒′_ 31 | ; _≡_ = _≡′_ 32 | ; _∘_ = _∘′_ 33 | ; id = slicearr identityʳ 34 | ; assoc = λ {_} {_} {_} {_} {f} {g} {h} → assoc′ {f = f} {g} {h} 35 | ; identityˡ = identityˡ 36 | ; identityʳ = identityʳ 37 | ; equiv = record 38 | { refl = refl 39 | ; sym = sym 40 | ; trans = trans 41 | } 42 | ; ∘-resp-≡ = λ {_} {_} {_} {f} {h} {g} {i} → ∘-resp-≡′ {f = f} {h} {g} {i} 43 | } 44 | where 45 | Obj′ = SliceObj x 46 | 47 | _⇒′_ : Rel Obj′ _ 48 | _⇒′_ = Slice⇒ 49 | 50 | infixr 9 _∘′_ 51 | infix 4 _≡′_ 52 | 53 | _≡′_ : ∀ {A B} → Rel (A ⇒′ B) _ 54 | slicearr {f} _ ≡′ slicearr {g} _ = f ≡ g 55 | 56 | _∘′_ : ∀ {A B C} → (B ⇒′ C) → (A ⇒′ B) → (A ⇒′ C) 57 | _∘′_ {sliceobj a⇒x} {sliceobj b⇒x} {sliceobj c⇒x} (slicearr {b⇒c} pf₁) (slicearr {a⇒b} pf₂) = slicearr pf 58 | where 59 | .pf : c⇒x ∘ (b⇒c ∘ a⇒b) ≡ a⇒x 60 | pf = begin 61 | c⇒x ∘ (b⇒c ∘ a⇒b) 62 | ↑⟨ assoc ⟩ 63 | (c⇒x ∘ b⇒c) ∘ a⇒b 64 | ↓⟨ ∘-resp-≡ˡ pf₁ ⟩ 65 | b⇒x ∘ a⇒b 66 | ↓⟨ pf₂ ⟩ 67 | a⇒x 68 | ∎ 69 | where 70 | open HomReasoning 71 | 72 | 73 | .assoc′ : ∀ {A B C D} {f : A ⇒′ B} {g : B ⇒′ C} {h : C ⇒′ D} 74 | → (h ∘′ g) ∘′ f ≡′ h ∘′ (g ∘′ f) 75 | assoc′ = assoc 76 | 77 | .∘-resp-≡′ : ∀ {A B C} {f h : B ⇒′ C} {g i : A ⇒′ B} 78 | → f ≡′ h 79 | → g ≡′ i 80 | → f ∘′ g ≡′ h ∘′ i 81 | ∘-resp-≡′ f≡h g≡i = ∘-resp-≡ f≡h g≡i 82 | -------------------------------------------------------------------------------- /Categories/SubCategory.agda: -------------------------------------------------------------------------------- 1 | module Categories.SubCategory where 2 | 3 | open import Categories.Category 4 | open import Data.Product 5 | 6 | sub-category : ∀ {o ℓ e o′ ℓ′} -> (C : Category o ℓ e) -> let module C = Category C in 7 | {A : Set o′} (U : A -> C.Obj) (R : ∀ {a b} -> U a C.⇒ U b -> Set ℓ′) -> 8 | (∀ {a} -> R (C.id {U a})) -> (∀ {a b c} {f : U b C.⇒ U c} {g : U a C.⇒ U b} -> R f -> R g -> R (f C.∘ g)) → 9 | Category _ _ _ 10 | sub-category C {A} U R Rid R∘ = record 11 | { Obj = A 12 | ; _⇒_ = λ a b → Σ (U a C.⇒ U b) R 13 | ; _≡_ = λ f g → proj₁ f C.≡ proj₁ g 14 | ; id = C.id , Rid 15 | ; _∘_ = λ f g → (proj₁ f C.∘ proj₁ g) , R∘ (proj₂ f) (proj₂ g) 16 | ; assoc = C.assoc 17 | ; identityˡ = C.identityˡ 18 | ; identityʳ = C.identityʳ 19 | ; equiv = record 20 | { refl = C.Equiv.refl 21 | ; sym = C.Equiv.sym 22 | ; trans = C.Equiv.trans 23 | } 24 | ; ∘-resp-≡ = C.∘-resp-≡ 25 | } 26 | where 27 | module C = Category C 28 | -------------------------------------------------------------------------------- /Categories/Support/EqReasoning.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Support.EqReasoning where 3 | 4 | open import Categories.Support.Equivalence using (Setoid; module Setoid) 5 | open import Relation.Binary.PropositionalEquality using () renaming (_≡_ to _≣_; trans to ≣-trans; sym to ≣-sym; refl to ≣-refl) 6 | 7 | module SetoidReasoning {s₁ s₂} (S : Setoid s₁ s₂) where 8 | open Setoid S 9 | 10 | infix 4 _IsRelatedTo_ 11 | infix 1 begin_ 12 | infixr 2 _≈⟨_⟩_ _↓⟨_⟩_ _↑⟨_⟩_ _↓≣⟨_⟩_ _↑≣⟨_⟩_ _↕_ 13 | infix 3 _∎ 14 | 15 | -- This seemingly unnecessary type is used to make it possible to 16 | -- infer arguments even if the underlying equality evaluates. 17 | 18 | data _IsRelatedTo_ (x y : Carrier) : Set s₂ where 19 | relTo : (x∼y : x ≈ y) → x IsRelatedTo y 20 | 21 | .begin_ : ∀ {x y} → x IsRelatedTo y → x ≈ y 22 | begin relTo x∼y = x∼y 23 | 24 | ._↓⟨_⟩_ : ∀ x {y z} → x ≈ y → y IsRelatedTo z → x IsRelatedTo z 25 | _ ↓⟨ x∼y ⟩ relTo y∼z = relTo (trans x∼y y∼z) 26 | -- where open IsEquivalence isEquivalence 27 | 28 | ._↑⟨_⟩_ : ∀ x {y z} → y ≈ x → y IsRelatedTo z → x IsRelatedTo z 29 | _ ↑⟨ y∼x ⟩ relTo y∼z = relTo (trans (sym y∼x) y∼z) 30 | -- where open IsEquivalence isEquivalence 31 | 32 | -- the syntax of the ancients, for compatibility 33 | ._≈⟨_⟩_ : ∀ x {y z} → x ≈ y → y IsRelatedTo z → x IsRelatedTo z 34 | _≈⟨_⟩_ = _↓⟨_⟩_ 35 | 36 | ._↓≣⟨_⟩_ : ∀ x {y z} → x ≣ y → y IsRelatedTo z → x IsRelatedTo z 37 | _ ↓≣⟨ ≣-refl ⟩ y∼z = y∼z 38 | 39 | ._↑≣⟨_⟩_ : ∀ x {y z} → y ≣ x → y IsRelatedTo z → x IsRelatedTo z 40 | _ ↑≣⟨ ≣-refl ⟩ y∼z = y∼z 41 | 42 | ._↕_ : ∀ x {z} → x IsRelatedTo z → x IsRelatedTo z 43 | _ ↕ x∼z = x∼z 44 | 45 | ._∎ : ∀ x → x IsRelatedTo x 46 | _∎ _ = relTo refl 47 | -- where open IsEquivalence isEquivalence 48 | 49 | module ≣-reasoning {ℓ} (S : Set ℓ) where 50 | infix 4 _IsRelatedTo_ 51 | infix 3 _∎ 52 | infixr 2 _≈⟨_⟩_ 53 | infixr 2 _↓⟨_⟩_ 54 | infixr 2 _↑⟨_⟩_ 55 | infixr 2 _↕_ 56 | infix 1 begin_ 57 | 58 | -- This seemingly unnecessary type is used to make it possible to 59 | -- infer arguments even if the underlying equality evaluates. 60 | 61 | data _IsRelatedTo_ (x y : S) : Set ℓ where 62 | relTo : (x∼y : x ≣ y) → x IsRelatedTo y 63 | 64 | begin_ : ∀ {x y} → x IsRelatedTo y → x ≣ y 65 | begin relTo x∼y = x∼y 66 | 67 | -- the syntax of the ancients, for compatibility 68 | _≈⟨_⟩_ : ∀ x {y z} → x ≣ y → y IsRelatedTo z → x IsRelatedTo z 69 | _ ≈⟨ x∼y ⟩ relTo y∼z = relTo (≣-trans x∼y y∼z) 70 | 71 | _↓⟨_⟩_ : ∀ x {y z} → x ≣ y → y IsRelatedTo z → x IsRelatedTo z 72 | _ ↓⟨ x∼y ⟩ relTo y∼z = relTo (≣-trans x∼y y∼z) 73 | 74 | _↑⟨_⟩_ : ∀ x {y z} → y ≣ x → y IsRelatedTo z → x IsRelatedTo z 75 | _ ↑⟨ y∼x ⟩ relTo y∼z = relTo (≣-trans (≣-sym y∼x) y∼z) 76 | 77 | _↕_ : ∀ x {z} → x IsRelatedTo z → x IsRelatedTo z 78 | _ ↕ x∼z = x∼z 79 | 80 | _∎ : ∀ x → x IsRelatedTo x 81 | _∎ _ = relTo ≣-refl 82 | 83 | -------------------------------------------------------------------------------- /Categories/Support/Equivalence.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism --irrelevant-projections #-} 2 | module Categories.Support.Equivalence where 3 | 4 | open import Level 5 | open import Relation.Binary using (Rel; IsEquivalence; module IsEquivalence; IsPreorder; Preorder; Reflexive; Transitive; Symmetric; _⇒_) renaming (Setoid to RSetoid; module Setoid to RSetoid) 6 | open import Data.Product using (_×_; _,_) 7 | open import Data.Product.Relation.Pointwise.NonDependent using (×-isEquivalence) 8 | open import Relation.Binary.PropositionalEquality as PropEq using (_≡_) 9 | 10 | ------------------------------------------------------------------------ 11 | -- Setoids 12 | 13 | -- Equivalence relations are defined in Relation.Binary.Core. 14 | 15 | record Setoid c ℓ : Set (suc (c ⊔ ℓ)) where 16 | infix 4 _≈_ 17 | field 18 | Carrier : Set c 19 | _≈_ : Rel Carrier ℓ 20 | .isEquivalence : IsEquivalence _≈_ 21 | 22 | .refl : Reflexive _≈_ 23 | refl = IsEquivalence.refl isEquivalence 24 | .trans : Transitive _≈_ 25 | trans = IsEquivalence.trans isEquivalence 26 | .sym : Symmetric _≈_ 27 | sym = IsEquivalence.sym isEquivalence 28 | .reflexive : _≡_ ⇒ _≈_ 29 | reflexive = IsEquivalence.reflexive isEquivalence 30 | 31 | .isPreorder : IsPreorder _≡_ _≈_ 32 | isPreorder = record 33 | { isEquivalence = PropEq.isEquivalence 34 | ; reflexive = reflexive 35 | ; trans = trans 36 | } 37 | 38 | .preorder : Preorder c c ℓ 39 | preorder = record { isPreorder = isPreorder } 40 | 41 | -- A trivially indexed setoid. 42 | 43 | {- 44 | indexedSetoid : ∀ {i} {I : Set i} → I.Setoid I c _ 45 | indexedSetoid = record 46 | { Carrier = λ _ → Carrier 47 | ; _≈_ = _≈_ 48 | ; isEquivalence = record 49 | { refl = refl 50 | ; sym = sym 51 | ; trans = trans 52 | } 53 | } 54 | -} 55 | 56 | setoid-r→i : ∀ {c ℓ} → RSetoid c ℓ → Setoid c ℓ 57 | setoid-r→i Base = record 58 | { Carrier = Base.Carrier 59 | ; _≈_ = Base._≈_ 60 | ; isEquivalence = Base.isEquivalence 61 | } 62 | where module Base = RSetoid Base 63 | 64 | record I→R-Wrapper {c ℓ} {Carrier : Set c} (Base : Rel Carrier ℓ) (x y : Carrier) : Set ℓ where 65 | constructor squash 66 | field 67 | .anonymous-witness : Base x y 68 | 69 | setoid-i→r : ∀ {c ℓ} → Setoid c ℓ → RSetoid c ℓ 70 | setoid-i→r Base = record 71 | { Carrier = Base.Carrier 72 | ; _≈_ = I→R-Wrapper Base._≈_ 73 | ; isEquivalence = record 74 | { refl = squash Base.refl 75 | ; sym = λ pf → squash (Base.sym (anonymous-witness pf)) 76 | ; trans = λ i≈j j≈k → squash (Base.trans (anonymous-witness i≈j) (anonymous-witness j≈k)) 77 | } 78 | } 79 | where 80 | module Base = Setoid Base 81 | open I→R-Wrapper 82 | 83 | set→setoid : ∀ {ℓ} → Set ℓ → Setoid ℓ ℓ 84 | set→setoid Base = record 85 | { Carrier = Base 86 | ; _≈_ = _≡_ 87 | ; isEquivalence = PropEq.isEquivalence 88 | } 89 | 90 | _×-setoid_ : ∀ {s₁ s₂ s₃ s₄} → Setoid s₁ s₂ → Setoid s₃ s₄ → Setoid _ _ 91 | S₁ ×-setoid S₂ = record 92 | { isEquivalence = ×-isEquivalence (isEquivalence S₁) (isEquivalence S₂) 93 | } where open Setoid 94 | 95 | Lift-setoid : ∀ {c ℓ a b} -> Setoid c ℓ -> Setoid (c ⊔ a) (ℓ ⊔ b) 96 | Lift-setoid {c} {ℓ} {a} {b} s = record { 97 | Carrier = Lift {c} a Carrier; 98 | _≈_ = λ x₁ x₂ → Lift {ℓ} b (lower x₁ ≈ lower x₂); 99 | isEquivalence = record { 100 | refl = lift refl; 101 | sym = λ x₁ → lift (sym (lower x₁)); 102 | trans = λ x₁ x₂ → lift (trans (lower x₁) (lower x₂))}} 103 | where 104 | open Setoid s 105 | 106 | ∀[_]-setoid_ : ∀ {ℓ s₁ s₂} → (A : Set ℓ) → (A → Setoid s₁ s₂) → Setoid _ _ 107 | ∀[ A ]-setoid B = record 108 | { Carrier = ∀ a → B.Carrier a 109 | ; _≈_ = λ f g → ∀ a → B._≈_ a (f a) (g a) 110 | ; isEquivalence = record 111 | { refl = λ a → B.refl a 112 | ; sym = λ f≈g a → B.sym a (f≈g a) 113 | ; trans = λ f≈g g≈h a → B.trans a (f≈g a) (g≈h a) 114 | } 115 | } 116 | where 117 | module B a = Setoid (B a) 118 | 119 | Fam-setoid : ∀ {ℓ s₁ s₂} → (A : Set ℓ) → (B : Setoid s₁ s₂) → (A → Setoid.Carrier B) → Setoid _ _ 120 | Fam-setoid A B p = record 121 | { Carrier = A 122 | ; _≈_ = λ a₁ a₂ → p a₁ ≈ p a₂ 123 | ; isEquivalence = record 124 | { refl = refl 125 | ; sym = sym 126 | ; trans = trans 127 | } 128 | } 129 | where 130 | open Setoid B 131 | -------------------------------------------------------------------------------- /Categories/Support/Experimental.agda: -------------------------------------------------------------------------------- 1 | module Categories.Support.Experimental where 2 | 3 | open import Relation.Binary.PropositionalEquality.TrustMe 4 | 5 | open import Categories.Support.PropositionalEquality 6 | 7 | ≣-relevant : ∀ {l} {A : Set l} {X Y : A} -> .(X ≣ Y) -> X ≣ Y 8 | ≣-relevant _ = trustMe 9 | 10 | private 11 | ≣-coe : ∀ {a} {A B : Set a} → (A ≣ B) -> A -> B 12 | ≣-coe ≣-refl a = a 13 | 14 | ≣-subst′ : ∀ {a p} {A : Set a} → (P : A -> Set p) -> {x y : A} -> .(x ≣ y) -> P x -> P y 15 | ≣-subst′ P eq Px = ≣-coe (≣-relevant (≣-cong P eq)) Px 16 | -------------------------------------------------------------------------------- /Categories/Support/IProduct.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | module Categories.Support.IProduct where 4 | 5 | open import Function 6 | open import Level 7 | open import Relation.Nullary 8 | open import Data.Product using (∃; _×_) 9 | 10 | infixr 4 _,_ _,″_ 11 | infix 4 ,′_ 12 | infixr 2 _×′_ _-×′-_ _-,″-_ 13 | 14 | ------------------------------------------------------------------------ 15 | -- Definition 16 | 17 | record Σ′ {a b} (A : Set a) (B : A → Set b) : Set (a ⊔ b) where 18 | constructor _,_ 19 | field 20 | proj₁′ : A 21 | .proj₂′ : B proj₁′ 22 | 23 | open Σ′ public 24 | 25 | syntax Σ′ A (λ x → B) = Σ′[ x ∶ A ] B 26 | 27 | ∃′ : ∀ {a b} {A : Set a} → (A → Set b) → Set (a ⊔ b) 28 | ∃′ = Σ′ _ 29 | 30 | ∄′ : ∀ {a b} {A : Set a} → (A → Set b) → Set (a ⊔ b) 31 | ∄′ P = ¬ ∃′ P 32 | 33 | ∃₂′ : ∀ {a b c} {A : Set a} {B : A → Set b} 34 | (C : (x : A) → B x → Set c) → Set (a ⊔ b ⊔ c) 35 | ∃₂′ C = ∃ λ a → ∃′ λ b → C a b 36 | 37 | _×′_ : ∀ {a b} (A : Set a) (B : Set b) → Set (a ⊔ b) 38 | A ×′ B = Σ′[ x ∶ A ] B 39 | 40 | _,″_ : ∀ {a b} {A : Set a} {B : Set b} → A → B → A ×′ B 41 | _,″_ = _,_ 42 | 43 | ------------------------------------------------------------------------ 44 | -- Unique existence 45 | 46 | -- Parametrised on the underlying equality. 47 | 48 | ∃′! : ∀ {a b ℓ} {A : Set a} → 49 | (A → A → Set ℓ) → (A → Set b) → Set (a ⊔ b ⊔ ℓ) 50 | ∃′! _≈_ B = ∃′ λ x → B x × (∀ {y} → B y → x ≈ y) 51 | 52 | ------------------------------------------------------------------------ 53 | -- Functions 54 | 55 | -- Sometimes the first component can be inferred. 56 | 57 | ,′_ : ∀ {a b} {A : Set a} {B : A → Set b} {x} → B x → ∃′ B 58 | ,′ y = _ , y 59 | 60 | <_,_>′ : ∀ {a b c} {A : Set a} {B : A → Set b} {C : ∀ {x} → B x → Set c} 61 | (f : (x : A) → B x) → ((x : A) → C (f x)) → 62 | ((x : A) → Σ′ (B x) C) 63 | < f , g >′ x = (f x , g x) 64 | 65 | map′ : ∀ {a b p q} 66 | {A : Set a} {B : Set b} {P : A → Set p} {Q : B → Set q} → 67 | (f : A → B) → (∀ {x} → P x → Q (f x)) → 68 | Σ′ A P → Σ′ B Q 69 | map′ f g (x , y) = (f x , g y) 70 | 71 | zip′ : ∀ {a b c p q r} 72 | {A : Set a} {B : Set b} {C : Set c} 73 | {P : A → Set p} {Q : B → Set q} {R : C → Set r} → 74 | (_∙_ : A → B → C) → 75 | (∀ {x y} → P x → Q y → R (x ∙ y)) → 76 | Σ′ A P → Σ′ B Q → Σ′ C R 77 | zip′ _∙_ _∘_ (a , p) (b , q) = (a ∙ b) , (p ∘ q) 78 | 79 | _-×′-_ : ∀ {a b i j} {A : Set a} {B : Set b} → 80 | (A → B → Set i) → (A → B → Set j) → (A → B → Set _) 81 | f -×′- g = f -[ _×′_ ]- g 82 | 83 | _-,″-_ : ∀ {a b c d} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → 84 | (A → B → C) → (A → B → D) → (A → B → C ×′ D) 85 | f -,″- g = f -[ _,″_ ]- g 86 | 87 | curry′ : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Σ′ A B → Set c} → 88 | ((p : Σ′ A B) → C p) → 89 | ((x : A) → (y : B x) → C (x , y)) 90 | curry′ f x y = f (x , y) 91 | 92 | {- 93 | uncurry : ∀ {a b c} {A : Set a} {B : A → Set b} {C : Σ A B → Set c} → 94 | ((x : A) → (y : B x) → C (x , y)) → 95 | ((p : Σ A B) → C p) 96 | uncurry f (x , y) = f x y 97 | 98 | uncurry′ : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → 99 | (A → B → C) → (A × B → C) 100 | uncurry′ = uncurry 101 | -} 102 | -------------------------------------------------------------------------------- /Categories/Support/PropositionalEquality.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Support.PropositionalEquality where 3 | 4 | open import Function using (flip; id) 5 | open import Relation.Binary.PropositionalEquality public using () renaming (_≡_ to _≣_; refl to ≣-refl; trans to ≣-trans; sym to ≣-sym; cong to ≣-cong; cong₂ to ≣-cong₂; subst to ≣-subst; subst₂ to ≣-subst₂; isEquivalence to ≣-isEquivalence; setoid to ≣-setoid) 6 | 7 | module ≣-reasoning = Relation.Binary.PropositionalEquality.≡-Reasoning renaming (_≡⟨_⟩_ to _≣⟨_⟩_) 8 | 9 | ≣-app : ∀ {a} {A : Set a} {b} {B : A → Set b} {f g : (x : A) → B x} → f ≣ g → (x : A) → f x ≣ g x 10 | ≣-app {f = f} {g} = flip (λ x → ≣-cong (flip id x)) 11 | 12 | ≣-appʰ : ∀ {a} {A : Set a} {b} {B : A → Set b} {f g : ∀ {x} → B x} → (λ {x} → f) ≣ g → ∀ {x} → f {x} ≣ g {x} 13 | ≣-appʰ {f = f} {g} = λ pf {x} → ≣-cong (λ h → h {x}) pf 14 | 15 | ≣-subst₂-breakdown-lr : ∀ {a} {A : Set a} {b} {B : Set b} {ℓ} (f : A → B → Set ℓ) {a₁ a₂ : A} (a₁≣a₂ : a₁ ≣ a₂) {b₁ b₂ : B} (b₁≣b₂ : b₁ ≣ b₂) (x : f a₁ b₁) → ≣-subst₂ f a₁≣a₂ b₁≣b₂ x ≣ ≣-subst (f a₂) b₁≣b₂ (≣-subst (λ y → f y b₁) a₁≣a₂ x) 16 | ≣-subst₂-breakdown-lr f ≣-refl ≣-refl x = ≣-refl 17 | 18 | ≣-subst₂-breakdown-rl : ∀ {a} {A : Set a} {b} {B : Set b} {ℓ} (f : A → B → Set ℓ) {a₁ a₂ : A} (a₁≣a₂ : a₁ ≣ a₂) {b₁ b₂ : B} (b₁≣b₂ : b₁ ≣ b₂) (x : f a₁ b₁) → ≣-subst₂ f a₁≣a₂ b₁≣b₂ x ≣ ≣-subst (λ y → f y b₂) a₁≣a₂ (≣-subst (f a₁) b₁≣b₂ x) 19 | ≣-subst₂-breakdown-rl f ≣-refl ≣-refl x = ≣-refl 20 | 21 | ≣-subst-trans : ∀ {a} {A : Set a} {ℓ} (f : A → Set ℓ) {a₁ a₂ a₃ : A} (a₁≣a₂ : a₁ ≣ a₂) (a₂≣a₃ : a₂ ≣ a₃) (x : f a₁) → ≣-subst f (≣-trans a₁≣a₂ a₂≣a₃) x ≣ ≣-subst f a₂≣a₃ (≣-subst f a₁≣a₂ x) 22 | ≣-subst-trans f ≣-refl ≣-refl x = ≣-refl 23 | 24 | ≣-cong₂₊ : ∀ {a b c r} {A : Set a} {B : Set b} {C : A → B → Set c} {R : Set r} (f : (x : A) → (u : B) → .(z : C x u) → R) {x y : A} {u v : B} .(z : C x u) → (x≣y : x ≣ y) (u≣v : u ≣ v) → f x u z ≣ f y v (≣-subst₂ C x≣y u≣v z) 25 | ≣-cong₂₊ f z ≣-refl ≣-refl = ≣-refl 26 | -------------------------------------------------------------------------------- /Categories/Support/SetoidFunctions.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | module Categories.Support.SetoidFunctions where 4 | 5 | open import Level 6 | open import Function using (_on_) 7 | open import Relation.Binary as B using (_=[_]⇒_) 8 | open import Relation.Binary.PropositionalEquality using () renaming (cong to ≡-cong; cong₂ to ≡-cong₂) 9 | open import Data.Product using (<_,_>; _×_; _,_) 10 | 11 | open import Categories.Support.Equivalence 12 | 13 | infixr 0 _⟶_ 14 | 15 | record _⟶_ {cf ℓf ct ℓt} (From : Setoid cf ℓf) (To : Setoid ct ℓt) : 16 | Set (cf ⊔ ℓf ⊔ ct ⊔ ℓt) where 17 | infixl 5 _⟨$⟩_ 18 | field 19 | _⟨$⟩_ : (x : Setoid.Carrier From) → Setoid.Carrier To 20 | .cong : Setoid._≈_ From =[ _⟨$⟩_ ]⇒ Setoid._≈_ To 21 | 22 | open _⟶_ public 23 | 24 | id : ∀ {a₁ a₂} {A : Setoid a₁ a₂} → A ⟶ A 25 | id = record { _⟨$⟩_ = Function.id; cong = Function.id } 26 | 27 | infixr 9 _∙_ 28 | 29 | _∙_ : ∀ {ca ℓa} {A : Setoid ca ℓa} 30 | {cb ℓb} {B : Setoid cb ℓb} 31 | {cc ℓc} {C : Setoid cc ℓc} → 32 | B ⟶ C → A ⟶ B → A ⟶ C 33 | f ∙ g = record 34 | { _⟨$⟩_ = Function._∘_ (_⟨$⟩_ f) (_⟨$⟩_ g) 35 | ; cong = Function._∘_ (cong f) (cong g) 36 | } 37 | 38 | const : ∀ {ca ℓa} {A : Setoid ca ℓa} 39 | {cb ℓb} {B : Setoid cb ℓb} → 40 | Setoid.Carrier B → A ⟶ B 41 | const {B = B} b = record 42 | { _⟨$⟩_ = Function.const b 43 | ; cong = Function.const (Setoid.refl B) 44 | } 45 | 46 | →-to-⟶ : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → (set→setoid A ⟶ set→setoid B) 47 | →-to-⟶ f = record { _⟨$⟩_ = f; cong = ≡-cong f } 48 | 49 | ------------------------------------------------------------------------ 50 | -- Function setoids 51 | 52 | setoid : ∀ {cf ℓf ct ℓt} 53 | (From : Setoid cf ℓf) → 54 | Setoid ct ℓt → 55 | Setoid _ _ 56 | setoid From To = record 57 | { Carrier = From ⟶ To 58 | ; _≈_ = λ f g → ∀ {x y} → x ≈₁ y → f ⟨$⟩ x ≈₂ g ⟨$⟩ y 59 | ; isEquivalence = record 60 | { refl = λ {f} → cong f 61 | ; sym = λ f∼g x∼y → To.sym (f∼g (From.sym x∼y)) 62 | ; trans = λ f∼g g∼h x∼y → To.trans (f∼g From.refl) (g∼h x∼y) 63 | } 64 | } 65 | where 66 | open module From = Setoid From using () renaming (_≈_ to _≈₁_) 67 | open module To = Setoid To using () renaming (_≈_ to _≈₂_) 68 | 69 | infixr 0 _⇨_ 70 | 71 | _⇨_ : ∀ {cf ℓf ct ℓt} → Setoid cf ℓf → Setoid ct ℓt → Setoid _ _ 72 | From ⇨ To = setoid From To 73 | 74 | ⟪_,_⟫ : ∀ {cf ℓf ct₁ ℓt₁ ct₂ ℓt₂} {From : Setoid cf ℓf} {To₁ : Setoid ct₁ ℓt₁} {To₂ : Setoid ct₂ ℓt₂} (f₁ : From ⟶ To₁) (f₂ : From ⟶ To₂) → (From ⟶ (To₁ ×-setoid To₂)) 75 | ⟪ f₁ , f₂ ⟫ = record { _⟨$⟩_ = < (_⟨$⟩_ f₁) , (_⟨$⟩_ f₂) > 76 | ; cong = < cong f₁ , cong f₂ > } 77 | 78 | ⟪_,_⟫′ : ∀ {cf ℓf ct₁ ct₂} {From : Setoid cf ℓf} {To₁ : Set ct₁} {To₂ : Set ct₂} (f₁ : From ⟶ set→setoid To₁) (f₂ : From ⟶ set→setoid To₂) → (From ⟶ set→setoid (To₁ × To₂)) 79 | ⟪ f₁ , f₂ ⟫′ = record { _⟨$⟩_ = < (_⟨$⟩_ f₁) , (_⟨$⟩_ f₂) > 80 | ; cong = λ x≈y → ≡-cong₂ _,_ (cong f₁ x≈y) (cong f₂ x≈y) } -------------------------------------------------------------------------------- /Categories/Support/StarEquality.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Categories.Support.StarEquality where 3 | 4 | open import Categories.Support.Equivalence 5 | open import Data.Star 6 | import Data.Star.Properties as Props 7 | open import Level 8 | open import Relation.Binary 9 | using ( Rel 10 | ; Reflexive; Symmetric; Transitive 11 | ; IsEquivalence 12 | ; _=[_]⇒_) 13 | open import Relation.Binary.PropositionalEquality 14 | using (_≡_) 15 | renaming (refl to ≡-refl) 16 | 17 | module StarEquality {o ℓ e} {Obj : Set o} (S : Obj → Obj → Setoid ℓ e) where 18 | private 19 | open module S i j = Setoid (S i j) 20 | using () renaming (Carrier to T) 21 | _≊_ : ∀ {i j} → Rel (T i j) e 22 | _≊_ {i}{j} = S._≈_ i j 23 | 24 | infix 4 _≈_ 25 | 26 | data _≈_ : {i₁ i₂ : Obj} → Rel (Star T i₁ i₂) (o ⊔ ℓ ⊔ e) where 27 | ε-cong : ∀ {i} → ε {x = i} ≈ ε {x = i} 28 | _◅-cong_ : {i j k : Obj}{x₁ x₂ : T i j} {xs₁ xs₂ : Star T j k} 29 | → x₁ ≊ x₂ 30 | → xs₁ ≈ xs₂ 31 | → (x₁ ◅ xs₁) ≈ (x₂ ◅ xs₂) 32 | 33 | _◅◅-cong_ : {i j k : Obj}{xs₁ xs₂ : Star T i j} {ys₁ ys₂ : Star T j k} 34 | → xs₁ ≈ xs₂ 35 | → ys₁ ≈ ys₂ 36 | → (xs₁ ◅◅ ys₁) ≈ (xs₂ ◅◅ ys₂) 37 | ε-cong ◅◅-cong p = p 38 | (p ◅-cong ps₁) ◅◅-cong ps₂ = p ◅-cong (ps₁ ◅◅-cong ps₂) 39 | 40 | _▻▻-cong_ : {i j k : Obj}{xs₁ xs₂ : Star T j k} {ys₁ ys₂ : Star T i j} 41 | → xs₁ ≈ xs₂ 42 | → ys₁ ≈ ys₂ 43 | → (xs₁ ▻▻ ys₁) ≈ (xs₂ ▻▻ ys₂) 44 | x ▻▻-cong y = y ◅◅-cong x 45 | 46 | private 47 | .refl : ∀ {i j} → Reflexive (_≈_ {i}{j}) 48 | refl {x = ε} = ε-cong 49 | refl {x = x ◅ xs} = S.refl _ _ ◅-cong refl 50 | 51 | .sym : ∀ {i j} → Symmetric (_≈_ {i}{j}) 52 | sym ε-cong = ε-cong 53 | sym (px ◅-cong pxs) = S.sym _ _ px ◅-cong sym pxs 54 | 55 | .trans : ∀ {i j} → Transitive (_≈_ {i}{j}) 56 | trans ε-cong ε-cong = ε-cong 57 | trans (px₁ ◅-cong pxs₁) (px₂ ◅-cong pxs₂) 58 | = S.trans _ _ px₁ px₂ ◅-cong trans pxs₁ pxs₂ 59 | 60 | .isEquivalence : ∀ {i j} → IsEquivalence (_≈_ {i}{j}) 61 | isEquivalence = record 62 | { refl = refl 63 | ; sym = sym 64 | ; trans = trans 65 | } 66 | 67 | private 68 | .reflexive : ∀ {i j} {x y : Star T i j} → x ≡ y → x ≈ y 69 | reflexive ≡-refl = refl 70 | 71 | StarSetoid : ∀ i j → Setoid (o ⊔ ℓ) (o ⊔ ℓ ⊔ e) 72 | StarSetoid i j = record 73 | { Carrier = Star T i j 74 | ; _≈_ = _≈_ 75 | ; isEquivalence = isEquivalence 76 | } 77 | 78 | .◅◅-assoc : {A B C D : Obj} (f : Star T A B) (g : Star T B C) (h : Star T C D) 79 | → ((f ◅◅ g) ◅◅ h) ≈ (f ◅◅ (g ◅◅ h)) 80 | ◅◅-assoc f g h = reflexive (Props.◅◅-assoc f g h) 81 | 82 | .▻▻-assoc : {A B C D : Obj} (f : Star T A B) (g : Star T B C) (h : Star T C D) 83 | → ((h ▻▻ g) ▻▻ f) ≈ (h ▻▻ (g ▻▻ f)) 84 | ▻▻-assoc f g h = sym (◅◅-assoc f g h) 85 | 86 | open StarEquality public using (StarSetoid) 87 | 88 | -- congruences involving Star lists of 2 relations 89 | module StarCong₂ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} 90 | {I : Set o₁} (T-setoid : I → I → Setoid ℓ₁ e₁) 91 | {J : Set o₂} (U-setoid : J → J → Setoid ℓ₂ e₂) 92 | 93 | where 94 | private 95 | module T i j = Setoid (T-setoid i j) 96 | T : Rel I ℓ₁ 97 | T = T.Carrier 98 | _≊₁_ : ∀ {i j} → Rel (T i j) e₁ 99 | _≊₁_ {i}{j} = T._≈_ i j 100 | module T* i j = Setoid (StarSetoid T-setoid i j) 101 | _≈₁_ : ∀ {i j} → Rel (Star T i j) (o₁ ⊔ ℓ₁ ⊔ e₁) 102 | _≈₁_ {i}{j} = T*._≈_ i j 103 | open StarEquality T-setoid 104 | using () 105 | renaming (ε-cong to ε-cong₁; _◅-cong_ to _◅-cong₁_) 106 | 107 | module U i j = Setoid (U-setoid i j) 108 | U : Rel J ℓ₂ 109 | U = U.Carrier 110 | _≊₂_ : ∀ {i j} → Rel (U i j) e₂ 111 | _≊₂_ {i}{j} = U._≈_ i j 112 | module U* i j = Setoid (StarSetoid U-setoid i j) 113 | _≈₂_ : ∀ {i j} → Rel (Star U i j) (o₂ ⊔ ℓ₂ ⊔ e₂) 114 | _≈₂_ {i}{j} = U*._≈_ i j 115 | open StarEquality U-setoid 116 | using () 117 | renaming (ε-cong to ε-cong₂; _◅-cong_ to _◅-cong₂_) 118 | 119 | gmap-cong : (f : I → J) (g : T =[ f ]⇒ U) (g′ : T =[ f ]⇒ U) 120 | → (∀ {i j} (x y : T i j) → x ≊₁ y → g x ≊₂ g′ y) 121 | → ∀ {i j} (xs ys : Star T i j) 122 | → xs ≈₁ ys 123 | → gmap {U = U} f g xs ≈₂ gmap f g′ ys 124 | gmap-cong f g g′ eq ε .ε ε-cong₁ = ε-cong₂ 125 | gmap-cong f g g′ eq (x ◅ xs) (y ◅ ys) (x≊y ◅-cong₁ xs≈ys) 126 | = (eq x y x≊y) ◅-cong₂ (gmap-cong f g g′ eq xs ys xs≈ys) 127 | gmap-cong f g g′ eq (x ◅ xs) ε () 128 | -------------------------------------------------------------------------------- /Categories/Terminal.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Level 4 | 5 | module Categories.Terminal {o ℓ e : Level} where 6 | 7 | open import Categories.Category 8 | open import Categories.Functor 9 | open import Categories.Categories 10 | import Categories.Object.Terminal as Terminal 11 | 12 | open Terminal (Categories o ℓ e) 13 | 14 | record Unit {x : _} : Set x where 15 | constructor unit 16 | 17 | OneC : Category o ℓ e 18 | OneC = record 19 | { Obj = Unit 20 | ; _⇒_ = λ _ _ → Unit 21 | ; _≡_ = λ _ _ → Unit 22 | ; _∘_ = λ _ _ → unit 23 | ; id = unit 24 | ; assoc = unit 25 | ; identityˡ = unit 26 | ; identityʳ = unit 27 | ; equiv = record 28 | { refl = unit 29 | ; sym = λ _ → unit 30 | ; trans = λ _ _ → unit 31 | } 32 | ; ∘-resp-≡ = λ _ _ → unit 33 | } 34 | 35 | -- I can probably use Discrete here once we get universe cumulativity in Agda 36 | One : Terminal 37 | One = record 38 | { ⊤ = OneC 39 | ; ! = record 40 | { F₀ = λ _ → unit 41 | ; F₁ = λ _ → unit 42 | ; identity = unit 43 | ; homomorphism = unit 44 | ; F-resp-≡ = λ _ → unit 45 | } 46 | ; !-unique = λ _ _ → Heterogeneous.≡⇒∼ unit 47 | } 48 | -------------------------------------------------------------------------------- /Categories/Topos.agda: -------------------------------------------------------------------------------- 1 | module Categories.Topos where 2 | -------------------------------------------------------------------------------- /Categories/WithFamilies.agda: -------------------------------------------------------------------------------- 1 | module Categories.WithFamilies where 2 | 3 | open import Level 4 | import Relation.Binary.HeterogeneousEquality as Het 5 | open Het using (_≅_) 6 | 7 | open import Categories.Support.PropositionalEquality 8 | open import Categories.Support.Experimental 9 | 10 | open import Categories.Category 11 | open import Categories.Functor 12 | open import Categories.NaturalIsomorphism 13 | open import Categories.Object.Terminal 14 | 15 | open import Categories.Fam 16 | 17 | module UnpackFam {o ℓ e a b} (C : Category o ℓ e) 18 | (T : Functor (Category.op C) (Fam a b)) where 19 | private module C = Category C 20 | private module T = Functor T 21 | 22 | Context = C.Obj 23 | 24 | Ty : C.Obj → Set a 25 | Ty Γ = Fam.U (T.F₀ Γ) 26 | 27 | _[_] : ∀ {Γ Δ} → Ty Γ → Δ C.⇒ Γ → Ty Δ 28 | _[_] A f = Fam.Hom.f (T.F₁ f) A 29 | 30 | Tm : ∀ Γ → Ty Γ → Set b 31 | Tm Γ = Fam.T (T.F₀ Γ) 32 | 33 | _[_⁺] : ∀ {Γ Δ A} → Tm Γ A → (f : Δ C.⇒ Γ) → Tm Δ (A [ f ]) 34 | _[_⁺] M f = Fam.Hom.φ (T.F₁ f) _ M 35 | 36 | record CwF {o ℓ e a b} : Set (suc e ⊔ (suc ℓ ⊔ suc o) ⊔ suc a ⊔ suc b) where 37 | field 38 | C : Category o ℓ e 39 | T : Functor (Category.op C) (Fam a b) 40 | Empty : Terminal C 41 | 42 | module C = Category C 43 | module T = Functor T 44 | open UnpackFam C T 45 | module Empty = Terminal C Empty 46 | 47 | field 48 | -- context snoc 49 | _>_ : ∀ Γ → Ty Γ → C.Obj 50 | -- projections 51 | p : ∀ {Γ A} → (Γ > A) C.⇒ Γ 52 | v : ∀ {Γ A} → Tm (Γ > A) (A [ p ]) 53 | -- constructor 54 | <_,_> : ∀ {Γ A} → ∀ {Δ} (γ : Δ C.⇒ Γ) (a : Tm Δ (A [ γ ])) → Δ C.⇒ (Γ > A) 55 | .p∘<γ,a>≡γ : ∀ {Γ A} → ∀ {Δ} {γ : Δ C.⇒ Γ} {a : Tm Δ (A [ γ ])} → p C.∘ < γ , a > C.≡ γ 56 | .v[<γ,a>]≡a : ∀ {Γ A} → ∀ {Δ} {γ : Δ C.⇒ Γ} {a : Tm Δ (A [ γ ])} → v [ < γ , a > ⁺] ≅ a 57 | .<γ,a>-unique : ∀ {Γ A} → ∀ {Δ} {γ : Δ C.⇒ Γ} {a : Tm Δ (A [ γ ])} → 58 | (δ : Δ C.⇒ (Γ > A)) → .(p C.∘ δ C.≡ γ) → .(v [ δ ⁺] ≅ a) → δ C.≡ < γ , a > 59 | 60 | v[_] : ∀ {Γ A Δ} → (γ : Δ C.⇒ Γ) -> Tm (Δ > A [ γ ]) (A [ γ C.∘ p ]) 61 | v[_] {Γ} {A} {Δ} γ = ≣-subst′ (Tm (Δ > A [ γ ])) (≣-sym (Fam.Eq.f≡g (T.homomorphism {Γ}) {A})) (v {Δ} {A [ γ ]}) 62 | 63 | _[id] : ∀ {Γ A} -> Tm Γ A -> Tm Γ (A [ C.id ]) 64 | _[id] {Γ} {A} x = ≣-subst′ (Tm Γ) (≣-sym (Fam.Eq.f≡g (T.identity {Γ}) {A})) x 65 | 66 | open UnpackFam C T public 67 | open Empty public using () renaming (⊤ to <>) 68 | 69 | 70 | 71 | record Pi {o ℓ e a b} (Cwf : CwF {o} {ℓ} {e} {a} {b}) : Set (o ⊔ ℓ ⊔ a ⊔ b) where 72 | open CwF Cwf 73 | field 74 | Π : ∀ {Γ} -> (A : Ty Γ) (B : Ty (Γ > A)) -> Ty Γ 75 | lam : ∀ {Γ} {A : Ty Γ} {B : Ty (Γ > A)} -> (b : Tm (Γ > A) B) -> Tm Γ (Π A B) 76 | _$_ : ∀ {Γ} {A : Ty Γ} {B : Ty (Γ > A)} -> 77 | (f : Tm Γ (Π A B)) (x : Tm Γ A) -> Tm Γ (B [ < C.id , x [id] > ]) 78 | 79 | -- naturality laws 80 | .Π-nat : ∀ {Γ} -> (A : Ty Γ) (B : Ty (Γ > A)) -> ∀ {Δ} (γ : Δ C.⇒ Γ) 81 | -> Π A B [ γ ] ≣ Π (A [ γ ]) (B [ < (γ C.∘ p) , v[ γ ] > ]) 82 | .lam-nat : ∀ {Γ} {A : Ty Γ} {B : Ty (Γ > A)} -> (b : Tm (Γ > A) B) -> ∀ {Δ} (γ : Δ C.⇒ Γ) 83 | -> lam b [ γ ⁺] ≅ lam {A = A [ γ ]} (b [ < γ C.∘ p , v[ γ ] > ⁺]) 84 | .app-nat : ∀ {Γ} {A : Ty Γ} {B : Ty (Γ > A)} -> (f : Tm Γ (Π A B)) (x : Tm Γ A) -> ∀ {Δ} (γ : Δ C.⇒ Γ) 85 | -> (f $ x) [ γ ⁺] ≅ ≣-subst′ (Tm Δ) (Π-nat A B γ) (f [ γ ⁺]) $ (x [ γ ⁺]) 86 | 87 | -- proofs of the lam/_$_ isomorphism 88 | .β : ∀ {Γ} {A : Ty Γ} {B : Ty (Γ > A)} -> (b : Tm (Γ > A) B) (x : Tm Γ A) 89 | -> (lam b $ x) ≣ b [ < C.id , x [id] > ⁺] 90 | 91 | .η : ∀ {Γ} {A : Ty Γ} {B : Ty (Γ > A)} -> (f : Tm Γ (Π A B)) 92 | -> lam (≣-subst′ (Tm (Γ > A)) (Π-nat A B p) (f [ p ⁺]) $ v) ≅ f 93 | 94 | 95 | -------------------------------------------------------------------------------- /Graphs/Graph.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Graphs.Graph where 3 | 4 | open import Level 5 | open import Categories.Support.Equivalence 6 | open import Relation.Binary 7 | using (Rel; IsEquivalence; module IsEquivalence; Reflexive; Symmetric; Transitive) 8 | renaming (_⇒_ to _⊆_) 9 | open import Relation.Binary.PropositionalEquality 10 | using () 11 | renaming (_≡_ to _≣_) 12 | 13 | record Graph (o ℓ e : Level) : Set (suc (o ⊔ ℓ ⊔ e)) where 14 | field 15 | Obj : Set o 16 | _↝_ : Obj → Obj → Set ℓ 17 | _≈_ : {A B : Obj} → Rel (A ↝ B) e 18 | .equiv : ∀ {A B : Obj} → IsEquivalence (_≈_ {A}{B}) 19 | 20 | edge-setoid : ∀ {A B : Obj} → Setoid ℓ e 21 | edge-setoid {A}{B} = record 22 | { Carrier = A ↝ B 23 | ; _≈_ = _≈_ 24 | ; isEquivalence = equiv 25 | } 26 | 27 | module Equiv {A B : Obj} where 28 | private 29 | module e = IsEquivalence 30 | .q : IsEquivalence _≈_ 31 | q = equiv {A} {B} 32 | 33 | .refl : Reflexive _≈_ 34 | refl = e.refl q 35 | .trans : Transitive _≈_ 36 | trans = e.trans q 37 | .sym : Symmetric _≈_ 38 | sym = e.sym q 39 | .reflexive : _≣_ ⊆ _≈_ 40 | reflexive = e.reflexive q 41 | 42 | _[_↝_] : ∀ {o ℓ e} → (G : Graph o ℓ e) → Graph.Obj G → Graph.Obj G → Set ℓ 43 | G [ A ↝ B ] = Graph._↝_ G A B 44 | 45 | _[_≈_] : ∀ {o ℓ e} → (G : Graph o ℓ e) → {A B : Graph.Obj G} → Rel (G [ A ↝ B ]) e 46 | G [ f ≈ g ] = Graph._≈_ G f g 47 | 48 | module Heterogeneous {o ℓ e} (G : Graph o ℓ e) where 49 | open Graph G 50 | open Equiv renaming (refl to refl′; sym to sym′; trans to trans′) 51 | 52 | data _~_ {A B} (f : A ↝ B) : ∀ {X Y} → (X ↝ Y) → Set (ℓ ⊔ e) where 53 | ≈⇒~ : {g : A ↝ B} → .(f ≈ g) → f ~ g 54 | 55 | refl : ∀ {A B} {f : A ↝ B} → f ~ f 56 | refl = ≈⇒~ refl′ 57 | 58 | sym : ∀ {A B} {f : A ↝ B} {D E} {g : D ↝ E} → f ~ g → g ~ f 59 | sym (≈⇒~ f≈g) = ≈⇒~ (sym′ f≈g) 60 | 61 | trans : ∀ {A B} {f : A ↝ B} 62 | {D E} {g : D ↝ E} 63 | {F G} {h : F ↝ G} 64 | → f ~ g → g ~ h → f ~ h 65 | trans (≈⇒~ f≈g) (≈⇒~ g≈h) = ≈⇒~ (trans′ f≈g g≈h) 66 | 67 | _[_~_] : ∀ {o ℓ e} (G : Graph o ℓ e) {A B} (f : G [ A ↝ B ]) {X Y} (g : G [ X ↝ Y ]) → Set (ℓ ⊔ e) 68 | G [ f ~ g ] = Heterogeneous._~_ G f g -------------------------------------------------------------------------------- /Graphs/GraphMorphism.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | module Graphs.GraphMorphism where 3 | 4 | open import Data.Product 5 | open import Level 6 | open import Relation.Binary 7 | renaming (_⇒_ to _⊆_) 8 | open import Relation.Binary.PropositionalEquality 9 | using () 10 | renaming 11 | (_≡_ to _≣_ 12 | ; refl to ≣-refl 13 | ; sym to ≣-sym 14 | ; trans to ≣-trans 15 | ; cong to ≣-cong 16 | ) 17 | 18 | open import Graphs.Graph 19 | 20 | record GraphMorphism {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} 21 | (A : Graph o₁ ℓ₁ e₁) (B : Graph o₂ ℓ₂ e₂) : Set (o₁ ⊔ o₂ ⊔ ℓ₁ ⊔ ℓ₂ ⊔ e₁ ⊔ e₂) where 22 | module A = Graph A 23 | module B = Graph B 24 | field 25 | F₀ : A.Obj → B.Obj 26 | F₁ : ∀ {X Y} → A [ X ↝ Y ] → B [ F₀ X ↝ F₀ Y ] 27 | .F-resp-≈ : ∀ {X Y}{f g : A [ X ↝ Y ]} → A [ f ≈ g ] → B [ F₁ f ≈ F₁ g ] 28 | 29 | infixr 9 _∘_ 30 | infix 4 _≈_ 31 | 32 | id : ∀{o ℓ e}{A : Graph o ℓ e} → GraphMorphism A A 33 | id = record 34 | { F₀ = λ A → A 35 | ; F₁ = λ f → f 36 | ; F-resp-≈ = λ f≈g → f≈g 37 | } 38 | 39 | _∘_ : ∀ {o₀ ℓ₀ e₀ o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} 40 | {A : Graph o₀ ℓ₀ e₀} {B : Graph o₁ ℓ₁ e₁} {C : Graph o₂ ℓ₂ e₂} 41 | → GraphMorphism B C → GraphMorphism A B → GraphMorphism A C 42 | G ∘ F = record 43 | { F₀ = λ A → G₀ (F₀ A) 44 | ; F₁ = λ f → G₁ (F₁ f) 45 | ; F-resp-≈ = λ f≈g → G-resp-≈ (F-resp-≈ f≈g) 46 | } where 47 | open GraphMorphism F 48 | open GraphMorphism G renaming (F₀ to G₀; F₁ to G₁; F-resp-≈ to G-resp-≈) 49 | 50 | _≈₀_ : ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} {A : Graph o₁ ℓ₁ e₁} {B : Graph o₂ ℓ₂ e₂} 51 | → (F G : GraphMorphism A B) 52 | → Set (o₁ ⊔ o₂) 53 | _≈₀_ {A = A}{B} F G 54 | = ∀ X → F₀ X ≣ G₀ X 55 | where 56 | open GraphMorphism F using (F₀) 57 | open GraphMorphism G renaming (F₀ to G₀) 58 | 59 | isEquivalence₀ : ∀{o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} {A : Graph o₁ ℓ₁ e₁} {B : Graph o₂ ℓ₂ e₂} 60 | → IsEquivalence (_≈₀_ {A = A}{B}) 61 | isEquivalence₀ {A = A}{B} = record 62 | { refl = λ x → ≣-refl 63 | ; sym = λ p x → ≣-sym (p x) 64 | ; trans = λ p q x → ≣-trans (p x) (q x) 65 | } 66 | 67 | _≈₁_ : ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} {A : Graph o₁ ℓ₁ e₁} {B : Graph o₂ ℓ₂ e₂} 68 | → (F G : GraphMorphism A B) 69 | → Set (o₁ ⊔ ℓ₁ ⊔ ℓ₂ ⊔ e₂) 70 | _≈₁_ {A = A}{B} F G 71 | = ∀ {X Y} (f : A [ X ↝ Y ]) → B [ F₁ f ~ G₁ f ] 72 | where 73 | open GraphMorphism F using (F₁) 74 | open GraphMorphism G renaming (F₁ to G₁) 75 | 76 | isEquivalence₁ : ∀{o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} {A : Graph o₁ ℓ₁ e₁} {B : Graph o₂ ℓ₂ e₂} 77 | → IsEquivalence (_≈₁_ {A = A}{B}) 78 | isEquivalence₁ {A = A}{B} = record 79 | { refl = λ f → refl 80 | ; sym = λ p f → sym (p f) 81 | ; trans = λ p q f → trans (p f) (q f) 82 | } where open Heterogeneous B 83 | 84 | _≈_ : ∀ {o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} {A : Graph o₁ ℓ₁ e₁} {B : Graph o₂ ℓ₂ e₂} 85 | → (F G : GraphMorphism A B) 86 | → Set (o₁ ⊔ o₂ ⊔ ℓ₁ ⊔ ℓ₂ ⊔ e₂) 87 | _≈_ = _≈₀_ -×- _≈₁_ 88 | 89 | isEquivalence : ∀{o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} {A : Graph o₁ ℓ₁ e₁} {B : Graph o₂ ℓ₂ e₂} 90 | → IsEquivalence (_≈_ {A = A}{B}) 91 | isEquivalence {A = A}{B} = record 92 | { refl = λ {x} → refl₀ {x} , λ f → refl₁ {x} f 93 | ; sym = λ {x}{y} → map (sym₀ {x}{y}) (sym₁ {x}{y}) 94 | ; trans = λ {x}{y}{z} → zip (trans₀ {x}{y}{z}) (trans₁ {x}{y}{z}) 95 | } 96 | where 97 | open IsEquivalence isEquivalence₀ 98 | renaming (refl to refl₀; sym to sym₀; trans to trans₀) 99 | open IsEquivalence isEquivalence₁ 100 | renaming (refl to refl₁; sym to sym₁; trans to trans₁) 101 | 102 | .∘-resp-≈ : ∀ {o₀ ℓ₀ e₀ o₁ ℓ₁ e₁ o₂ ℓ₂ e₂} 103 | {A : Graph o₀ ℓ₀ e₀} {B : Graph o₁ ℓ₁ e₁} {C : Graph o₂ ℓ₂ e₂} 104 | {F H : GraphMorphism B C} {G I : GraphMorphism A B} 105 | → F ≈ H → G ≈ I → F ∘ G ≈ H ∘ I 106 | ∘-resp-≈ {B = B} {C} {F}{H}{G}{I} F≈H G≈I 107 | = (λ x → helper₁ (proj₁ G≈I x) (proj₁ F≈H (GraphMorphism.F₀ I x))) 108 | , (λ q → helper₂ (proj₂ G≈I q) (proj₂ F≈H (GraphMorphism.F₁ I q))) 109 | where 110 | open Heterogeneous C 111 | module C = Graph C 112 | helper₁ : ∀ {x} 113 | → GraphMorphism.F₀ G x ≣ GraphMorphism.F₀ I x 114 | → GraphMorphism.F₀ F (GraphMorphism.F₀ I x) ≣ GraphMorphism.F₀ H (GraphMorphism.F₀ I x) 115 | → GraphMorphism.F₀ F (GraphMorphism.F₀ G x) ≣ GraphMorphism.F₀ H (GraphMorphism.F₀ I x) 116 | helper₁ g≣i f≣h = ≣-trans (≣-cong (GraphMorphism.F₀ F) g≣i) f≣h 117 | 118 | helper₂ : ∀ {a b c d} {z w} {f : B [ a ↝ b ]} {h : B [ c ↝ d ]} {i : C [ z ↝ w ]} 119 | → B [ f ~ h ] → C [ GraphMorphism.F₁ F h ~ i ] → C [ GraphMorphism.F₁ F f ~ i ] 120 | helper₂ (≈⇒~ f≈h) (≈⇒~ g≈i) = ≈⇒~ (C.Equiv.trans (GraphMorphism.F-resp-≈ F f≈h) g≈i) 121 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Copyright (c) 2014, Dan Peebles, James Deikun, Andrea Vezzosi, James 2 | Cook, and other contributors. 3 | 4 | All rights reserved. 5 | 6 | Redistribution and use in source and binary forms, with or without 7 | modification, are permitted provided that the following conditions are met: 8 | 9 | * Redistributions of source code must retain the above copyright 10 | notice, this list of conditions and the following disclaimer. 11 | 12 | * Redistributions in binary form must reproduce the above 13 | copyright notice, this list of conditions and the following 14 | disclaimer in the documentation and/or other materials provided 15 | with the distribution. 16 | 17 | * Neither the name of Dan Peebles nor the names of other 18 | contributors may be used to endorse or promote products derived 19 | from this software without specific prior written permission. 20 | 21 | THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS 22 | "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT 23 | LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR 24 | A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT 25 | OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, 26 | SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT 27 | LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, 28 | DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY 29 | THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT 30 | (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE 31 | OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. 32 | -------------------------------------------------------------------------------- /README: -------------------------------------------------------------------------------- 1 | This library is **obsolete** and does not work with newer versions of Agda. It sort-of works with 2.5.2, but less and less so afterwards. It is completely broken from 2.6 onwards. 2 | 3 | Users should instead try https://github.com/agda/agda-categories. It is not completely compatible, but close. And it is supported. 4 | 5 | ---------------------------------------------------------------------------------------------------------------------------------- 6 | 7 | One of the main goals of this library is to be as general as possible by abstracting over notions of equality between morphisms. 8 | 9 | Another is to keep the definitions of categorical structures as simple as possible and then where possible to prove that the simple definition is equivalent to more interesting formulations. 10 | For example, we can define a monad as recording containing a functor and two natural transformations. Separately, we can show that a monoid object in the monoidal category of endofunctors is an equivalent definition, and that the composition of an adjoint pair of functors leads to monads, and so on. 11 | 12 | 13 | The module structure is a mess, I realize. A lot of the parametrized modules should not be. Naming of things in general could also be made cleaner, but I've been more interested in definitions and proofs so far. 14 | 15 | 16 | A lot of this is based on http://web.student.chalmers.se/~stevan/ctfp/html/README.html, but with some design changes that I thought were necessary. It's still very much a work in progress. 17 | 18 | Other parts (mostly produts) are borrowed from Dan Doel 19 | -------------------------------------------------------------------------------- /categories.agda-lib: -------------------------------------------------------------------------------- 1 | name: categories 2 | depend: standard-library 3 | include: . 4 | --------------------------------------------------------------------------------