X-Git-Url: http://git.megacz.com/?p=coq-categories.git;a=blobdiff_plain;f=src%2FEnrichment_ch2_8.v;h=3d3aac9ada1dd2f0b4fa42dc7c357a3b3757710b;hp=87d9aa036780c12a14e2cdaf14cc81ecc46eb543;hb=f494aec0e2a8f5ccb7a5f560dd4f8a0b302feb40;hpb=bce224f0a9ddd24cfe16a710b4f7d5be64b32929 diff --git a/src/Enrichment_ch2_8.v b/src/Enrichment_ch2_8.v index 87d9aa0..3d3aac9 100644 --- a/src/Enrichment_ch2_8.v +++ b/src/Enrichment_ch2_8.v @@ -1,6 +1,5 @@ Generalizable All Variables. -Require Import Preamble. -Require Import General. +Require Import Notations. Require Import Categories_ch1_3. Require Import Functors_ch1_4. Require Import Isomorphisms_ch1_5. @@ -364,11 +363,11 @@ Instance UnderlyingFunctor `(EF:@EFunctor Ob Hom V bin_obj' bc EI mn Eob1 EHom1 Defined. Coercion UnderlyingFunctor : EFunctor >-> Functor. -Class EBinoidalCat `(ec:ECategory) := -{ ebc_bobj : ec -> ec -> ec -; ebc_first : forall a:ec, EFunctor ec ec (fun x => ebc_bobj x a) -; ebc_second : forall a:ec, EFunctor ec ec (fun x => ebc_bobj a x) +Class EBinoidalCat `(ec:ECategory)(bobj : ec -> ec -> ec) := +{ ebc_first : forall a:ec, EFunctor ec ec (fun x => bobj x a) +; ebc_second : forall a:ec, EFunctor ec ec (fun x => bobj a x) ; ebc_ec := ec (* this isn't a coercion - avoids duplicate paths *) +; ebc_bobj := bobj }. Instance EBinoidalCat_is_binoidal `(ebc:EBinoidalCat(ec:=ec)) : BinoidalCat (Underlying ec) ebc_bobj.