(************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* Gram.symbol val gram_token_of_string : string -> Gram.symbol (** The superclass of all grammar entries *) type grammar_object (** Type of reinitialization data *) type gram_reinit = gram_assoc * gram_position (** General entry keys *) (** This intermediate abstract representation of entries can both be reified into mlexpr for the ML extensions and dynamically interpreted as entries for the Coq level extensions *) type ('self, 'a) entry_key = ('self, 'a) Extend.symbol = | Atoken : Tok.t -> ('self, string) entry_key | Alist1 : ('self, 'a) entry_key -> ('self, 'a list) entry_key | Alist1sep : ('self, 'a) entry_key * string -> ('self, 'a list) entry_key | Alist0 : ('self, 'a) entry_key -> ('self, 'a list) entry_key | Alist0sep : ('self, 'a) entry_key * string -> ('self, 'a list) entry_key | Aopt : ('self, 'a) entry_key -> ('self, 'a option) entry_key | Amodifiers : ('self, 'a) entry_key -> ('self, 'a list) entry_key | Aself : ('self, 'self) entry_key | Anext : ('self, 'self) entry_key | Aentry : 'a Entry.t -> ('self, 'a) entry_key | Aentryl : 'a Entry.t * int -> ('self, 'a) entry_key (** Add one extension at some camlp4 position of some camlp4 entry *) val unsafe_grammar_extend : grammar_object Gram.entry -> gram_reinit option (** for reinitialization if ever needed *) -> Gram.extend_statment -> unit val grammar_extend : 'a Gram.entry -> gram_reinit option (** for reinitialization if ever needed *) -> 'a Extend.extend_statment -> unit (** Remove the last n extensions *) val remove_grammars : int -> unit (** Temporary activate camlp4 verbosity *) val camlp4_verbosity : bool -> ('a -> unit) -> 'a -> unit (** Parse a string *) val parse_string : 'a Gram.entry -> string -> 'a val eoi_entry : 'a Gram.entry -> 'a Gram.entry val map_entry : ('a -> 'b) -> 'a Gram.entry -> 'b Gram.entry type gram_universe = Entry.universe val get_univ : string -> gram_universe val uprim : gram_universe val uconstr : gram_universe val utactic : gram_universe val uvernac : gram_universe val create_generic_entry : string -> ('a, rlevel) abstract_argument_type -> 'a Gram.entry module Prim : sig open Names open Libnames val preident : string Gram.entry val ident : Id.t Gram.entry val name : Name.t located Gram.entry val identref : Id.t located Gram.entry val pattern_ident : Id.t Gram.entry val pattern_identref : Id.t located Gram.entry val base_ident : Id.t Gram.entry val natural : int Gram.entry val bigint : Bigint.bigint Gram.entry val integer : int Gram.entry val string : string Gram.entry val qualid : qualid located Gram.entry val fullyqualid : Id.t list located Gram.entry val reference : reference Gram.entry val by_notation : (Loc.t * string * string option) Gram.entry val smart_global : reference or_by_notation Gram.entry val dirpath : DirPath.t Gram.entry val ne_string : string Gram.entry val ne_lstring : string located Gram.entry val var : Id.t located Gram.entry end module Constr : sig val constr : constr_expr Gram.entry val constr_eoi : constr_expr Gram.entry val lconstr : constr_expr Gram.entry val binder_constr : constr_expr Gram.entry val operconstr : constr_expr Gram.entry val ident : Id.t Gram.entry val global : reference Gram.entry val sort : glob_sort Gram.entry val pattern : cases_pattern_expr Gram.entry val constr_pattern : constr_expr Gram.entry val lconstr_pattern : constr_expr Gram.entry val closed_binder : local_binder list Gram.entry val binder : local_binder list Gram.entry (* closed_binder or variable *) val binders : local_binder list Gram.entry (* list of binder *) val open_binders : local_binder list Gram.entry val binders_fixannot : (local_binder list * (Id.t located option * recursion_order_expr)) Gram.entry val typeclass_constraint : (Name.t located * bool * constr_expr) Gram.entry val record_declaration : constr_expr Gram.entry val appl_arg : (constr_expr * explicitation located option) Gram.entry end module Module : sig val module_expr : module_ast Gram.entry val module_type : module_ast Gram.entry end module Tactic : sig val open_constr : open_constr_expr Gram.entry val constr_with_bindings : constr_expr with_bindings Gram.entry val bindings : constr_expr bindings Gram.entry val hypident : (Id.t located * Locus.hyp_location_flag) Gram.entry val constr_may_eval : (constr_expr,reference or_by_notation,constr_expr) may_eval Gram.entry val constr_eval : (constr_expr,reference or_by_notation,constr_expr) may_eval Gram.entry val uconstr : constr_expr Gram.entry val quantified_hypothesis : quantified_hypothesis Gram.entry val int_or_var : int or_var Gram.entry val red_expr : raw_red_expr Gram.entry val simple_tactic : raw_tactic_expr Gram.entry val simple_intropattern : constr_expr intro_pattern_expr located Gram.entry val clause_dft_concl : Names.Id.t Loc.located Locus.clause_expr Gram.entry val tactic_arg : raw_tactic_arg Gram.entry val tactic_expr : raw_tactic_expr Gram.entry val binder_tactic : raw_tactic_expr Gram.entry val tactic : raw_tactic_expr Gram.entry val tactic_eoi : raw_tactic_expr Gram.entry val tacdef_body : (reference * bool * raw_tactic_expr) Gram.entry end module Vernac_ : sig val gallina : vernac_expr Gram.entry val gallina_ext : vernac_expr Gram.entry val command : vernac_expr Gram.entry val syntax : vernac_expr Gram.entry val vernac : vernac_expr Gram.entry val rec_definition : (fixpoint_expr * decl_notation list) Gram.entry val vernac_eoi : vernac_expr Gram.entry end (** The main entry: reads an optional vernac command *) val main_entry : (Loc.t * vernac_expr) option Gram.entry (** Mapping formal entries into concrete ones *) (** Binding constr entry keys to entries and symbols *) val interp_constr_entry_key : bool (** true for cases_pattern *) -> constr_entry_key -> grammar_object Gram.entry * int option val symbol_of_constr_prod_entry_key : gram_assoc option -> constr_entry_key -> bool -> constr_prod_entry_key -> Gram.symbol val name_of_entry : 'a Gram.entry -> 'a Entry.t (** Binding general entry keys to symbols *) val symbol_of_prod_entry_key : ('self, 'a) entry_key -> Gram.symbol type 's entry_name = EntryName : 'a raw_abstract_argument_type * ('s, 'a) entry_key -> 's entry_name (** Interpret entry names of the form "ne_constr_list" as entry keys *) type _ target = TgAny : 's target | TgTactic : int -> raw_tactic_expr target val interp_entry_name : bool (** true to fail on unknown entry *) -> 's target -> string -> string -> 's entry_name (** Recover the list of all known tactic notation entries. *) val list_entry_names : unit -> (string * argument_type) list (** Registering/resetting the level of a constr entry *) val find_position : bool (** true if for creation in pattern entry; false if in constr entry *) -> Extend.gram_assoc option -> int option -> Extend.gram_position option * Extend.gram_assoc option * string option * (** for reinitialization: *) gram_reinit option val synchronize_level_positions : unit -> unit val register_empty_levels : bool -> int list -> (Extend.gram_position option * Extend.gram_assoc option * string option * gram_reinit option) list val remove_levels : int -> unit val level_of_snterml : Gram.symbol -> int