File: prettyp.mli

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: forky, 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 (98 lines) | stat: -rw-r--r-- 4,000 bytes parent folder | download | duplicates (2)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *         Copyright INRIA, CNRS and contributors             *)
(* <O___,, * (see version control and CREDITS file for authors & dates) *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)

open Names
open Environ
open Libnames

(** This module implements Print/About/Locate commands *)

val assumptions_for_print : Name.t list -> Termops.names_context (* unused? *)

val print_closed_sections : bool ref
val print_context (* unused? *)
  : env
  -> Evd.evar_map
  -> with_values:Global.indirect_accessor option -> int option -> 'a Lib.library_segment -> Pp.t
val print_library_leaf
  : env
  -> Evd.evar_map
  -> with_values:Global.indirect_accessor option
  -> ModPath.t
  -> Libobject.t
  -> Pp.t option
val print_library_node : Summary.Interp.frozen Lib.node -> Pp.t (* unused? *)
val print_full_context : Global.indirect_accessor -> env -> Evd.evar_map -> Pp.t
val print_full_context_typ : env -> Evd.evar_map -> Pp.t

val print_full_pure_context : Global.indirect_accessor -> env -> Evd.evar_map -> Pp.t

val print_sec_context : Global.indirect_accessor -> env -> Evd.evar_map -> qualid -> Pp.t
val print_sec_context_typ : env -> Evd.evar_map -> qualid -> Pp.t
val print_judgment : env -> Evd.evar_map -> EConstr.unsafe_judgment -> Pp.t
val print_safe_judgment : Safe_typing.judgment -> Pp.t

val print_name : Global.indirect_accessor -> env -> Evd.evar_map
  -> qualid Constrexpr.or_by_notation
  -> UnivNames.full_name_list option
  -> Pp.t
val print_notation : env -> Evd.evar_map
  -> Constrexpr.notation_entry
  -> string
  -> Pp.t

val print_abbreviation : Global.indirect_accessor -> env -> Evd.evar_map -> KerName.t -> Pp.t

val print_about : env -> Evd.evar_map -> qualid Constrexpr.or_by_notation ->
  UnivNames.full_name_list option -> Pp.t
val print_impargs : env -> GlobRef.t -> Pp.t

(** Pretty-printing functions for classes and coercions *)
val print_graph : unit -> Pp.t
val print_classes : unit -> Pp.t
val print_coercions : unit -> Pp.t
val print_coercion_paths : Coercionops.cl_typ -> Coercionops.cl_typ -> Pp.t
val print_canonical_projections : env -> Evd.evar_map -> GlobRef.t list -> Pp.t

(** Pretty-printing functions for type classes and instances *)
val print_typeclasses : unit -> Pp.t
val print_instances : GlobRef.t -> Pp.t
val print_all_instances : unit -> Pp.t

val inspect : env -> Evd.evar_map -> int -> Pp.t

(** {5 Locate} *)

type 'a locatable_info = {
  locate : qualid -> 'a option;
  (** Locate the most precise object with the provided name if any. *)
  locate_all : qualid -> 'a list;
  (** Locate all objects whose name is a suffix of the provided name *)
  shortest_qualid : 'a -> qualid;
  (** Return the shortest name in the current context *)
  name : 'a -> Pp.t;
  (** Data as printed by the Locate command *)
  print : 'a -> Pp.t;
  (** Data as printed by the Print command *)
  about : 'a -> Pp.t;
  (** Data as printed by the About command *)
}
(** Generic data structure representing locatable objects. *)

val register_locatable : string -> 'a locatable_info -> unit
(** Define a new type of locatable objects that can be reached via the
    corresponding generic vernacular commands. The string should be a unique
    name describing the kind of objects considered and that is added as a
    grammar command prefix for vernacular commands Locate. *)

val print_located_qualid : env -> qualid -> Pp.t
val print_located_term : env -> qualid -> Pp.t
val print_located_module : env -> qualid -> Pp.t
val print_located_other : env -> string -> qualid -> Pp.t