Notation_ops
notation_constr
val eq_notation_constr : (Names.Id.t list * Names.Id.t list) -> Notation_term.notation_constr -> Notation_term.notation_constr -> bool
val strictly_finer_notation_constr : (Names.Id.t list * Names.Id.t list) -> Notation_term.notation_constr -> Notation_term.notation_constr -> bool
Tell if t1
is a strict refinement of t2
(this is a partial order and returning false
does not mean that t2
is finer than t1
)
val strictly_finer_interpretation_than : Notation_term.interpretation -> Notation_term.interpretation -> bool
Tell if a notation interpretation is strictly included in another one
val finer_interpretation_than : Notation_term.interpretation -> Notation_term.interpretation -> bool
Tell if a notation interpretation is included in another one
Substitution of kernel names in interpretation data
val subst_interpretation : Mod_subst.substitution -> Notation_term.interpretation -> Notation_term.interpretation
Name of the special identifier used to encode recursive notations
val ldots_var : Names.Id.t
glob_constr
and notation_constr
Translate a glob_constr
into a notation given the list of variables bound by the notation; also interpret recursive patterns
val notation_constr_of_glob_constr : Notation_term.notation_interp_env -> Glob_term.glob_constr -> Notation_term.notation_constr * Notation_term.reversibility_status
Re-interpret a notation as a glob_constr
, taking care of binders
val apply_cases_pattern : ?loc:Loc.t -> ((Names.Id.t list * Glob_term.cases_pattern_disjunction) * Names.Id.t) -> Glob_term.glob_constr -> Glob_term.glob_constr
val glob_constr_of_notation_constr_with_binders : ?loc:Loc.t ->
('a -> Names.Name.t -> Glob_term.glob_constr option -> 'a * ((Names.Id.t list * Glob_term.cases_pattern_disjunction) * Names.Id.t) option * Names.Name.t * Glob_term.binding_kind * Glob_term.glob_constr option) -> ('a -> Notation_term.notation_constr -> Glob_term.glob_constr) -> ?h:'a binder_status_fun
-> 'a -> Notation_term.notation_constr -> Glob_term.glob_constr
val glob_constr_of_notation_constr : ?loc:Loc.t -> Notation_term.notation_constr -> Glob_term.glob_constr
glob_constr
match_notation_constr
matches a glob_constr
against a notation interpretation; raise No_match
if the matching fails
val match_notation_constr : print_univ:bool -> 'a Glob_term.glob_constr_g -> vars:Names.Id.Set.t ->
Notation_term.interpretation -> ((Names.Id.Set.t * 'a Glob_term.glob_constr_g) * Notation_term.extended_subscopes) list * ((Names.Id.Set.t * 'a Glob_term.glob_constr_g list) * Notation_term.extended_subscopes) list * ((Names.Id.Set.t * 'a Glob_term.cases_pattern_disjunction_g) * Notation_term.extended_subscopes) list * ((Names.Id.Set.t * 'a Glob_term.extended_glob_local_binder_g list) * Notation_term.extended_subscopes) list
val match_notation_constr_cases_pattern : 'a Glob_term.cases_pattern_g -> Notation_term.interpretation -> (('a Glob_term.cases_pattern_g * Notation_term.extended_subscopes) list * ('a Glob_term.cases_pattern_g list * Notation_term.extended_subscopes) list) * (bool * int * 'a Glob_term.cases_pattern_g list)
val match_notation_constr_ind_pattern : Names.inductive -> 'a Glob_term.cases_pattern_g list -> Notation_term.interpretation -> (('a Glob_term.cases_pattern_g * Notation_term.extended_subscopes) list * ('a Glob_term.cases_pattern_g list * Notation_term.extended_subscopes) list) * (bool * int * 'a Glob_term.cases_pattern_g list)
glob_constr