11Require Import Basics.Overture Basics.Tactics.
22Require Import Types .Forall.
3- Require Import WildCat.Core WildCat.Opposite WildCat. Universe WildCat. Prod.
3+ Require Import WildCat.Core WildCat.Prod.
44
55(** * Bifunctors between WildCats *)
66
@@ -114,9 +114,10 @@ Global Instance is0functor_bifunctor_uncurried01 {A B C : Type}
114114 (F : A * B -> C) `{!Is0Functor F} (a : A)
115115 : Is0Functor (fun b => F (a, b)).
116116Proof .
117- apply Build_Is0Functor.
118- intros x y f.
119- rapply (fmap F).
117+ nrapply (is0functor_compose (fun b => (a, b)) F).
118+ 2: exact _.
119+ nrapply Build_Is0Functor.
120+ intros b c f.
120121 exact (Id a, f).
121122Defined .
122123
@@ -125,27 +126,25 @@ Global Instance is1functor_bifunctor_uncurried01 {A B C : Type}
125126 (F : A * B -> C) `{!Is0Functor F, !Is1Functor F} (a : A)
126127 : Is1Functor (fun b => F (a, b)).
127128Proof .
128- apply Build_Is1Functor.
129- - intros x y f g p.
130- rapply (fmap2 F).
129+ nrapply (is1functor_compose (fun b => (a, b)) F).
130+ 2: exact _.
131+ nrapply Build_Is1Functor.
132+ - intros b c f g p.
131133 exact (Id _, p).
132- - intros b.
133- exact (fmap_id F (a, b)).
134- - intros x y z f g.
135- refine (_ $@ _).
136- 2: rapply (fmap_comp F).
137- rapply (fmap2 F).
138- exact ((cat_idl (Id a))^$, Id _).
134+ - intros b; reflexivity.
135+ - intros b c d f g.
136+ exact ((cat_idl _)^$, Id _).
139137Defined .
140138
141139Global Instance is0functor_bifunctor_uncurried10 {A B C : Type }
142140 `{IsGraph A} `{Is01Cat B} `{IsGraph C}
143141 (F : A * B -> C) `{!Is0Functor F} (b : B)
144142 : Is0Functor (fun a => F (a, b)).
145143Proof .
146- apply Build_Is0Functor.
147- intros x y f.
148- rapply (fmap F).
144+ nrapply (is0functor_compose (fun a => (a, b)) F).
145+ 2: exact _.
146+ nrapply Build_Is0Functor.
147+ intros a c f.
149148 exact (f, Id b).
150149Defined .
151150
@@ -154,16 +153,14 @@ Global Instance is1functor_bifunctor_uncurried10 {A B C : Type}
154153 (F : A * B -> C) `{!Is0Functor F, !Is1Functor F} (b : B)
155154 : Is1Functor (fun a => F (a, b)).
156155Proof .
157- apply Build_Is1Functor.
158- - intros x y f g p.
159- rapply (fmap2 F).
156+ nrapply (is1functor_compose (fun a => (a, b)) F).
157+ 2: exact _.
158+ nrapply Build_Is1Functor.
159+ - intros a c f g p.
160160 exact (p, Id _).
161- - intros a.
162- exact (fmap_id F (a, b)).
163- - intros x y z f g.
164- refine (_ $@ _).
165- 2: rapply (fmap_comp F).
166- rapply (fmap2 F).
161+ - intros a; reflexivity.
162+ - intros a c d f g.
163+ cbn.
167164 exact (Id _, (cat_idl _)^$).
168165Defined .
169166
@@ -181,14 +178,11 @@ Global Instance is1bifunctor_bifunctor_uncurried {A B C : Type}
181178 : Is1Bifunctor (fun a b => F (a, b)).
182179Proof .
183180 apply Build_Is1Bifunctor.
184- - intros a.
185- exact (is1functor_bifunctor_uncurried01 F a).
186- - intros b.
187- exact (is1functor_bifunctor_uncurried10 F b).
188- - intros a b f c d g.
189- refine ((fmap_comp F _ _)^$ $@ _).
190- refine (_ $@ (fmap_comp F _ _)).
191- rapply (fmap2 F).
192- refine (cat_idl f $@ (cat_idr f)^$, _).
193- exact (cat_idr g $@ (cat_idl g)^$).
181+ 1,2: exact _.
182+ intros a b f c d g.
183+ refine ((fmap_comp F _ _)^$ $@ _).
184+ refine (_ $@ (fmap_comp F _ _)).
185+ rapply (fmap2 F).
186+ refine (cat_idl f $@ (cat_idr f)^$, _).
187+ exact (cat_idr g $@ (cat_idl g)^$).
194188Defined .
0 commit comments