Module Notation
val pr_notation : Constrexpr.notation -> Pp.tPrinting
val notation_entry_eq : Constrexpr.notation_entry -> Constrexpr.notation_entry -> boolEquality on
notation_entry.
val notation_entry_level_eq : Constrexpr.notation_entry_level -> Constrexpr.notation_entry_level -> boolEquality on
notation_entry_level.
val notation_eq : Constrexpr.notation -> Constrexpr.notation -> boolEquality on
notation.
module NotationSet : Stdlib.Set.S with type NotationSet.elt = Constrexpr.notationmodule NotationMap : CMap.ExtS with type key = Constrexpr.notation and module Set := NotationSetScopes
val declare_scope : Notation_term.scope_name -> unitval ensure_scope : Notation_term.scope_name -> unitval current_scopes : unit -> scopesval scope_is_open_in_scopes : Notation_term.scope_name -> scopes -> boolCheck where a scope is opened or not in a scope list, or in * the current opened scopes
val scope_is_open : Notation_term.scope_name -> bool
val open_close_scope : (bool * bool * Notation_term.scope_name) -> unitval empty_scope_stack : scopesExtend a list of scopes
val push_scope : Notation_term.scope_name -> scopes -> scopesval find_scope : Notation_term.scope_name -> scope
val declare_delimiters : Notation_term.scope_name -> delimiters -> unitval remove_delimiters : Notation_term.scope_name -> unitval find_delimiters_scope : ?loc:Loc.t -> delimiters -> Notation_term.scope_name
Declare and uses back and forth an interpretation of primitive token
type notation_location= (Names.DirPath.t * Names.DirPath.t) * stringtype required_module= Libnames.full_path * string listtype rawnum= Constrexpr.sign * Constrexpr.raw_numeral
type prim_token_uid= stringtype 'a prim_token_interpreter= ?loc:Loc.t -> 'a -> Glob_term.glob_constrtype 'a prim_token_uninterpreter= Glob_term.any_glob_constr -> 'a optiontype 'a prim_token_interpretation= 'a prim_token_interpreter * 'a prim_token_uninterpreter
val register_rawnumeral_interpretation : ?allow_overwrite:bool -> prim_token_uid -> rawnum prim_token_interpretation -> unitval register_bignumeral_interpretation : ?allow_overwrite:bool -> prim_token_uid -> Bigint.bigint prim_token_interpretation -> unitval register_string_interpretation : ?allow_overwrite:bool -> prim_token_uid -> string prim_token_interpretation -> unit
exceptionPrimTokenNotationError of string * Environ.env * Evd.evar_map * prim_token_notation_error
type numnot_option=|Nop|Warning of string|Abstract of stringtype int_ty={uint : Names.inductive;int : Names.inductive;}type z_pos_ty={z_ty : Names.inductive;pos_ty : Names.inductive;}type decimal_ty={int : int_ty;decimal : Names.inductive;}type target_kind=|Int of int_ty|UInt of Names.inductive|Z of z_pos_ty|Int63|Decimal of decimal_tytype string_target_kind=|ListByte|Bytetype option_kind=|Option|Directtype 'target conversion_kind= 'target * option_kindtype ('target, 'warning) prim_token_notation_obj={to_kind : 'target conversion_kind;to_ty : Names.GlobRef.t;of_kind : 'target conversion_kind;of_ty : Names.GlobRef.t;ty_name : Libnames.qualid;warning : 'warning;}type numeral_notation_obj= (target_kind, numnot_option) prim_token_notation_objtype string_notation_obj= (string_target_kind, unit) prim_token_notation_objtype prim_token_interp_info=|Uid of prim_token_uid|NumeralNotation of numeral_notation_obj|StringNotation of string_notation_objtype prim_token_infos={pt_local : bool;Is this interpretation local?
pt_scope : Notation_term.scope_name;Concerned scope
pt_interp_info : prim_token_interp_info;Unique id "pointing" to (un)interp functions, OR a numeral notation object describing (un)interp functions
pt_required : required_module;Module that should be loaded first
pt_refs : Names.GlobRef.t list;Entry points during uninterpretation
pt_in_match : bool;Is this prim token legal in match patterns ?
}
val enable_prim_token_interpretation : prim_token_infos -> unit
val declare_numeral_interpreter : ?local:bool -> Notation_term.scope_name -> required_module -> Bigint.bigint prim_token_interpreter -> (Glob_term.glob_constr list * Bigint.bigint prim_token_uninterpreter * bool) -> unitval 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
val interp_prim_token : ?loc:Loc.t -> Constrexpr.prim_token -> Notation_term.subscopes -> Glob_term.glob_constr * (notation_location * Notation_term.scope_name option)val 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)
val uninterp_prim_token : 'a Glob_term.glob_constr_g -> Notation_term.scope_name * Constrexpr.prim_tokenval uninterp_prim_token_cases_pattern : 'a Glob_term.cases_pattern_g -> Names.Name.t * Notation_term.scope_name * Constrexpr.prim_tokenval 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
type interp_rule=|NotationRule of Notation_term.scope_name option * Constrexpr.notation|SynDefRule of Names.KerName.tBinds a notation in a given scope to an interpretation
val declare_notation_interpretation : Constrexpr.notation -> Notation_term.scope_name option -> Notation_term.interpretation -> notation_location -> onlyprint:bool -> Deprecation.t option -> unitval declare_uninterpretation : interp_rule -> Notation_term.interpretation -> unitval interp_notation : ?loc:Loc.t -> Constrexpr.notation -> Notation_term.subscopes -> Notation_term.interpretation * (notation_location * Notation_term.scope_name option)Return the interpretation bound to a notation
type notation_rule= interp_rule * Notation_term.interpretation * int option
val uninterp_notations : 'a Glob_term.glob_constr_g -> notation_rule listReturn the possible notations for a given term
val uninterp_cases_pattern_notations : 'a Glob_term.cases_pattern_g -> notation_rule listval uninterp_ind_pattern_notations : Names.inductive -> notation_rule listval availability_of_notation : (Notation_term.scope_name option * Constrexpr.notation) -> Notation_term.subscopes -> (Notation_term.scope_name option * delimiters option) optionTest 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
Miscellaneous
val interp_notation_as_global_reference : ?loc:Loc.t -> (Names.GlobRef.t -> bool) -> Constrexpr.notation_key -> delimiters option -> Names.GlobRef.tval exists_notation_in_scope : Notation_term.scope_name option -> Constrexpr.notation -> bool -> Notation_term.interpretation -> boolChecks for already existing notations
val declare_arguments_scope : bool -> Names.GlobRef.t -> Notation_term.scope_name option list -> unitDeclares and looks for scopes associated to arguments of a global ref
val find_arguments_scope : Names.GlobRef.t -> Notation_term.scope_name option list
val scope_class_compare : scope_class -> scope_class -> intComparison of scope_class
val subst_scope_class : Mod_subst.substitution -> scope_class -> scope_class optionval declare_scope_class : Notation_term.scope_name -> scope_class -> unitval declare_ref_arguments_scope : Evd.evar_map -> Names.GlobRef.t -> unitval compute_arguments_scope : Evd.evar_map -> EConstr.types -> Notation_term.scope_name option listval compute_type_scope : Evd.evar_map -> EConstr.types -> Notation_term.scope_name optionval current_type_scope_name : unit -> Notation_term.scope_name optionGet the current scope bound to Sortclass, if it exists
val scope_class_of_class : Classops.cl_typ -> scope_class
type symbol=|Terminal of string|NonTerminal of Names.Id.t|SProdList of Names.Id.t * symbol list|Break of int
val symbol_eq : symbol -> symbol -> boolval make_notation_key : Constrexpr.notation_entry_level -> symbol list -> Constrexpr.notationMake/decompose a notation of the form "_ U _"
val decompose_notation_key : Constrexpr.notation -> Constrexpr.notation_entry_level * symbol listval decompose_raw_notation : string -> symbol listDecompose a notation of the form "a 'U' b"
val pr_scope_class : scope_class -> Pp.tPrints scopes (expects a pure aconstr printer)
val pr_scope : (Glob_term.glob_constr -> Pp.t) -> Notation_term.scope_name -> Pp.tval pr_scopes : (Glob_term.glob_constr -> Pp.t) -> Pp.tval locate_notation : (Glob_term.glob_constr -> Pp.t) -> Constrexpr.notation_key -> Notation_term.scope_name option -> Pp.tval pr_visibility : (Glob_term.glob_constr -> Pp.t) -> Notation_term.scope_name option -> Pp.t
type entry_coercion= Constrexpr.notation list
val declare_entry_coercion : Constrexpr.notation -> Constrexpr.notation_entry_level -> unitval availability_of_entry_coercion : Constrexpr.notation_entry_level -> Constrexpr.notation_entry_level -> entry_coercion optionval declare_custom_entry_has_global : string -> int -> unitval declare_custom_entry_has_ident : string -> int -> unitval entry_has_global : Constrexpr.notation_entry_level -> boolval entry_has_ident : Constrexpr.notation_entry_level -> bool
val with_notation_protection : ('a -> 'b) -> 'a -> 'bval int63_of_pos_bigint : Bigint.bigint -> Uint63.tConversion from bigint to int63