File: bug_4498.v

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: sid, trixie
  • size: 44,116 kB
  • sloc: ml: 234,160; sh: 4,301; python: 3,270; ansic: 2,644; makefile: 882; lisp: 172; javascript: 63; xml: 24; sed: 2
file content (23 lines) | stat: -rw-r--r-- 831 bytes parent folder | download | duplicates (5)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
Require Export Coq.Unicode.Utf8.
Require Export Coq.Classes.Morphisms.
Require Export Coq.Relations.Relation_Definitions.
Require Export Coq.Setoids.Setoid.

Set Universe Polymorphism.

Reserved Notation "a ~> b" (at level 90, right associativity).

Class Category := {
  ob            : Type;
  uhom          := Type : Type;
  hom           : ob → ob → uhom where "a ~> b" := (hom a b);
  compose       : ∀ {A B C}, (B ~> C) → (A ~> B) → (A ~> C);
  equiv         : ∀ {A B}, relation (A ~> B);
  is_equiv      : ∀ {A B}, @Equivalence (A ~> B) equiv;
  comp_respects : ∀ {A B C},
      Proper (@equiv B C ==> @equiv A B ==> @equiv A C) (@compose A B C);
}.

Add Parametric Morphism `{Category} {A B C} : (@compose _ A B C) with
  signature equiv ==> equiv ==> equiv as compose_mor.
Proof. apply comp_respects. Qed.