Module ComInductive
Inductive and coinductive types
val do_mutual_inductive : template:bool option -> Constrexpr.cumul_univ_decl_expr option -> (Vernacexpr.one_inductive_expr * Vernacexpr.decl_notation list) list -> cumulative:bool -> poly:bool -> ?typing_flags:Declarations.typing_flags -> private_ind:bool -> uniform:uniform_inductive_flag -> Declarations.recursivity_kind -> unit
val make_cases : Names.inductive -> string list listval interp_mutual_inductive_constr : sigma:Evd.evar_map -> template:bool option -> udecl:UState.universe_decl -> variances:Entries.variance_entry -> ctx_params:(EConstr.t, EConstr.t) Context.Rel.Declaration.pt list -> indnames:Names.Id.t list -> arities:EConstr.t list -> arityconcl:(bool * EConstr.ESorts.t) option list -> constructors:(Names.Id.t list * Constr.constr list) list -> env_ar_params:Environ.env -> cumulative:bool -> poly:bool -> private_ind:bool -> finite:Declarations.recursivity_kind -> Entries.mutual_inductive_entry * UnivNames.universe_binders
val should_auto_template : Names.Id.t -> bool -> boolshould_auto_template x bistruewhenbistrueand we automatically use template polymorphism.xis the name of the inductive under consideration.
val template_polymorphism_candidate : ctor_levels:Univ.LSet.t -> Entries.universes_entry -> Constr.rel_context -> Sorts.t option -> booltemplate_polymorphism_candidate ~ctor_levels uctx params conclsortistrueiff an inductive with paramsparams, conclusionconclsortand universe levels appearing in the constructor argumentsctor_levelswould be definable as template polymorphic. It should have at least one universe in its monomorphic universe context that can be made parametric in its conclusion sort, if one is given.
val maybe_unify_params_in : Environ.env -> Evd.evar_map -> ninds:int -> nparams:int -> binders:int -> EConstr.t -> Evd.evar_mapnparamsis the number of parameters which aren't treated as uniform, ie the length of params (including letins) where the env isuniform params, inductives, params, binders.
val variance_of_entry : cumulative:bool -> variances:Entries.variance_entry -> Entries.universes_entry -> Entries.variance_entry optionWill return None if non-cumulative, and resize if there are more universes than originally specified. If monomorphic,
cumulativeis treated asfalse.