Move properties of categories to Cat.Category.Properties

This commit is contained in:
Frederik Hanghøj Iversen 2018-01-21 15:01:01 +01:00
parent 316de7e4f9
commit 793fc30534
2 changed files with 63 additions and 71 deletions

View file

@ -53,68 +53,26 @@ record Category ( ' : Level) : Set (lsuc (' ⊔ )) where
open Category
module _ { ' : Level} { : Category '} { A B : .Object } where
private
open module = Category
_+_ = ._⊕_
module _ { ' : Level} { : Category '} where
module _ { A B : .Object } where
Isomorphism : (f : .Arrow A B) Set '
Isomorphism f = Σ[ g .Arrow B A ] ._⊕_ g f .𝟙 × ._⊕_ f g .𝟙
Isomorphism : (f : .Arrow A B) Set '
Isomorphism f = Σ[ g .Arrow B A ] g .⊕ f .𝟙 × f + g .𝟙
Epimorphism : {X : .Object } (f : .Arrow A B) Set '
Epimorphism {X} f = ( g₀ g₁ : .Arrow B X ) ._⊕_ g₀ f ._⊕_ g₁ f g₀ g₁
Epimorphism : {X : .Object } (f : .Arrow A B) Set '
Epimorphism {X} f = ( g₀ g₁ : .Arrow B X ) g₀ + f g₁ + f g₀ g₁
Monomorphism : {X : .Object} (f : .Arrow A B) Set '
Monomorphism {X} f = ( g₀ g₁ : .Arrow X A ) ._⊕_ f g₀ ._⊕_ f g₁ g₀ g₁
Monomorphism : {X : .Object} (f : .Arrow A B) Set '
Monomorphism {X} f = ( g₀ g₁ : .Arrow X A ) f + g₀ f + g₁ g₀ g₁
-- Isomorphism of objects
_≅_ : (A B : Object ) Set '
_≅_ A B = Σ[ f .Arrow A B ] (Isomorphism f)
iso-is-epi : {X} (f : .Arrow A B) Isomorphism f Epimorphism {X = X} f
iso-is-epi f (f- , left-inv , right-inv) g₀ g₁ eq =
begin
g₀ ≡⟨ sym (fst ident)
g₀ + .𝟙 ≡⟨ cong (_+_ g₀) (sym right-inv)
g₀ + (f + f-) ≡⟨ assoc
(g₀ + f) + f- ≡⟨ cong (λ x x + f-) eq
(g₁ + f) + f- ≡⟨ sym assoc
g₁ + (f + f-) ≡⟨ cong (_+_ g₁) right-inv
g₁ + .𝟙 ≡⟨ fst ident
g₁
where
open IsCategory .isCategory
iso-is-mono : {X} (f : .Arrow A B ) Isomorphism f Monomorphism {X = X} f
iso-is-mono f (f- , (left-inv , right-inv)) g₀ g₁ eq =
begin
g₀ ≡⟨ sym (snd ident)
.𝟙 + g₀ ≡⟨ cong (λ x x + g₀) (sym left-inv)
(f- + f) + g₀ ≡⟨ sym assoc
f- + (f + g₀) ≡⟨ cong (_+_ f-) eq
f- + (f + g₁) ≡⟨ assoc
(f- + f) + g₁ ≡⟨ cong (λ x x + g₁) left-inv
.𝟙 + g₁ ≡⟨ snd ident
g₁
where
open IsCategory .isCategory
iso-is-epi-mono : {X} (f : .Arrow A B ) Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f
iso-is-epi-mono f iso = iso-is-epi f iso , iso-is-mono f iso
{-
epi-mono-is-not-iso : { '} ¬ (( : Category {} {'}) {A B X : Object } (f : Arrow A B ) Epimorphism { = } {X = X} f Monomorphism { = } {X = X} f Isomorphism { = } f)
epi-mono-is-not-iso f =
let k = f {!!} {!!} {!!} {!!}
in {!!}
-}
-- Isomorphism of objects
_≅_ : { '} { : Category '} (A B : Object ) Set '
_≅_ { = } A B = Σ[ f .Arrow A B ] (Isomorphism { = } f)
IsProduct : { '} ( : Category ') {A B obj : Object } (π₁ : Arrow obj A) (π₂ : Arrow obj B) Set ( ')
IsProduct {A = A} {B = B} π₁ π₂
= {X : .Object} (x₁ : .Arrow X A) (x₂ : .Arrow X B)
∃![ x ] (π₁ .⊕ x x₁ × π₂ .⊕ x x₂)
where
open module = Category
module _ { ' : Level} ( : Category ') {A B obj : Object } where
IsProduct : (π₁ : Arrow obj A) (π₂ : Arrow obj B) Set ( ')
IsProduct π₁ π₂
= {X : .Object} (x₁ : .Arrow X A) (x₂ : .Arrow X B)
∃![ x ] ( ._⊕_ π₁ x x₁ × ._⊕_ π₂ x x₂)
-- Tip from Andrea; Consider this style for efficiency:
-- record IsProduct { ' : Level} ( : Category {} {'})
@ -131,19 +89,7 @@ record Product { ' : Level} { : Category '} (A B : .Object)
proj₂ : .Arrow obj B
{{isProduct}} : IsProduct proj₁ proj₂
-- Two pairs are equal if their components are equal.
eqpair : {a b} {A : Set a} {B : Set b} {a a' : A} {b b' : B}
a a' b b' (a , b) (a' , b')
eqpair eqa eqb i = eqa i , eqb i
module _ { ' : Level} ( : Category ') where
private
instance
_ : IsCategory ( .Object) (flip ( .Arrow)) ( .𝟙) (flip ( ._⊕_))
_ = record { assoc = sym assoc ; ident = swap ident }
where
open IsCategory ( .isCategory)
Opposite : Category '
Opposite =
record
@ -151,7 +97,10 @@ module _ { ' : Level} ( : Category ') where
; Arrow = flip ( .Arrow)
; 𝟙 = .𝟙
; _⊕_ = flip ( ._⊕_)
; isCategory = record { assoc = sym assoc ; ident = swap ident }
}
where
open IsCategory ( .isCategory)
-- A consequence of no-eta-equality; `Opposite-is-involution` is no longer
-- definitional - i.e.; you must match on the fields:

View file

@ -2,10 +2,53 @@
module Cat.Category.Properties where
open import Agda.Primitive
open import Data.Product
open import Cubical.PathPrelude
open import Cat.Category
open import Cat.Functor
open import Cat.Categories.Sets
module _ { ' : Level} { : Category '} { A B : .Category.Object } {X : .Category.Object} (f : .Category.Arrow A B) where
open Category
open IsCategory (isCategory)
iso-is-epi : Isomorphism { = } f Epimorphism { = } {X = X} f
iso-is-epi (f- , left-inv , right-inv) g₀ g₁ eq =
begin
g₀ ≡⟨ sym (proj₁ ident)
g₀ 𝟙 ≡⟨ cong (_⊕_ g₀) (sym right-inv)
g₀ (f f-) ≡⟨ assoc
(g₀ f) f- ≡⟨ cong (λ φ φ f-) eq
(g₁ f) f- ≡⟨ sym assoc
g₁ (f f-) ≡⟨ cong (_⊕_ g₁) right-inv
g₁ 𝟙 ≡⟨ proj₁ ident
g₁
iso-is-mono : Isomorphism { = } f Monomorphism { = } {X = X} f
iso-is-mono (f- , (left-inv , right-inv)) g₀ g₁ eq =
begin
g₀ ≡⟨ sym (proj₂ ident)
𝟙 g₀ ≡⟨ cong (λ φ φ g₀) (sym left-inv)
(f- f) g₀ ≡⟨ sym assoc
f- (f g₀) ≡⟨ cong (_⊕_ f-) eq
f- (f g₁) ≡⟨ assoc
(f- f) g₁ ≡⟨ cong (λ φ φ g₁) left-inv
𝟙 g₁ ≡⟨ proj₂ ident
g₁
iso-is-epi-mono : Isomorphism { = } f Epimorphism { = } {X = X} f × Monomorphism { = } {X = X} f
iso-is-epi-mono iso = iso-is-epi iso , iso-is-mono iso
{-
epi-mono-is-not-iso : { '} ¬ (( : Category {} {'}) {A B X : Object } (f : Arrow A B ) Epimorphism { = } {X = X} f Monomorphism { = } {X = X} f Isomorphism { = } f)
epi-mono-is-not-iso f =
let k = f {!!} {!!} {!!} {!!}
in {!!}
-}
module _ {a a' b b'} where
Exponential : Category a a' Category b b' Category {!!} {!!}
Exponential A B = record
@ -13,7 +56,7 @@ module _ {a a' b b'} where
; Arrow = {!!}
; 𝟙 = {!!}
; _⊕_ = {!!}
; isCategory = ?
; isCategory = {!!}
}
_⇑_ = Exponential