Git Product home page Git Product logo

jwiegley / category-theory Goto Github PK

View Code? Open in Web Editor NEW
740.0 19.0 68.0 3.29 MB

An axiom-free formalization of category theory in Coq for personal study and practical work

License: BSD 3-Clause "New" or "Revised" License

Coq 99.00% Makefile 0.11% Haskell 0.27% Nix 0.53% Shell 0.09%
category-theory coq monad profunctor comonads monoid profunctor-composition construction category categories functor cartesian cartesian-closed-category

category-theory's Introduction

Category Theory in Coq

This development encodes category theory in Coq, with the primary aim being to allow representation and manipulation of categorical terms, as well realization of those terms in various target categories.

  • Versions used: Coq 8.14.1, 8.15.2, 8.16+rc1.
  • Some parts depend on Coq-Equations 1.2.4, 1.3.

Usage

It is recommended to include this library in your developments by adding the following to your _CoqProject file:

-R <path to this library> Category

Then include the primary elements of the library using:

Require Import Category.Theory.

Library structure

This library is broken up into several major areas:

  • Core Theory, covering topics such as categories, functors, natural transformations, adjunctions, kan extensions, etc.

  • Categorical Structure, which reveals internal structure of a category by way of morphisms related to some universal property.

  • Categorical Construction, which introduces external structure by building new categories out of existing ones.

  • Species of different kinds of Functor, Natural.Transformation, Adjunction and Kan.Extension; for example: categories with monoidal structure give rise to monoidal functors preserving this structure, which in turn leads to monoidal transformations that transform functors while preserving their monoidal mapping property.

  • The Instance directory defines various categories; some of these are fairly general, such as the category of preorders, applicable to any PreOrder relation, but in general these are either not defined in terms of other categories, or are sufficiently specific.

  • When a concept, such as limits, can be defined using more fundamental terms, that version of limits can be found in a subdirectory of the derived concept, for example there is Category.Structure.Limit and Category.Limit.Kan.Extension. This is done to demonstrate the relationship of ideas; for example: Category.Construction.Comma.Natural.Transformation. As a result, files with the same name occur often, with the parent directory establishing intent.

Programming sub-library

Within Theory.Coq there is now a sub-library that continues work started in the coq-haskell library. This sub-library is specifically aimed at "applied category theory" for programmers in the category of Coq types and functions. The aim is to mimic the utility of Haskell's monad hierarchy -- but for Coq users, similar to what ext-lib achieves. I've also adjusted a few of my notations to more closely match ext-lib.

Where this library differs, and what is considered the main contribution of this work, is that laws are not defined for these structures in the sub-library. Rather, the programmer establishs that her Monad is lawful by proving that a mapping exists from that definition into the general definition of monads (found in Theory.Monad) specialized to the category Coq. In this way the rest of the category-theory library is leveraged to discharge these proof obligations, while keeping the programmatic side as simple as possible.

For example, it is trivial to define the composition of two Applicatives. What is not so trivial is proving that this is truly an Applicative, based on that simple definition. While this proof was done in coq-haskell, it requires a bit of Ltac magic to keep the size down:

Program Instance Compose_ApplicativeLaws
  `{ApplicativeLaws F} `{ApplicativeLaws G} : ApplicativeLaws (F \o G).
Obligation 2. (* ap_composition *)
  (* Discharge w *)
  rewrite <- ap_comp; f_equal.
  (* Discharge v *)
  rewrite <- !ap_fmap, <- ap_comp.
  symmetry.
  rewrite <- ap_comp; f_equal.
  (* Discharge u *)
  apply_applicative_laws.
  f_equal.
  extensionality y.
  extensionality x.
  extensionality x0.
  rewrite <- ap_comp, ap_fmap.
  reflexivity.
Qed.

What's ill-gotten about this proof is that it's confined to the very specific case of Coq applicative endo-functors. However, there is a more general truth, which is that any two lax monoidal functors in any monoidal category also compose. So why not appeal to that proof to establish that our programmatic applicative in Coq is lawful by construction?

This is what the new sub-library does. Since the more general proof is already completed (and is too large to paste here), one may appeal to it directly to establish the desired fact:

(* The composition of two applicatives is itself applicative. We establish
   this by appeal to the general proofs that:

   1. every Coq functor has strength;
   2. (also, but not needed: any two strong functors compose to a strong
      functor; if it is a Coq functor it is known to have strength); and
   3. any two lax monoidal functors compose to a lax monoidal functor. *)

Theorem Compose_IsApplicative
  `(HF : EndoFunctor F)
  `(AF : @Functor.Applicative.Applicative _ _ (FromAFunctor HF))
  `(HG : EndoFunctor G)
  `(AG : @Functor.Applicative.Applicative _ _ (FromAFunctor HG)) :
  IsApplicative (Compose_IsFunctor HF HG)
    (@Compose_Applicative
       F HF (EndoApplicative_Applicative HF AF)
       G HG (EndoApplicative_Applicative HG AG)).
Proof.
  construct.
  - apply (@Compose_LaxMonoidalFunctor _ _ _ _ _ _ _ _ AF AG).
  - apply Coq_StrongFunctor.
Qed.

This proof pulls in several instances to establish that the category Coq is cartesian, closed, and thus closed monoidal, etc.

The hope is this will become a happy marriage of simple, useful computational constructions for Coq programmers, with relevant proof results from what category theory tells us about these structures in general, such as the above fact concerning composition of monoidal functors.

Duality

The core theory is defined in such a way that "the dual of the dual" is evident to Coq by mere simplification (that is, "C^op^op = C" follows by reflexivity for the opposite of the opposite of categories, functors, natural transformation, adjunctions, isomorphisms, etc.).

For this to be true, certain artificial constructions are necessary, such as repeating the associativity law for categories in symmetric form, and likewise the naturality condition of natural transformations. This repeats proof obligations when constructing these objects, but pays off in the ability to avoid repetition when stating the dual of whole structures.

As a result, the definition of comonads, for example, is reduced to one line:

Definition Comonad `{M : C ⟶ C} := @Monad (C^op) (M^op).

Most dual constructions are similarly defined, with the exception of Initial and Cocartesian structures. Although the core classes are indeed defined in terms of their dual construction, an alternate surface syntax and set of theorems is provided (for example, using a + b to indicate coproducts) to make life is a little less confusing for the reader. For instance, it follows from duality that 0 + X ≅ X is just 1 × X ≅ X in the opposite category, but using separate notations makes it easier to see that these two isomorphisms in the same category are not identical. This is especially important because Coq hides implicit parameters that would usually let you know duality is involved.

Design decisions

Some features and choices made in this library:

  • Type classes are used throughout to present concepts. When a type class instance would be too general -- and thus overlap with other instances -- it is presented as a definition that can later be added to instance resolution with Existing Instance.

  • All definitions are in Type, so that Prop is not used anywhere except specific category instances defined over Prop, such as the category Rel with heterogeneous relations as arrows.

  • No axioms are used anywhere in the core theory; they appear only at times when defining specific category instances, mostly for the Coq category.

  • Homsets are defined as computationally-relevant homsetoids (that is, using crelation). This is done using a variant of the Setoid type class defined for this library; likewise, the category of Sets is actually the category of such setoids. This increases the proof work necessary to establish definitions -- since preservation of the equivalence relation is required at all levels -- but allows categories to be flexible in what it means for two arrows to be equivalent.

Notations

There are many notations used through the library, which are chosen in an attempt to make complex constructions appear familiar to those reading modern texts on category theory. Some of the key notations are:

  • is equivalence; equality is almost never used.
  • ≈[Cat] is equivalence between arrows of some category, here Cat; this is only needed when type inference fails because it tries to find both the type of the arguments, and the type class instance for the equivalence
  • is isomorphism; apply it with to or from
  • is used specifically for isomorphisms between homsets in Sets
  • iso⁻¹ also indicates the reverse direction of an isomorphism
  • X ~> Y: a squiggly arrow between objects is a morphism
  • X ~{C}~> Y: squiggly arrows may also specify the intended category
  • id[C]: many known morphisms allow specifying the intended category; sometimes this is even used in the printing format
  • C ⟶ D: a long right arrow between categories is a functor
  • F ⟹ G: a long right double arrow between functors is a natural transformation
  • f ∘ g: a small centered circle is composition of morphisms, or horizontal composition generally
  • f ∘[Cat] g: composition can specify the intended category, as an aid to type inference composition generally
  • f ○ g: a larger hollow circle is composition of functors
  • f ⊙ g: a larger circle with a dot is composition of isomorphisms
  • f ∙ g: a solid composition dot is composition of natural transformations, or vertical composition generally
  • f ⊚ g: a larger circle with a smaller circle is composition of adjunctions
  • ([C, D]): A pair of categories in square brackets is another way to give the type of a functor, being an object of the category Fun(C, D)
  • F ~{[C, D]}~> G: An arrow in a functor category is a natural transformation
  • F ⊣ G: the turnstile is used for adjunctions
  • Cartesian categories use as the fork operation and × for products
  • Cocartesian categories use as the merge operation and + for coproducts
  • Closed categories use ^ for exponents and ≈> for the internal hom
  • As objects, the numbers 0 and 1 refer to initial and terminal objects
  • As categories, the numbers 0 and 1 refer to the initial and terminal objects of Cat
  • Products of categories can be specified using , which does not require pulling in the cartesian definition of Cat
  • Coproducts of categories can be specified using , which does not require pulling in the cocartesian definition of Cat
  • Products of functors are given with F ∏⟶ G, combining product and functor notations; the same for ∐⟶
  • Comma categories of two functors are given by (F ↓ G)
  • Likewise, the arrow category of C⃗
  • Slice categories use a unicode forward slash C̸c, since the normal forward slash is not considered an operator
  • Coslice categories use c ̸co C, to avoid ambiguity

Future directions

Computational Solver

There are some equivalences in category-theory that are very easily expressed and proven, but slow to establish in Coq using only symbolic term rewriting. For example:

(f ∘ g) △ (h ∘ i) ≈ split f h ∘ g △ i

This is solved by unfolding the definition of split, and then rewriting so that the fork operation (here given by ) absorbs the terms to its left, followed by observing the associativity of composition, and then simplify based on the universal properties of products. This is repeated several times until the prove is trivially completed.

Although this is easy to state, and even to write a tactic for, it can be extremely slow, especially when the types of the terms involved becomes large. A single rewrite can take several seconds to complete for some terms, in my experience.

The goal of this solver is to reify the above equivalence in terms of its fundamental operations, and then, using what we know about the laws of category theory, to compute the equivalence down to an equation on indices between the reduced terms. This is called computational reflection, and encodes the fact that our solution only depends on the categorical structure of the terms, and not their type.

That is, an incorrectly-built structure will simply fail to solve; but since we're reflecting over well-typed expressions to build the structure we pass to the solver, combined with a proof of soundness for that solver, we can know that solvable, well-typed, terms always give correct solutions. In this way, we transfer the problem to a domain without types, only indices, solve the structural problem there, and then bring the solution back to the domain of full types by way of the soundness proof.

Compiling to categories

Work has started in Tools/Abstraction for compiling monomorphic Gallina functions into "categorical terms" that can then be instantiated in any supporting target category using Coq's type class instance resolution.

This is as a Coq implementation of an idea developed by Conal Elliott, which he first implemented in and for Haskell. It is hoped that the medium of categories may provide a sound means for transporting Gallina terms into other syntactic domains, without relying on Coq's extraction mechanism.

License

This library is made available under the MIT license, a copy of which is included in the file LICENSE. Basically: you are free to use it for any purpose, personal or commercial (including proprietary derivates), so long as a copy of the license file is maintained in the derived work. Further, any acknowledgement referring back to this repository, while not necessary, is certainly appreciated.

John Wiegley

category-theory's People

Contributors

alizter avatar andres-erbsen avatar ekmett avatar gares avatar gmalecha avatar herbelin avatar jonsterling avatar jwiegley avatar nickmertin avatar patrick-nicodemus avatar ppedrot avatar proux01 avatar skyskimmer avatar t-wissmann avatar timjb avatar

Stargazers

 avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar

Watchers

 avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar  avatar

category-theory's Issues

Definition of full is nonstandard

From Theory/Functor.v -

Class Full `(F : C ⟶ D) := {
  prefmap {x y} (g : F x ~> F y) : x ~> y;

  prefmap_respects {x y} : Proper (equiv ==> equiv) (@prefmap x y);

  prefmap_id : ∀ x, @prefmap x x id ≈ id;
  prefmap_comp : ∀ x y z (f : F y ~> F z) (g : F x ~> F y),
    prefmap (f ∘ g) ≈ prefmap f ∘ prefmap g;

  fmap_sur {x y} (g : F x ~> F y) : fmap[F] (prefmap g) ≈ g
}.

This definition of 'full' is nonstandard. Do you have a motivation/justification for this choice of definition?

I would advocate instead

Class Full `(F : C ⟶ D) := {
  prefmap {x y} (g : F x ~> F y) : x ~> y;
  fmap_sur {x y} (g : F x ~> F y) : fmap[F] (prefmap g) ≈ g
}.

Or for modularity, define what it means for a map of setoids to be surjective and use that definition.

The question of whether prefmap should respect equality is more an issue of constructive mathematics than of category theory but in Bishop-style constructive math you do not require existence to respect equality.

Pdf link is not available.

Hi, I came across your repo and tried to get to the pdf file mentioned. However it is not active anymore. Is it possible to include it in the repo itself? Thanks for your help.

Can you explain uhom?

So here is some code in Category.v
Class Category := {
ob : Type;
uhom := Type : Type;
hom : ob → ob → uhom where "a ~> b" := (hom a b);
...

what will happend if hom is ob -> ob -> Type instead, why wont universe polymorphism work?

Compatibility with Coq 8.8?

Hi, I'm trying to update the version of this library in nixpkgs, but when I add the following lines to the default.nix there:

    "8.8" = {
      version = "20180705";
      rev = "2685dc91f528dc3e82f17a1c32804a94d2ee8ed7";
      sha256 = "0ll22lfzjglnvvf4p0vwk8crwn2c5lkawd85wr3qzcrwnpp5dc43";
    };

and

    compatibleCoqVersions = v: builtins.elem v [ "8.6" "8.7" "8.8" ];

I get this error when building:

File "./Lib/Nomega.v", line 92, characters 2-21:
Warning: Nge is N.ge [compatibility-notation,deprecated]
COQC Lib/Tactics.v
COQC Lib/Datatypes.v
COQC Lib/Equality.v
Closed under the global context
     = tt
     : typeD (fun _ : type => unit) Ty
     = 5
     : typeD (fun _ : type => nat) Ty
File "./Lib/Datatypes.v", line 212, characters 0-4:
Error:
Illegal application:
The term "list_equivalence_obligation_1" of type
 "∀ (A : Type) (H : Setoid A), Reflexive list_equiv"
cannot be applied to the terms
 "A" : "Type"
 "H" : "Setoid A"
The 1st term has type "Type@{Category.Lib.Datatypes.1007}"
which should be coercible to "Type@{Category.Lib.Datatypes.1020}".

Is this library compatible with Coq 8.8?

How to build?

Sorry to ask a stupid question but I just can't build from master branch. I am wonder how am I supposed to build? I am using Coq 8.8.1 and equations, but it keeps fail to build with follow messages:

NOT IN _CoqProject: Instance/LambdaOld.v
NOT IN _CoqProject: Instance/Lambda/Lec1.v
NOT IN _CoqProject: Instance/Lambda/Connect.v
NOT IN _CoqProject: Instance/Lambda/Nominal.v
NOT IN _CoqProject: Instance/Lambda/Lemmas.v
NOT IN _CoqProject: Instance/Lambda/Definitions.v
NOT IN _CoqProject: Instance/Lambda/Lec3.v
NOT IN _CoqProject: Instance/Lambda/Lec2.v
make -j -f Makefile.coq  # TIMECMD=time
make[1]: Entering directory '/home/hu/projects/category-theory'
make[2]: *** No rule to make target 'Instance/Lambda/Definitions.vo', needed by 'Instance/Lambda.vo'.  Stop.
COQC Lib/Datatypes.v
make[2]: *** Waiting for unfinished jobs....
File "./Lib/Datatypes.v", line 212, characters 0-4:
Error:
Illegal application: 
The term "list_equivalence_obligation_1" of type
 "∀ (A : Type) (H : Setoid A), Reflexive list_equiv"
cannot be applied to the terms
 "A" : "Type"
 "H" : "Setoid A"
The 1st term has type "Type@{Category.Lib.Datatypes.1007}"
which should be coercible to "Type@{Category.Lib.Datatypes.1020}".

Makefile.coq:656: recipe for target 'Lib/Datatypes.vo' failed
make[2]: *** [Lib/Datatypes.vo] Error 1
Makefile.coq:317: recipe for target 'all' failed
make[1]: *** [all] Error 2
make[1]: Leaving directory '/home/hu/projects/category-theory'
Makefile:2: recipe for target 'all' failed
make: *** [all] Error 2

Even if I added those missing files to _CoqProject, it still fails to build, not to mention the universe inconsistency. How should I proceed?

Also, it seems Lambda depends on Metalib, which is not stated as dependency in README. Can I just skip Lambda entirely?

CartesianMonoidal cleanup

I suggest some reorganization and renaming in Structure/Monoidal/Cartesian.v and Structure/Monoidal/Cartesian/Cartesian.v.

There are two relations between monoidal products and Cartesian products:

  1. Every Cartesian product is a monoidal product.
  2. A monoidal product which is both semicartesian and relevant is a Cartesian product, this is the result by by Heunen and Vicary proven in Structure/Monoidal/Cartesian.v.

The second theorem is important and interesting, but I don't expect it to be actually heavily used by users of the library. Instead, I suspect that if a user wants to prove that their category is Cartesian, they will show that for each pair of objects x, y there is an object z with the universal property of the Cartesian product. I have never personally needed or used theorem 2 to prove that a concrete category has a Cartesian product.

On the other hand, 1. is obviously very important, because it allows us to prove theorems about general monoidal products which automatically hold for Cartesian products. I'm not aware of a proof of 1. in the library at this moment. It should be added. I can try to do this when I get a chance.

I also suggest reorganizing things to deemphasize the proof of theorem 2 in order to make it easier to find the proof of theorem 1. I think a file named Monoidal/Cartesian.v which contains a typeclass called "CartesianMonoidal" and proves that a "CartesianMonoidal" product is Cartesian is likely to be misleading to readers who are looking for a proof that a Cartesian category is also a monoidal category. Furthermore it is defined as a typeclass which is probably unnecessary as, realistically, this theorem is applied rarely enough that it doesn't need typeclass resolution.

I would suggest we simply delete the typeclass CartesianMonoidal and just prove the result by separately assuming the two typeclasses RelevantMonoidal and Semicartesian. A pair of only two assumptions does not need to be bundled together, and the choice of bundle name pollutes the library namespace. We should also rename the files "Monoidal/Cartesian.v" and "Monoidal/Cartesian/Cartesian.v" to something that does less overloading of the heavily used terms "Monoidal" and "Cartesian" such as, for example, "Heunen_Vicary.v".

StrictCat

Tl; dr: The category Cat is defined "incorrectly", in particular its setoid hom equivalence relation is wrong (it does not agree with the ZFC definition of equality between functors). I propose that we redefine the category Cat with the "correct" setoid hom equivalence relation defined in StrictCat. This would be a long term project as obviously the category Cat is used in many places in the library.

Longer version:
As it stands, most of the library is consistent with ZFC-style classical mathematics. It does not use the law of excluded middle or UIP, so it admits interpretations in other theories such as homotopy type theory, etc. (Although due to technical issues with the propositional erasure in the extraction backend, the univalence axiom is not compatible with use of the sort Prop, so it would require a substantial and fairly unrealistic rewrite of the library to assume that axiom.)

In interpreting this library in the context of ZFC, you would generally quotient out by the setoid relation on the homsets of a category, because ZFC can prove function extensionality. This is why we want to use setoid homs, because Coq lacks function extensionality. Thus, when defining any concrete category whose objects are types with additional structure and whose morphisms are functions, we would generally want the setoid equivalence relationship on morphisms to be given by f \approx g := \forall x. f x = g x.
Indeed this is the case for most concrete categories in the library. An exception is the category Cat, where we identify two functors if they are naturally isomorphic. This is inconsistent with equality of functors in ZFC. The ZFC definition of equality between functors would be that two functors F, G are equal if Fx = Gx for all objects and F(f) = G(f) for all morphisms, appropriately transported along equalities.
In homotopy type theory this would be an acceptable definition for univalent categories but because we are not assuming the univalence axiom we do not get any benefits from defining things this way, so there is no good reason to appeal to homotopy type theory principles here. The downside is that it becomes impossible to define functors from the category Cat into any other category if it does not send naturally isomorphic functors to setoid-equivalent morphisms.
In a previous PR I gave an example of such a functor, namely the functor Cat -> Graph which sends each category to its underlying directed graph and sends each functor to its graph isomorphism. This is a well defined functor in ZFC but is not a well-defined functor in the current library, because it does not respect natural isomorphism of functors, which is why I created the alternative category StrictCat of categories and functors satisfying strict equality.

Conflicts in notations with Coq 8.19

Lib/Foundation.v defines notations for ∀ , ∃, and λ

If user code imports both Category.Lib and Utf8 from the Coq standard library, we get the following error message

Error: Notation "∀ _ .. _ , _" is already defined at level 200 with arguments
binder, constr at level 200 while it is now required to be at level 10
with arguments binder, constr at level 200.

The following patch seems enough to work around this issue: I confirmed category-theory compiles and so does my code.
Before proposing a pull request, I am curious about your feedback on this. I must admit my patch is pretty arbitrary, it ressembles the notation definition from the Coq standard library to avoid conflicts.

diff --git a/Lib/Foundation.v b/Lib/Foundation.v
index 82f7455..560a8c2 100644
--- a/Lib/Foundation.v
+++ b/Lib/Foundation.v
@@ -31,7 +31,7 @@ Arguments Basics.arrow _ _ /.
 Arguments Datatypes.id {_} _ /.
 
 Notation "∀  x .. y , P" := (forall x, .. (forall y, P) ..)
-  (at level 200, x binder, y binder, right associativity) :
+  (at level 10, x binder, y binder, P at level 200, right associativity) :
   category_theory_scope.
 
 Notation "'exists' x .. y , p" := (sigT (fun x => .. (sigT (fun y => p)) ..))
@@ -40,7 +40,7 @@ Notation "'exists' x .. y , p" := (sigT (fun x => .. (sigT (fun y => p)) ..))
   category_theory_scope.
 
 Notation "∃  x .. y , P" := (exists x, .. (exists y, P) ..)
-  (at level 200, x binder, y binder, right associativity) :
+  (at level 10, x binder, y binder, P at level 200, right associativity) :
   category_theory_scope.
 
 Notation "x → y" := (x -> y)
@@ -55,5 +55,5 @@ Infix "∧" := prod (at level 80, right associativity) : category_theory_scope.
 Infix "∨" := sum (at level 85, right associativity) : category_theory_scope.
 
 Notation "'λ'  x .. y , t" := (fun x => .. (fun y => t) ..)
-  (at level 200, x binder, y binder, right associativity) :
+  (at level 10, x binder, y binder, t at level 200, right associativity) :
   category_theory_scope.

Thanks,

Sets_Terminal is insufficiently universe-polymorphic

With Printing Universes enabled, the full definition is as follows:

Sets_Terminal@{u} = 
{|
  Terminal.terminal_obj := {| carrier := poly_unit@{Set}; is_setoid := Unit_Setoid@{Set} |};
  Terminal.one :=
    λ x : obj[Sets@{Set Set u Set Set}],
    {| morphism := λ _ : x, ttt@{Set}; proper_morphism := Sets.Sets_Terminal_obligation_1@{u} x |};
  Terminal.one_unique :=
    λ (x : obj[Sets@{Set Set u Set Set}])
    (f
     g : x ~{ Sets@{Set Set u Set Set}
         }~> {| carrier := poly_unit@{Set}; is_setoid := Unit_Setoid@{Set} |}),
    Sets.Sets_Terminal_obligation_2@{u} x f g
|}
     : Terminal.Terminal@{u Set}
(* u |= Set < u *)

This imposes a lot of unnecessary constraints on various universes to be Set, which then quickly propagates throughout the code from wherever Sets_Terminal is used.

not build for 8.12

the commit seems to indicate the library would work for 8.12 but it doesn't work for me:

File "./Theory/Isomorphism.v", line 89, characters 29-33:
Error:
 (in proof isomorphism_equiv_equivalence_obligation_2): Attempt to save an incomplete proof

Makefile.coq:715: recipe for target 'Theory/Isomorphism.vo' failed
make[2]: *** [Theory/Isomorphism.vo] Error 1
File "./Lib/MapDecide.v", line 333, characters 29-30:
Error:
Syntax error: [constr:operconstr level 200] expected after '[' (in [constr:operconstr]).
coqc -v
The Coq Proof Assistant, version 8.12.0 (September 2020)
compiled on Sep 1 2020 1:01:58 with OCaml 4.07.1

which Coq version works for you?

The reference X0 was not found in the current environment

Hi guys! I ran make and got the following error:

File "./Structure/Limit/Kan/Extension.v", line 65, characters 17-19:
Error: The reference X0 was not found in the current environment.

Coq version is 8.9.0.

Best regards,
Vladimir.

Duplicated proofs in Isomorphism.v

Hi, the file defines identity isos and iso composition. The equivalence relation for isos defines the same isos for reflexivity and transitivity. You can remove one of these definitions.
Best,
Sven

Mistakes in the definition of Poset

The definition of Poset in Instance/Poset.v currently is:

Definition Poset {C : Category} `{R : relation C}
           `(P : PreOrder C R) `{Asymmetric C R} : Category := Proset P.

I think there are two issues:

  1. C should only be a Type and not a category (the morphism structure isn't used anywhere as far as I can see)
  2. The relation R should be Antisymmetric instead of Asymmetric.

This second issue implies that all inhabitants of Poset are empty, i.e. any element of a poset leads to a contradiction:

Require Import Category.Lib.
Require Import Category.Theory.Category.
Require Import Category.Instance.Proset.

Require Import Coq.Classes.Equivalence.
Require Import Coq.Relations.Relation_Definitions.

Generalizable All Variables.
Set Implicit Arguments.

Definition Poset {C : Type} `{R : relation C}
           `(P : PreOrder C R) `{Asymmetric C R} : Category := Proset P.

Theorem there_is_no_poset
  {C : Type} `(P : PreOrder C R) `{Asymmetric C R} :
  Poset P -> False.
Proof.
  intro x.
  simpl in *.
  destruct P as [refl _].
  eapply H; unshelve apply (refl x); apply (refl x).
Qed.

I assume the definition should be revised as follows:

Definition Poset {C : Type} `{R : relation C}
           `(P : PreOrder C R) `{Antisymmetric C R} : Category := Proset P.

If there aren't any further subtleties that I missed, I can prepare a PR.

Recommend Projects

  • React photo React

    A declarative, efficient, and flexible JavaScript library for building user interfaces.

  • Vue.js photo Vue.js

    🖖 Vue.js is a progressive, incrementally-adoptable JavaScript framework for building UI on the web.

  • Typescript photo Typescript

    TypeScript is a superset of JavaScript that compiles to clean JavaScript output.

  • TensorFlow photo TensorFlow

    An Open Source Machine Learning Framework for Everyone

  • Django photo Django

    The Web framework for perfectionists with deadlines.

  • D3 photo D3

    Bring data to life with SVG, Canvas and HTML. 📊📈🎉

Recommend Topics

  • javascript

    JavaScript (JS) is a lightweight interpreted programming language with first-class functions.

  • web

    Some thing interesting about web. New door for the world.

  • server

    A server is a program made to process requests and deliver data to clients.

  • Machine learning

    Machine learning is a way of modeling and interpreting data that allows a piece of software to respond intelligently.

  • Game

    Some thing interesting about game, make everyone happy.

Recommend Org

  • Facebook photo Facebook

    We are working to build community through open source technology. NB: members must have two-factor auth.

  • Microsoft photo Microsoft

    Open source projects and samples from Microsoft.

  • Google photo Google

    Google ❤️ Open Source for everyone.

  • D3 photo D3

    Data-Driven Documents codes.