Update Fun according to new naming policy

This commit is contained in:
Frederik Hanghøj Iversen 2018-02-21 13:40:24 +01:00
parent 57d7eab4cb
commit 9e96e704e8
2 changed files with 6 additions and 7 deletions

View file

@ -21,7 +21,7 @@ open import Cat.Equality
open Equality.Data.Product
module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Category d d'} where
open Category hiding ( _∘_ ; Arrow )
open Category using (Object ; 𝟙)
open Functor
module _ (F G : Functor 𝔻) where
@ -70,7 +70,7 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
where
module F = Functor F
F→ = F.func→
module 𝔻 = IsCategory (isCategory 𝔻)
module 𝔻 = Category 𝔻
identityNat : (F : Functor 𝔻) NaturalTransformation F F
identityNat F = identityTrans F , identityNatural F
@ -95,13 +95,13 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
𝔻 [ H.func→ f 𝔻 [ θ A η A ] ] ≡⟨⟩
𝔻 [ H.func→ f (θ ∘nt η) A ]
where
open IsCategory (isCategory 𝔻)
open Category 𝔻
NatComp = _:⊕:_
private
module _ {F G : Functor 𝔻} where
module 𝔻 = IsCategory (isCategory 𝔻)
module 𝔻 = Category 𝔻
transformationIsSet : isSet (Transformation F G)
transformationIsSet _ _ p q i j C = 𝔻.arrowIsSet _ _ (λ l p l C) (λ l q l C) i j
@ -139,7 +139,7 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
:assoc: = lemSig (naturalIsProp {F = A} {D})
L R (funExt (λ x assoc))
where
open IsCategory (isCategory 𝔻)
open Category 𝔻
module _ {A B : Functor 𝔻} {f : NaturalTransformation A B} where
private
@ -181,7 +181,7 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
}
Fun : Category (c c' d d') (c c' d')
raw Fun = RawFun
Category.raw Fun = RawFun
module _ { ' : Level} ( : Category ') where
open import Cat.Categories.Sets

View file

@ -14,7 +14,6 @@ open Equality.Data.Product
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