view nat.agda @ 4:05087d3aeac0

fix
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sat, 06 Jul 2013 08:50:32 +0900
parents dce706edd66b
children 16572013c362
line wrap: on
line source

module nat  where 

-- Monad
-- Category A
-- A = Category
-- Functor T : A -> A
--T(a) = t(a)
--T(f) = tf(f)

open import Category -- https://github.com/konn/category-agda
open import Level
open Functor

--T(g f) = T(g) T(f)

Lemma1 : {c₁ c₂ l : Level} {A : Category c₁ c₂ l} (T : Functor A A) ->  {a b c : Obj A} {g : Hom A b c} { f : Hom A a b }
     -> A [ ( FMap T (A [ g o f ] ))  ≈ (A [ FMap T g o FMap T f ]) ]
Lemma1 = \t -> IsFunctor.distr ( isFunctor t )

--        F(f)
--   F(a) ----> F(b)
--    |          |
--    |t(a)      |t(b)    G(f)t(a) = t(b)F(f)
--    |          |
--    v          v
--   G(a) ----> G(b)
--        G(f)

record IsNTrans {c₁ c₂ ℓ c₁′ c₂′ ℓ′ : Level} (D : Category c₁ c₂ ℓ) (C : Category c₁′ c₂′ ℓ′)
                 ( F G : Functor D C )
                 (Trans : (A : Obj D) → Hom C (FObj F A) (FObj G A))
                 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ ⊔ c₁′ ⊔ c₂′ ⊔ ℓ′)) where
  field
    naturality : {a b : Obj D} {f : Hom D a b} 
      → C [ C [ (  FMap G f ) o  ( Trans a ) ]  ≈ C [ (Trans b ) o  (FMap F f)  ] ]
--    uniqness : {d : Obj D} 
--      →  C [ Trans d  ≈ Trans d ]


record NTrans {c₁ c₂ ℓ c₁′ c₂′ ℓ′ : Level} (domain : Category c₁ c₂ ℓ) (codomain : Category c₁′ c₂′ ℓ′) (F G : Functor domain codomain )
       : Set (suc (c₁ ⊔ c₂ ⊔ ℓ ⊔ c₁′ ⊔ c₂′ ⊔ ℓ′)) where
  field
    Trans :  (A : Obj domain) → Hom codomain (FObj F A) (FObj G A)
    isNTrans : IsNTrans domain codomain F G Trans

open NTrans
Lemma2 : {c₁ c₂ l : Level} {A : Category c₁ c₂ l} {F G : Functor A A} 
      -> (μ : NTrans A A F G) -> {a b : Obj A} { f : Hom A a b }
      -> A [ A [  FMap G f o Trans μ a ]  ≈ A [ Trans μ b o  FMap F f ] ]
Lemma2 = \n -> IsNTrans.naturality ( isNTrans  n  )

open import Category.Cat

-- η :   1_A -> T
-- μ :   TT -> T
-- μ(a)η(T(a)) = a
-- μ(a)T(η(a)) = a
-- μ(a)(μ(T(a))) = μ(a)T(μ(a))

record IsMonad {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) 
                 ( T : Functor A A )
                 ( η : NTrans A A identityFunctor T )
                 ( μ : NTrans A A (T ○ T) T)
                 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ )) where
   field
      assoc  : {a : Obj A} -> A [ A [ Trans μ a o Trans μ ( FObj T a ) ] ≈  A [  Trans μ a o FMap T (Trans μ a) ] ]
      unity1 : {a : Obj A} -> A [ A [ Trans μ a o Trans η ( FObj T a ) ] ≈ Id {_} {_} {_} {A} (FObj T a) ]
      unity2 : {a : Obj A} -> A [ A [ Trans μ a o (FMap T (Trans η a ))] ≈ Id {_} {_} {_} {A} (FObj T a) ]

record Monad {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (T : Functor A A) (η : NTrans A A identityFunctor T) (μ : NTrans A A (T ○ T) T)
       : Set (suc (c₁ ⊔ c₂ ⊔ ℓ )) where
  field
    isMonad : IsMonad A T η μ

open Monad
Lemma3 : {c₁ c₂ ℓ : Level} {A : Category c₁ c₂ ℓ}
                 { T : Functor A A }
                 { η : NTrans A A identityFunctor T }
                 { μ : NTrans A A (T ○ T) T }
                 { a : Obj A } ->
                 ( M : Monad A T η μ )
    -> A [ A [  Trans μ a o Trans μ ( FObj T a ) ] ≈  A [  Trans μ a o FMap T (Trans μ a) ] ]
Lemma3 = \m -> IsMonad.assoc ( isMonad m )


Lemma4 : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) {a b : Obj A } { f : Hom A a b}
    -> A [ A [ Id {_} {_} {_} {A} b o f ] ≈ f ]
Lemma4 = \a -> IsCategory.identityL ( Category.isCategory a )

Lemma5 : {c₁ c₂ ℓ : Level} {A : Category c₁ c₂ ℓ}
                 { T : Functor A A }
                 { η : NTrans A A identityFunctor T }
                 { μ : NTrans A A (T ○ T) T }
                 { a : Obj A } ->
                 ( M : Monad A T η μ )
    ->  A [ A [ Trans μ a o Trans η ( FObj T a )  ] ≈ Id {_} {_} {_} {A} (FObj T a) ]
Lemma5 = \m -> IsMonad.unity1 ( isMonad m )

Lemma6 : {c₁ c₂ ℓ : Level} {A : Category c₁ c₂ ℓ}
                 { T : Functor A A }
                 { η : NTrans A A identityFunctor T }
                 { μ : NTrans A A (T ○ T) T }
                 { a : Obj A } ->
                 ( M : Monad A T η μ )
    ->  A [ A [ Trans μ a o (FMap T (Trans η a )) ] ≈ Id {_} {_} {_} {A} (FObj T a) ]
Lemma6 = \m -> IsMonad.unity2 ( isMonad m )

-- T = M x A

-- open import Data.Product -- renaming (_×_ to _*_)

-- MonoidalMonad : 
-- MonoidalMonad = record { Obj = M × A 
--                  ; Hom = _⟶_
--                  ; Id = IdProd
--                  ; _o_ = _∘_
--                  ; _≈_ = _≈_
--                  ; isCategory = record { isEquivalence = record { refl  = λ {φ} → ≈-refl {φ = φ}
--                                                                 ; sym   = λ {φ ψ} → ≈-symm {φ = φ} {ψ}
--                                                                 ; trans = λ {φ ψ σ} → ≈-trans {φ = φ} {ψ} {σ}
--                                                                 }
--                                        ; identityL = identityL
--                                        ; identityR = identityR
--                                        ; o-resp-≈ = o-resp-≈
--                                        ; associative = associative
--                                        }
--                  }


-- nat of η


-- g ○ f = μ(c) T(g) f
-- h ○ (g ○ f) = (h ○ g) ○ f

-- η(b) ○ f = f
-- f ○ η(a) = f


record Kleisli  { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ )
                 ( T : Functor A A )
                 ( η : NTrans A A identityFunctor T )
                 ( μ : NTrans A A (T ○ T) T )
                 ( M : Monad A T η μ ) : Set (suc (c₁ ⊔ c₂ ⊔ ℓ )) where
     join : { a b : Obj A } -> ( c : Obj A ) ->
                      ( Hom A b ( FObj T c )) -> (  Hom A a ( FObj T b)) -> Hom A a ( FObj T c )
     join c g f = A [ Trans μ c  o A [ FMap T g  o f ] ]

open Kleisli
Lemma7 : {c₁ c₂ ℓ : Level} {A : Category c₁ c₂ ℓ}
                 { T : Functor A A }
                 { η : NTrans A A identityFunctor T }
                 { μ : NTrans A A (T ○ T) T }
                 { a b : Obj A } 
                 { f : Hom A a ( FObj T b) } 
                 { M : Monad A T η μ } 
                 ( K : Kleisli A T η μ M) 
                      -> A  [ join ? ? ?  ≈ f ]
Lemma7 = \m -> {!!}




-- Kleisli :
-- Kleisli = record { Hom = 
--                 ; Hom = _⟶_
--                  ; Id = IdProd
--                  ; _o_ = _∘_
--                  ; _≈_ = _≈_
--                  ; isCategory = record { isEquivalence = record { refl  = λ {φ} → ≈-refl {φ = φ}
--                                                                 ; sym   = λ {φ ψ} → ≈-symm {φ = φ} {ψ}
--                                                                 ; trans = λ {φ ψ σ} → ≈-trans {φ = φ} {ψ} {σ}
--                                                                 }
--                                        ; identityL = identityL
--                                        ; identityR = identityR
--                                        ; o-resp-≈ = o-resp-≈
--                                        ; associative = associative
--                                        }
--                  }