package coq

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

Module NotationSource

Notations

Sourceval pr_notation : Constrexpr.notation -> Pp.t

Printing

Sourceval notation_entry_eq : Constrexpr.notation_entry -> Constrexpr.notation_entry -> bool

Equality on notation_entry.

Sourceval notation_entry_level_eq : Constrexpr.notation_entry_level -> Constrexpr.notation_entry_level -> bool

Equality on notation_entry_level.

Sourceval notation_eq : Constrexpr.notation -> Constrexpr.notation -> bool

Equality on notation.

Scopes

A scope is a set of interpreters for symbols + optional interpreter and printers for integers + optional delimiters

Sourcetype delimiters = string
Sourcetype scope
Sourcetype scopes

= scope_name list

Sourceval declare_scope : Notation_term.scope_name -> unit
Sourceval ensure_scope : Notation_term.scope_name -> unit
Sourceval current_scopes : unit -> scopes
Sourceval scope_is_open_in_scopes : Notation_term.scope_name -> scopes -> bool

Check where a scope is opened or not in a scope list, or in * the current opened scopes

Sourceval scope_is_open : Notation_term.scope_name -> bool

Open scope

Sourceval open_close_scope : (bool * bool * Notation_term.scope_name) -> unit
Sourceval empty_scope_stack : scopes

Extend a list of scopes

Declare delimiters for printing

Sourceval declare_delimiters : Notation_term.scope_name -> delimiters -> unit
Sourceval remove_delimiters : Notation_term.scope_name -> unit
Sourceval find_delimiters_scope : ?loc:Loc.t -> delimiters -> Notation_term.scope_name
Declare and uses back and forth an interpretation of primitive token

A number interpreter is the pair of an interpreter for **(hexa)decimal** numbers in terms and an optional interpreter in pattern, if non integer or negative numbers are not supported, the interpreter must fail with an appropriate error message

Sourcetype notation_location = (Names.DirPath.t * Names.DirPath.t) * string
Sourcetype required_module = Libnames.full_path * string list
Sourcetype rawnum = NumTok.Signed.t

The unique id string below will be used to refer to a particular registered interpreter/uninterpreter of number or string notation. Using the same uid for different (un)interpreters will fail. If at most one interpretation of prim token is used per scope, then the scope name could be used as unique id.

Sourcetype prim_token_uid = string
Sourcetype 'a prim_token_interpreter = ?loc:Loc.t -> 'a -> Glob_term.glob_constr
Sourcetype 'a prim_token_uninterpreter = Glob_term.any_glob_constr -> 'a option
Sourcetype 'a prim_token_interpretation = 'a prim_token_interpreter * 'a prim_token_uninterpreter
Sourceval register_rawnumeral_interpretation : ?allow_overwrite:bool -> prim_token_uid -> rawnum prim_token_interpretation -> unit
Sourceval register_bignumeral_interpretation : ?allow_overwrite:bool -> prim_token_uid -> Z.t prim_token_interpretation -> unit
Sourceval register_string_interpretation : ?allow_overwrite:bool -> prim_token_uid -> string prim_token_interpretation -> unit

* Number notation

Sourcetype prim_token_notation_error =
  1. | UnexpectedTerm of Constr.t
  2. | UnexpectedNonOptionTerm of Constr.t
Sourceexception PrimTokenNotationError of string * Environ.env * Evd.evar_map * prim_token_notation_error
Sourcetype numnot_option =
  1. | Nop
  2. | Warning of NumTok.UnsignedNat.t
  3. | Abstract of NumTok.UnsignedNat.t
Sourcetype int_ty = {
  1. dec_uint : Names.inductive;
  2. dec_int : Names.inductive;
  3. hex_uint : Names.inductive;
  4. hex_int : Names.inductive;
  5. uint : Names.inductive;
  6. int : Names.inductive;
}
Sourcetype z_pos_ty = {
  1. z_ty : Names.inductive;
  2. pos_ty : Names.inductive;
}
Sourcetype number_ty = {
  1. int : int_ty;
  2. decimal : Names.inductive;
  3. hexadecimal : Names.inductive;
  4. number : Names.inductive;
}
Sourcetype pos_neg_int63_ty = {
  1. pos_neg_int63_ty : Names.inductive;
}
Sourcetype target_kind =
  1. | Int of int_ty
  2. | UInt of int_ty
  3. | Z of z_pos_ty
  4. | Int63 of pos_neg_int63_ty
  5. | Number of number_ty
Sourcetype string_target_kind =
  1. | ListByte
  2. | Byte
Sourcetype option_kind =
  1. | Option
  2. | Direct
Sourcetype 'target conversion_kind = 'target * option_kind
Sourcetype to_post_arg =
  1. | ToPostCopy
  2. | ToPostAs of int
  3. | ToPostHole
  4. | ToPostCheck of Names.GlobRef.t

A postprocessing translation to_post can be done after execution of the to_ty interpreter. The reverse translation is performed before the of_ty uninterpreter.

to_post is an array of n lists l_i of tuples (f, t, args). When the head symbol of the translated term matches one of the f in the list l_0 it is replaced by t and its arguments are translated acording to args where ToPostCopy means that the argument is kept unchanged and ToPostAs k means that the argument is recursively translated according to l_k. ToPostHole introduces an additional implicit argument hole (in the reverse translation, the corresponding argument is removed). ToPostCheck r behaves as ToPostCopy except in the reverse translation which fails if the copied term is not r. When n is null, no translation is performed.

Sourcetype ('target, 'warning) prim_token_notation_obj = {
  1. to_kind : 'target conversion_kind;
  2. to_ty : Names.GlobRef.t;
  3. to_post : (Names.GlobRef.t * Names.GlobRef.t * to_post_arg list) list array;
  4. of_kind : 'target conversion_kind;
  5. of_ty : Names.GlobRef.t;
  6. ty_name : Libnames.qualid;
  7. warning : 'warning;
}
Sourcetype string_notation_obj = (string_target_kind, unit) prim_token_notation_obj
Sourcetype prim_token_interp_info =
  1. | Uid of prim_token_uid
  2. | NumberNotation of number_notation_obj
  3. | StringNotation of string_notation_obj
Sourcetype prim_token_infos = {
  1. pt_local : bool;
    (*

    Is this interpretation local?

    *)
  2. pt_scope : Notation_term.scope_name;
    (*

    Concerned scope

    *)
  3. pt_interp_info : prim_token_interp_info;
    (*

    Unique id "pointing" to (un)interp functions, OR a number notation object describing (un)interp functions

    *)
  4. pt_required : required_module;
    (*

    Module that should be loaded first

    *)
  5. pt_refs : Names.GlobRef.t list;
    (*

    Entry points during uninterpretation

    *)
  6. pt_in_match : bool;
    (*

    Is this prim token legal in match patterns ?

    *)
}

Note: most of the time, the pt_refs field above will contain inductive constructors (e.g. O and S for nat). But it could also be injection functions such as IZR for reals.

Activate a prim token interpretation whose unique id and functions have already been registered.

Sourceval enable_prim_token_interpretation : prim_token_infos -> unit

Compatibility. Avoid the next two functions, they will now store unnecessary objects in the library segment. Instead, combine register_*_interpretation and enable_prim_token_interpretation (the latter inside a Mltop.declare_cache_obj).

Sourceval declare_numeral_interpreter : ?local:bool -> Notation_term.scope_name -> required_module -> Z.t prim_token_interpreter -> (Glob_term.glob_constr list * Z.t prim_token_uninterpreter * bool) -> unit
Sourceval declare_string_interpreter : ?local:bool -> Notation_term.scope_name -> required_module -> string prim_token_interpreter -> (Glob_term.glob_constr list * string prim_token_uninterpreter * bool) -> unit

Return the term/cases_pattern bound to a primitive token in a given scope context

Sourceval interp_prim_token_cases_pattern_expr : ?loc:Loc.t -> (Names.GlobRef.t -> unit) -> Constrexpr.prim_token -> Notation_term.subscopes -> Glob_term.glob_constr * (notation_location * Notation_term.scope_name option)

Return the primitive token associated to a term/cases_pattern; raise No_match if no such token

Sourceval availability_of_prim_token : Constrexpr.prim_token -> Notation_term.scope_name -> Notation_term.subscopes -> delimiters option option
Declare and interpret back and forth a notation
Sourcetype interp_rule =
  1. | NotationRule of Constrexpr.specific_notation
  2. | SynDefRule of Names.KerName.t

Binds a notation in a given scope to an interpretation

Sourcetype notation_use =
  1. | OnlyPrinting
  2. | OnlyParsing
  3. | ParsingAndPrinting
Sourceval declare_uninterpretation : ?also_in_cases_pattern:bool -> interp_rule -> Notation_term.interpretation -> unit
Sourcetype entry_coercion_kind =
  1. | IsEntryCoercion of Constrexpr.notation_entry_level
  2. | IsEntryGlobal of string * int
  3. | IsEntryIdent of string * int
Sourceval declare_notation : (Constrexpr.notation_with_optional_scope * Constrexpr.notation) -> Notation_term.interpretation -> notation_location -> use:notation_use -> also_in_cases_pattern:bool -> entry_coercion_kind option -> Deprecation.t option -> unit

Return the interpretation bound to a notation

Sourcetype notation_applicative_status =
  1. | AppBoundedNotation of int
  2. | AppUnboundedNotation
  3. | NotAppNotation
Sourceval uninterp_notations : 'a Glob_term.glob_constr_g -> notation_rule list

Return the possible notations for a given term

Sourceval uninterp_cases_pattern_notations : 'a Glob_term.cases_pattern_g -> notation_rule list
Sourceval uninterp_ind_pattern_notations : Names.inductive -> notation_rule list
Sourceval availability_of_notation : Constrexpr.specific_notation -> Notation_term.subscopes -> (Notation_term.scope_name option * delimiters option) option

Test if a notation is available in the scopes context scopes; if available, the result is not None; the first argument is itself not None if a delimiters is needed

Sourceval is_printing_inactive_rule : interp_rule -> Notation_term.interpretation -> bool
Miscellaneous
Sourceval interp_notation_as_global_reference : ?loc:Loc.t -> head:bool -> (Names.GlobRef.t -> bool) -> Constrexpr.notation_key -> delimiters option -> Names.GlobRef.t

If head is true, also allows applied global references.

Sourceval declare_arguments_scope : bool -> Names.GlobRef.t -> Notation_term.scope_name option list -> unit

Declares and looks for scopes associated to arguments of a global ref

Sourceval find_arguments_scope : Names.GlobRef.t -> Notation_term.scope_name option list
Sourcetype scope_class
Sourceval scope_class_compare : scope_class -> scope_class -> int

Comparison of scope_class

Sourceval subst_scope_class : Environ.env -> Mod_subst.substitution -> scope_class -> scope_class option
Sourceval declare_scope_class : Notation_term.scope_name -> scope_class -> unit
Sourceval declare_ref_arguments_scope : Evd.evar_map -> Names.GlobRef.t -> unit
Sourceval compute_arguments_scope : Environ.env -> Evd.evar_map -> EConstr.types -> Notation_term.scope_name option list
Sourceval compute_type_scope : Environ.env -> Evd.evar_map -> EConstr.types -> Notation_term.scope_name option
Sourceval current_type_scope_name : unit -> Notation_term.scope_name option

Get the current scope bound to Sortclass, if it exists

Sourceval scope_class_of_class : Coercionops.cl_typ -> scope_class

Building notation key

Sourcetype symbol =
  1. | Terminal of string
  2. | NonTerminal of Names.Id.t
  3. | SProdList of Names.Id.t * symbol list
  4. | Break of int
Sourceval symbol_eq : symbol -> symbol -> bool

Make/decompose a notation of the form "_ U _"

Sourceval decompose_notation_key : Constrexpr.notation -> Constrexpr.notation_entry * symbol list
Sourceval decompose_raw_notation : string -> (Names.Id.t * Names.Id.t) list * Names.Id.t list * symbol list

Decompose a notation of the form "a 'U' b" together with the lists of pairs of recursive variables and the list of all variables binding in the notation

Sourceval pr_scope_class : scope_class -> Pp.t

Prints scopes (expects a pure aconstr printer)

Sourceval pr_scopes : (Glob_term.glob_constr -> Pp.t) -> Pp.t
Sourceval pr_visibility : (Glob_term.glob_constr -> Pp.t) -> Notation_term.scope_name option -> Pp.t
Sourceval declare_custom_entry_has_global : string -> int -> unit
Sourceval declare_custom_entry_has_ident : string -> int -> unit
Sourceval entry_has_global : Constrexpr.notation_entry_level -> bool
Sourceval entry_has_ident : Constrexpr.notation_entry_level -> bool

Dealing with precedences

Sourceval level_eq : level -> level -> bool
Declare and test the level of a (possibly uninterpreted) notation
Sourceval declare_notation_level : Constrexpr.notation -> level -> unit
Sourceval level_of_notation : Constrexpr.notation -> level

raise Not_found if not declared

Rem: printing rules for primitive token are canonical

Sourceval with_notation_protection : ('a -> 'b) -> 'a -> 'b
Sourceval int63_of_pos_bigint : Z.t -> Uint63.t

Conversion from bigint to int63

OCaml

Innovation. Community. Security.