File: bug_3368.v

package info (click to toggle)
coq 9.1.0%2Bdfsg-3
  • links: PTS, VCS
  • area: main
  • in suites: experimental
  • size: 35,964 kB
  • sloc: ml: 239,908; sh: 4,355; python: 2,985; ansic: 2,644; makefile: 874; lisp: 171; javascript: 63; xml: 24; sed: 2
file content (16 lines) | stat: -rw-r--r-- 870 bytes parent folder | download | duplicates (6)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
(* File reduced by coq-bug-finder from 7411 lines to 2271 lines., then from 889 lines to 119 lines, then from 76 lines to 19 lines *)
Set Universe Polymorphism.
Set Implicit Arguments.
Set Primitive Projections.
Record PreCategory := { object :> Type; morphism : object -> object -> Type }.
Record Functor (C D : PreCategory) :=
  { object_of :> C -> D;
    morphism_of : forall s d, morphism C s d -> morphism D (object_of s) (object_of d) }.
Definition opposite (C : PreCategory) : PreCategory := @Build_PreCategory C (fun s d => morphism C d s).
Definition opposite' C D (F : Functor C D)
  := Build_Functor (opposite C) (opposite D)
                   (object_of F)
                   (fun s d => @morphism_of C D F d s).
(* Toplevel input, characters 15-191:
Anomaly: File "pretyping/reductionops.ml", line 149, characters 4-10: Assertion failed.
Please report. *)