comparison sandbox/FunctorExample.agda @ 10:7c7659d4521d

Improve NaturalTransformation definition
author Yasutaka Higa <e115763@ie.u-ryukyu.ac.jp>
date Tue, 20 Jan 2015 10:51:33 +0900
parents 4a0841123810
children 26e64661b969
comparison
equal deleted inserted replaced
9:4a0841123810 10:7c7659d4521d
6 module FunctorExample where 6 module FunctorExample where
7 7
8 id : {l : Level} {A : Set l} -> A -> A 8 id : {l : Level} {A : Set l} -> A -> A
9 id x = x 9 id x = x
10 10
11 _∙_ : {l ll lll : Level} {A : Set l} {B : Set ll} {C : Set lll} -> (B -> C) -> (A -> B) -> (A -> C) 11 _∙_ : {l : Level} {A B C : Set l} -> (B -> C) -> (A -> B) -> (A -> C)
12 f ∙ g = \x -> f (g x) 12 f ∙ g = \x -> f (g x)
13 13
14 record Functor {l : Level} {A : Set l} (F : {l' : Level} -> Set l' -> Set l') : Set (suc l) where
15 field
16 fmap : {A B : Set l} -> (A -> B) -> (F A) -> (F B)
17 field
18 preserve-id : {A : Set l} (x : F A) → fmap id x ≡ id x
19 covariant : {A B C : Set l} (f : A -> B) -> (g : B -> C) -> (x : F A)
20 -> fmap (g ∙ f) x ≡ ((fmap g) ∙ (fmap f)) x
21 open Functor
14 22
15 23
16 record Functor {l : Level} (F : Set l -> Set l) : (Set (suc l)) where
17 field
18 fmap : ∀{A B} -> (A -> B) -> (F A) -> (F B)
19 field
20 preserve-id : ∀{A} (Fa : F A) → fmap id Fa ≡ id Fa
21 covariant : ∀{A B C} (f : A → B) → (g : B → C) → (x : F A)
22 → fmap (g ∙ f) x ≡ fmap g (fmap f x)
23 24
24 data List {l : Level} (A : Set l) : (Set l) where 25 data List {l : Level} (A : Set l) : (Set l) where
25 nil : List A 26 nil : List A
26 cons : A -> List A -> List A 27 cons : A -> List A -> List A
27 28
31 32
32 list-preserve-id : {l : Level} {A : Set l} -> (xs : List A) -> list-fmap id xs ≡ id xs 33 list-preserve-id : {l : Level} {A : Set l} -> (xs : List A) -> list-fmap id xs ≡ id xs
33 list-preserve-id nil = refl 34 list-preserve-id nil = refl
34 list-preserve-id (cons x xs) = cong (\li -> cons x li) (list-preserve-id xs) 35 list-preserve-id (cons x xs) = cong (\li -> cons x li) (list-preserve-id xs)
35 36
36 list-covariant : {l ll lll : Level} {A : Set l} {B : Set ll} {C : Set lll} -> 37 list-covariant : {l : Level} {A B C : Set l} ->
37 (f : A -> B) → (g : B -> C) → (x : List A) → list-fmap (g ∙ f) x ≡ list-fmap g (list-fmap f x) 38 (f : A -> B) → (g : B -> C) → (x : List A) → list-fmap (g ∙ f) x ≡ list-fmap g (list-fmap f x)
38 list-covariant f g nil = refl 39 list-covariant f g nil = refl
39 list-covariant f g (cons x xs) = cong (\li -> cons (g (f x)) li) (list-covariant f g xs) 40 list-covariant f g (cons x xs) = cong (\li -> cons (g (f x)) li) (list-covariant f g xs)
40 41
41 42
42 list-is-functor : {l : Level} -> Functor (List {l}) 43 list-is-functor : {l : Level} {A : Set l}-> Functor {l} {A} List
43 list-is-functor = record { fmap = list-fmap ; 44 list-is-functor = record { fmap = list-fmap ;
44 preserve-id = list-preserve-id ; 45 preserve-id = list-preserve-id ;
45 covariant = list-covariant} 46 covariant = list-covariant}
46 47
47 fmap-to-nest-list : {l ll : Level} {A : Set l} {B : Set l} {fl : Functor List} 48 fmap-to-nest-list : {l : Level} {A B : Set l}
48 -> (A -> B) -> (List (List A)) -> (List (List B)) 49 -> (A -> B) -> (List (List A)) -> (List (List B))
49 fmap-to-nest-list {_} {_} {_} {_} {fl} f xs = Functor.fmap fl (Functor.fmap fl f) xs 50 fmap-to-nest-list {l} {A} {B} f xs = Functor.fmap (list-is-functor {l} {List A}) (Functor.fmap {l} {A} list-is-functor f) xs
50 51
51 data Identity {l : Level} (A : Set l) : Set l where 52 data Identity {l : Level} (A : Set l) : Set l where
52 identity : A -> Identity A 53 identity : A -> Identity A
53 54
54 identity-fmap : {l ll : Level} {A : Set l} {B : Set ll} -> (A -> B) -> Identity A -> Identity B 55 identity-fmap : {l ll : Level} {A : Set l} {B : Set ll} -> (A -> B) -> Identity A -> Identity B
55 identity-fmap f (identity a) = identity (f a) 56 identity-fmap f (identity a) = identity (f a)
56 57
57 identity-preserve-id : {l : Level} {A : Set l} -> (x : Identity A) -> identity-fmap id x ≡ id x 58 identity-preserve-id : {l : Level} {A : Set l} -> (x : Identity A) -> identity-fmap id x ≡ id x
58 identity-preserve-id (identity x) = refl 59 identity-preserve-id (identity x) = refl
59 60
60 identity-covariant : {l ll lll : Level} {A : Set l} {B : Set ll} {C : Set lll} -> 61 identity-covariant : {l : Level} {A B C : Set l}
61 (f : A -> B) → (g : B -> C) → (x : Identity A) → identity-fmap (g ∙ f) x ≡ identity-fmap g (identity-fmap f x) 62 (f : A -> B) → (g : B -> C) → (x : Identity A) → identity-fmap (g ∙ f) x ≡ identity-fmap g (identity-fmap f x)
62 identity-covariant f g (identity x) = refl 63 identity-covariant f g (identity x) = refl
63 64
64 identity-is-functor : {l : Level} -> Functor (Identity {l}) 65 identity-is-functor : {l : Level} {A : Set l} -> Functor {l} {A} Identity
65 identity-is-functor {l} = record { fmap = identity-fmap {l}; 66 identity-is-functor {l} = record { fmap = identity-fmap {l};
66 preserve-id = identity-preserve-id ; 67 preserve-id = identity-preserve-id ;
67 covariant = identity-covariant } 68 covariant = identity-covariant }
68 69
69 70
70 71
71 72
72 record NaturalTransformation {l ll : Level} (F G : Set l -> Set l) 73 record NaturalTransformation {l : Level} (F G : {l' : Level} -> Set l' -> Set l')
73 (functorF : Functor F) 74 {fmapF : {A B : Set l} -> (A -> B) -> (F A) -> (F B)}
74 (functorG : Functor G) : Set (suc (l ⊔ ll)) where 75 {fmapG : {A B : Set l} -> (A -> B) -> (G A) -> (G B)}
76 (natural-transformation : {A : Set l} -> F A -> G A)
77 : Set (suc l) where
75 field 78 field
76 natural-transformation : {A : Set l} -> F A -> G A 79 commute : {A B : Set l} -> (f : A -> B) -> (x : F A) ->
77 field 80 natural-transformation (fmapF f x) ≡ fmapG f (natural-transformation x)
78 commute : ∀ {A B} -> (function : A -> B) -> (x : F A) -> 81 open NaturalTransformation
79 natural-transformation (Functor.fmap functorF function x) ≡ Functor.fmap functorG function (natural-transformation x)
80 82
81 tail : {l : Level} {A : Set l} -> List A -> List A 83 tail : {l : Level} {A : Set l} -> List A -> List A
82 tail nil = nil 84 tail nil = nil
83 tail (cons _ xs) = xs 85 tail (cons _ xs) = xs
84 86
85 tail-commute : {l ll : Level} {A : Set l} {B : Set ll} -> (f : A -> B) -> (xs : List A) -> 87 tail-commute : {l : Level} {A B : Set l} -> (f : A -> B) -> (xs : List A) ->
86 tail (list-fmap f xs) ≡ list-fmap f (tail xs) 88 tail (list-fmap f xs) ≡ list-fmap f (tail xs)
87 tail-commute f nil = refl 89 tail-commute f nil = refl
88 tail-commute f (cons x xs) = refl 90 tail-commute f (cons x xs) = refl
89 91
90 tail-is-natural-transformation : {l ll : Level} -> NaturalTransformation {l} {ll} List List list-is-functor list-is-functor 92
91 tail-is-natural-transformation = record { natural-transformation = tail; 93 tail-is-natural-transformation : {l : Level} -> NaturalTransformation List List {list-fmap} {list-fmap {l}} tail
92 commute = tail-commute} 94 tail-is-natural-transformation = record { commute = tail-commute}
95
96
97 append : {l : Level} {A : Set l} -> List A -> List A -> List A
98 append nil ys = ys
99 append (cons x xs) ys = cons x (append xs ys)
100
101 append-with-fmap : {l : Level} {A B : Set l} -> (f : A -> B) -> (xs : List A) -> (ys : List A) ->
102 append (list-fmap f xs) (list-fmap f ys) ≡ list-fmap f (append xs ys)
103 append-with-fmap f nil ys = refl
104 append-with-fmap f (cons x xs) ys = begin
105 append (list-fmap f (cons x xs)) (list-fmap f ys) ≡⟨ refl ⟩
106 append (cons (f x) (list-fmap f xs)) (list-fmap f ys) ≡⟨ refl ⟩
107 cons (f x) (append (list-fmap f xs) (list-fmap f ys)) ≡⟨ cong (\li -> cons (f x) li) (append-with-fmap f xs ys) ⟩
108 list-fmap f (append (cons x xs) ys) ∎
109
110
111
112 concat : {l : Level} {A : Set l} -> List (List A) -> List A
113 concat nil = nil
114 concat (cons x xs) = append x (concat xs)
115
116 concat-commute : {l : Level} {A B : Set l} -> (f : A -> B) -> (xs : List (List A)) ->
117 concat ((list-fmap ∙ list-fmap) f xs) ≡ list-fmap f (concat xs)
118 concat-commute f nil = refl
119 concat-commute f (cons x xs) = begin
120 concat ((list-fmap ∙ list-fmap) f (cons x xs)) ≡⟨ refl ⟩
121 concat (cons (list-fmap f x) ((list-fmap ∙ list-fmap) f xs)) ≡⟨ refl ⟩
122 append (list-fmap f x) (concat ((list-fmap ∙ list-fmap) f xs)) ≡⟨ cong (\li -> append (list-fmap f x) li) (concat-commute f xs) ⟩
123 append (list-fmap f x) (list-fmap f (concat xs)) ≡⟨ append-with-fmap f x (concat xs) ⟩
124 list-fmap f (append x (concat xs)) ≡⟨ refl ⟩
125 list-fmap f (concat (cons x xs))
126
127
128 concat-is-natural-transformation : {l : Level} -> NaturalTransformation (\A -> List (List A)) List
129 {list-fmap ∙ list-fmap} {list-fmap {l}} concat
130 concat-is-natural-transformation = record {commute = concat-commute}