package coq

  1. Overview
  2. Docs
Legend:
Page
Library
Module
Module type
Parameter
Class
Class type
Source

Module PrinterSource

These are the entry points for printing terms, context, tac, ...

Sourceval print_goal_tag_opt_name : string list

Terms

Printers for terms.

The "lconstr" variant does not require parentheses to isolate the expression from the surrounding context (for instance 3 + 4 will be written 3 + 4). The "constr" variant (w/o "l") enforces parentheses whenever the term is not an atom (for instance, 3 will be written 3 but 3 + 4 will be written (3 + 4).

~inctx:true indicates that the term is intended to be printed in a context where its type is known so that a head coercion would be skipped, or implicit arguments inferable from the context will not be made explicit. For instance, if foo is declared as a coercion, foo bar will be printed as bar if inctx is true and as foo bar otherwise.

~scope:some_scope_name indicates that the head of the term is intended to be printed in scope some_scope_name. It defaults to None.

~impargs:some_list_of_binding_kind indicates the implicit arguments of the external quatification. Only used for printing types (not terms), and at toplevel (only "l" versions). It defaults to None.

~lax:true is for debugging purpose. It defaults to ~lax:false.

Sourceval pr_constr_env : ?lax:bool -> ?inctx:bool -> ?scope:Notation_term.scope_name -> Environ.env -> Evd.evar_map -> Constr.constr -> Pp.t
Sourceval pr_lconstr_env : ?lax:bool -> ?inctx:bool -> ?scope:Notation_term.scope_name -> Environ.env -> Evd.evar_map -> Constr.constr -> Pp.t
Sourceval pr_constr_n_env : ?lax:bool -> ?inctx:bool -> ?scope:Notation_term.scope_name -> Environ.env -> Evd.evar_map -> Constrexpr.entry_relative_level -> Constr.constr -> Pp.t

Same, but resilient to Nametab errors. Prints fully-qualified names when shortest_qualid_of_global has failed. Prints "??" in case of remaining issues (such as reference not in env).

Sourceval safe_pr_constr_env : Environ.env -> Evd.evar_map -> Constr.constr -> Pp.t
Sourceval safe_pr_lconstr_env : Environ.env -> Evd.evar_map -> Constr.constr -> Pp.t
Sourceval pr_econstr_env : ?lax:bool -> ?inctx:bool -> ?scope:Notation_term.scope_name -> Environ.env -> Evd.evar_map -> EConstr.t -> Pp.t
Sourceval pr_leconstr_env : ?lax:bool -> ?inctx:bool -> ?scope:Notation_term.scope_name -> Environ.env -> Evd.evar_map -> EConstr.t -> Pp.t
Sourceval pr_econstr_n_env : ?lax:bool -> ?inctx:bool -> ?scope:Notation_term.scope_name -> Environ.env -> Evd.evar_map -> Constrexpr.entry_relative_level -> EConstr.t -> Pp.t
Sourceval pr_etype_env : ?lax:bool -> ?goal_concl_style:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Pp.t
Sourceval pr_letype_env : ?lax:bool -> ?goal_concl_style:bool -> Environ.env -> Evd.evar_map -> ?impargs:Glob_term.binding_kind list -> EConstr.types -> Pp.t
Sourceval pr_open_constr_env : ?lax:bool -> ?inctx:bool -> ?scope:Notation_term.scope_name -> Environ.env -> Evd.evar_map -> Evd.open_constr -> Pp.t
Sourceval pr_open_lconstr_env : ?lax:bool -> ?inctx:bool -> ?scope:Notation_term.scope_name -> Environ.env -> Evd.evar_map -> Evd.open_constr -> Pp.t
Sourceval pr_constr_under_binders_env : Environ.env -> Evd.evar_map -> Ltac_pretype.constr_under_binders -> Pp.t
Sourceval pr_lconstr_under_binders_env : Environ.env -> Evd.evar_map -> Ltac_pretype.constr_under_binders -> Pp.t

Printers for types. Types are printed in scope "type_scope" and under the constraint of being of type a sort.

The "ltype" variant does not require parentheses to isolate the expression from the surrounding context (for instance nat * bool will be written nat * bool). The "type" variant (w/o "l") enforces parentheses whenever the term is not an atom (for instance, nat will be written nat but nat * bool will be written (nat * bool).

~goal_concl_style:true tells to print the type the same way as command Show would print a goal. Concretely, it means that all names of goal/section variables and all names of variables referred by de Bruijn indices (if any) in the given environment and all short names of global definitions of the current module must be avoided while printing bound variables. Otherwise, short names of global definitions are printed qualified and only names of goal/section variables and rel names that do _not_ occur in the scope of the binder to be printed are avoided.

~lax:true is for debugging purpose.

Sourceval pr_ltype_env : ?lax:bool -> ?goal_concl_style:bool -> Environ.env -> Evd.evar_map -> ?impargs:Glob_term.binding_kind list -> Constr.types -> Pp.t
Sourceval pr_type_env : ?lax:bool -> ?goal_concl_style:bool -> Environ.env -> Evd.evar_map -> Constr.types -> Pp.t
Sourceval pr_closed_glob_n_env : ?lax:bool -> ?goal_concl_style:bool -> ?inctx:bool -> ?scope:Notation_term.scope_name -> Environ.env -> Evd.evar_map -> Constrexpr.entry_relative_level -> Ltac_pretype.closed_glob_constr -> Pp.t
Sourceval pr_closed_glob_env : ?lax:bool -> ?goal_concl_style:bool -> ?inctx:bool -> ?scope:Notation_term.scope_name -> Environ.env -> Evd.evar_map -> Ltac_pretype.closed_glob_constr -> Pp.t
Sourceval pr_lglob_constr_env : Environ.env -> Evd.evar_map -> 'a Glob_term.glob_constr_g -> Pp.t
Sourceval pr_glob_constr_env : Environ.env -> Evd.evar_map -> 'a Glob_term.glob_constr_g -> Pp.t
Sourceval pr_lconstr_pattern_env : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> Pp.t
Sourceval pr_constr_pattern_env : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> Pp.t
Sourceval pr_cases_pattern : Glob_term.cases_pattern -> Pp.t
Sourceval pr_sort : Evd.evar_map -> Sorts.t -> Pp.t

Universe constraints

Sourceval pr_universe_instance : Evd.evar_map -> Univ.Instance.t -> Pp.t
Sourceval pr_universe_instance_constraints : Evd.evar_map -> Univ.Instance.t -> Univ.Constraint.t -> Pp.t
Sourceval pr_universe_ctx : Evd.evar_map -> ?variance:Univ.Variance.t array -> Univ.UContext.t -> Pp.t
Sourceval pr_abstract_universe_ctx : Evd.evar_map -> ?variance:Univ.Variance.t array -> ?priv:Univ.ContextSet.t -> Univ.AUContext.t -> Pp.t
Sourceval pr_universe_ctx_set : Evd.evar_map -> Univ.ContextSet.t -> Pp.t
Sourceval pr_universes : Evd.evar_map -> ?variance:Univ.Variance.t array -> ?priv:Univ.ContextSet.t -> Declarations.universes -> Pp.t
Sourceval universe_binders_with_opt_names : Univ.AUContext.t -> UnivNames.univ_name_list option -> UnivNames.universe_binders

universe_binders_with_opt_names ref l

If l is Some univs return the universe binders naming the bound levels of ref by univs (generating names for Anonymous). May error if the lengths mismatch.

Otherwise return the bound universe names registered for ref.

Inefficient on large contexts due to name generation.

Printing global references using names as short as possible

Sourceval pr_global_env : Names.Id.Set.t -> Names.GlobRef.t -> Pp.t
Sourceval pr_global : Names.GlobRef.t -> Pp.t
Sourceval pr_constant : Environ.env -> Names.Constant.t -> Pp.t
Sourceval pr_existential_key : Evd.evar_map -> Evar.t -> Pp.t
Sourceval pr_constructor : Environ.env -> Names.constructor -> Pp.t
Sourceval pr_inductive : Environ.env -> Names.inductive -> Pp.t
Sourceval pr_evaluable_reference : Tacred.evaluable_global_reference -> Pp.t

Contexts

Sourceval set_compact_context : bool -> unit

Display compact contexts of goals (simple hyps on the same line)

Sourceval get_compact_context : unit -> bool
Sourceval pr_context_unlimited : Environ.env -> Evd.evar_map -> Pp.t
Sourceval pr_ne_context_of : Pp.t -> Environ.env -> Evd.evar_map -> Pp.t
Sourceval pr_named_context_of : Environ.env -> Evd.evar_map -> Pp.t
Sourceval pr_rel_context_of : Environ.env -> Evd.evar_map -> Pp.t
Sourceval pr_context_of : Environ.env -> Evd.evar_map -> Pp.t

Predicates

Sourceval pr_predicate : ('a -> Pp.t) -> (bool * 'a list) -> Pp.t
Sourceval pr_cpred : Names.Cpred.t -> Pp.t
Sourceval pr_idpred : Names.Id.Pred.t -> Pp.t
Sourceval pr_transparent_state : TransparentState.t -> Pp.t

Proofs, these functions obey Hyps Limit and Compact contexts.

Sourceval pr_goal : ?diffs:bool -> ?og_s:Goal.goal Evd.sigma -> Goal.goal Evd.sigma -> Pp.t

pr_goal ~diffs ~og_s g_s prints the goal specified by g_s. If diffs is true, highlight the differences between the old goal, og_s, and g_s. g_s and og_s are records containing the goal and sigma for, respectively, the new and old proof steps, e.g. { it = g ; sigma = sigma }.

Sourceval pr_subgoals : ?pr_first:bool -> ?diffs:bool -> ?os_map:(Evd.evar_map * Goal.goal Evar.Map.t) -> Pp.t option -> Evd.evar_map -> seeds:Goal.goal list -> shelf:Goal.goal list -> stack:int list -> unfocused:Goal.goal list -> goals:Goal.goal list -> Pp.t

pr_subgoals ~pr_first ~diffs ~os_map close_cmd sigma ~seeds ~shelf ~stack ~unfocused ~goals prints the goals in goals followed by the goals in unfocused in a compact form (typically only the conclusion). If pr_first is true, print the first goal in full. close_cmd is printed afterwards verbatim.

If diffs is true, then highlight diffs relative to os_map in the output for first goal. os_map contains sigma for the old proof step and the goal map created by Proof_diffs.make_goal_map.

This function prints only the focused goals unless the corresponding option enable_unfocused_goal_printing is set. seeds is for printing dependent evars (mainly for emacs proof tree mode). shelf is from Proof.proof and is used to identify shelved goals in a message if there are no more subgoals but there are non-instantiated existential variables. stack is used to print summary info on unfocused goals.

Sourceval pr_subgoal : int -> Evd.evar_map -> Goal.goal list -> Pp.t
Sourceval pr_concl : int -> ?diffs:bool -> ?og_s:Goal.goal Evd.sigma -> Evd.evar_map -> Goal.goal -> Pp.t

pr_concl n ~diffs ~og_s sigma g prints the conclusion of the goal g using sigma. The output is labelled "subgoal n". If diffs is true, highlight the differences between the old conclusion, og_s, and g+sigma. og_s is a record containing the old goal and sigma, e.g. { it = g ; sigma = sigma }.

Sourceval pr_open_subgoals_diff : ?quiet:bool -> ?diffs:bool -> ?oproof:Proof.t -> Proof.t -> Pp.t

pr_open_subgoals_diff ~quiet ~diffs ~oproof proof shows the context for proof as used by, for example, coqtop. The first active goal is printed with all its antecedents and the conclusion. The other active goals only show their conclusions. If diffs is true, highlight the differences between the old proof, oproof, and proof. quiet disables printing messages as Feedback.

Sourceval pr_open_subgoals : proof:Proof.t -> Pp.t
Sourceval pr_nth_open_subgoal : proof:Proof.t -> int -> Pp.t
Sourceval pr_evars_int : Evd.evar_map -> shelf:Goal.goal list -> given_up:Goal.goal list -> int -> Evd.evar_info Evar.Map.t -> Pp.t
Sourceval pr_ne_evar_set : Pp.t -> Pp.t -> Evd.evar_map -> Evar.Set.t -> Pp.t
Sourceval print_and_diff : Proof.t option -> Proof.t option -> unit
Sourceval print_dependent_evars : Evar.t option -> Evd.evar_map -> Evar.t list -> Pp.t
Sourcetype axiom =
  1. | Constant of Names.Constant.t
  2. | Positive of Names.MutInd.t
  3. | Guarded of Names.GlobRef.t
  4. | TypeInType of Names.GlobRef.t
  5. | UIP of Names.MutInd.t

Declarations for the "Print Assumption" command

Sourcetype context_object =
  1. | Variable of Names.Id.t
  2. | Axiom of axiom * (Names.Label.t * Constr.rel_context * Constr.types) list
  3. | Opaque of Names.Constant.t
  4. | Transparent of Names.Constant.t
Sourceval pr_goal_by_id : proof:Proof.t -> Names.Id.t -> Pp.t
Sourceval pr_goal_emacs : proof:Proof.t option -> int -> int -> Pp.t
Sourceval pr_typing_flags : Declarations.typing_flags -> Pp.t
Sourceval print_goal_names : unit -> bool

Tells if flag "Printing Goal Names" is activated

OCaml

Innovation. Community. Security.