Module Type_errors
type 'constr pfix_guard_error
=
|
NotEnoughAbstractionInFixBody
|
RecursionNotOnInductiveType of 'constr
|
RecursionOnIllegalTerm of int * Environ.env * 'constr * (int list * int list) Stdlib.Lazy.t
|
NotEnoughArgumentsForFixCall of int
|
FixpointOnIrrelevantInductive
type 'constr pcofix_guard_error
=
type 'constr pguard_error
=
|
FixGuardError of 'constr pfix_guard_error
|
CoFixGuardError of 'constr pcofix_guard_error
type fix_guard_error
= Constr.constr pfix_guard_error
type cofix_guard_error
= Constr.constr pcofix_guard_error
type guard_error
= Constr.constr pguard_error
type ('constr, 'types) pcant_apply_bad_type
= (int * 'constr * 'constr) * ('constr, 'types) Environ.punsafe_judgment * ('constr, 'types) Environ.punsafe_judgment array
type ('constr, 'types) ptype_error
=
|
UnboundRel of int
|
UnboundVar of Names.variable
|
NotAType of ('constr, 'types) Environ.punsafe_judgment
|
BadAssumption of ('constr, 'types) Environ.punsafe_judgment
|
ReferenceVariables of Names.Id.t * Names.GlobRef.t
|
ElimArity of Constr.pinductive * 'constr * (('constr, 'types) Environ.punsafe_judgment * Sorts.t) option
|
CaseNotInductive of ('constr, 'types) Environ.punsafe_judgment
|
CaseOnPrivateInd of Names.inductive
|
WrongCaseInfo of Constr.pinductive * Constr.case_info
|
NumberBranches of ('constr, 'types) Environ.punsafe_judgment * int
|
IllFormedCaseParams
|
IllFormedBranch of 'constr * Constr.pconstructor * 'constr * 'constr
|
Generalization of Names.Name.t * 'types * ('constr, 'types) Environ.punsafe_judgment
|
ActualType of ('constr, 'types) Environ.punsafe_judgment * 'types
|
IncorrectPrimitive of (CPrimitives.op_or_type, 'types) Environ.punsafe_judgment * 'types
|
CantApplyBadType of ('constr, 'types) pcant_apply_bad_type
|
CantApplyNonFunctional of ('constr, 'types) Environ.punsafe_judgment * ('constr, 'types) Environ.punsafe_judgment array
|
IllFormedRecBody of 'constr pguard_error * Names.Name.t Context.binder_annot array * int * Environ.env * ('constr, 'types) Environ.punsafe_judgment array
|
IllTypedRecBody of int * Names.Name.t Context.binder_annot array * ('constr, 'types) Environ.punsafe_judgment array * 'types array
|
UnsatisfiedQConstraints of Sorts.QConstraints.t
|
UnsatisfiedConstraints of Univ.Constraints.t
|
UndeclaredQualities of Sorts.QVar.Set.t
|
UndeclaredUniverse of Univ.Level.t
|
DisallowedSProp
|
BadBinderRelevance of Sorts.relevance * ('constr, 'types) Context.Rel.Declaration.pt
|
BadCaseRelevance of Sorts.relevance * 'constr
|
BadInvert
|
BadVariance of
{
lev : Univ.Level.t;
expected : UVars.Variance.t;
actual : UVars.Variance.t;
}
|
UndeclaredUsedVariables of
{
declared_vars : Names.Id.Set.t;
inferred_vars : Names.Id.Set.t;
}
type type_error
= (Constr.constr, Constr.types) ptype_error
exception
TypeError of Environ.env * type_error
type inductive_error
=
|
NonPos of Environ.env * Constr.constr * Constr.constr
|
NotEnoughArgs of Environ.env * Constr.constr * Constr.constr
|
NotConstructor of Environ.env * Names.Id.t * Constr.constr * Constr.constr * int * int
|
NonPar of Environ.env * Constr.constr * int * Constr.constr * Constr.constr
|
SameNamesTypes of Names.Id.t
|
SameNamesConstructors of Names.Id.t
|
SameNamesOverlap of Names.Id.t list
|
NotAnArity of Environ.env * Constr.constr
|
BadEntry
|
LargeNonPropInductiveNotInType
|
MissingConstraints of Sorts.t list * Sorts.t
The different kinds of errors that may result of a malformed inductive definition.
exception
InductiveError of inductive_error
val error_unbound_rel : Environ.env -> int -> 'a
val error_unbound_var : Environ.env -> Names.variable -> 'a
val error_not_type : Environ.env -> Environ.unsafe_judgment -> 'a
val error_assumption : Environ.env -> Environ.unsafe_judgment -> 'a
val error_reference_variables : Environ.env -> Names.Id.t -> Names.GlobRef.t -> 'a
val error_elim_arity : Environ.env -> Constr.pinductive -> Constr.constr -> (Environ.unsafe_judgment * Sorts.t) option -> 'a
val error_case_not_inductive : Environ.env -> Environ.unsafe_judgment -> 'a
val error_case_on_private_ind : Environ.env -> Names.inductive -> 'a
val error_number_branches : Environ.env -> Environ.unsafe_judgment -> int -> 'a
val error_ill_formed_branch : Environ.env -> Constr.constr -> Constr.pconstructor -> Constr.constr -> Constr.constr -> 'a
val error_generalization : Environ.env -> (Names.Name.t * Constr.types) -> Environ.unsafe_judgment -> 'a
val error_actual_type : Environ.env -> Environ.unsafe_judgment -> Constr.types -> 'a
val error_incorrect_primitive : Environ.env -> (CPrimitives.op_or_type, Constr.types) Environ.punsafe_judgment -> Constr.types -> 'a
val error_cant_apply_not_functional : Environ.env -> Environ.unsafe_judgment -> Environ.unsafe_judgment array -> 'a
val error_cant_apply_bad_type : Environ.env -> (int * Constr.constr * Constr.constr) -> Environ.unsafe_judgment -> Environ.unsafe_judgment array -> 'a
val error_ill_formed_rec_body : Environ.env -> guard_error -> Names.Name.t Context.binder_annot array -> int -> Environ.env -> Environ.unsafe_judgment array -> 'a
val error_ill_typed_rec_body : Environ.env -> int -> Names.Name.t Context.binder_annot array -> Environ.unsafe_judgment array -> Constr.types array -> 'a
val error_unsatisfied_qconstraints : Environ.env -> Sorts.QConstraints.t -> 'a
val error_unsatisfied_constraints : Environ.env -> Univ.Constraints.t -> 'a
val error_undeclared_qualities : Environ.env -> Sorts.QVar.Set.t -> 'a
val error_undeclared_universe : Environ.env -> Univ.Level.t -> 'a
val error_disallowed_sprop : Environ.env -> 'a
val error_bad_binder_relevance : Environ.env -> Sorts.relevance -> Constr.rel_declaration -> 'a
val error_bad_case_relevance : Environ.env -> Sorts.relevance -> Constr.case -> 'a
val error_bad_invert : Environ.env -> 'a
val error_bad_variance : Environ.env -> lev:Univ.Level.t -> expected:UVars.Variance.t -> actual:UVars.Variance.t -> 'a
val error_undeclared_used_variables : Environ.env -> declared_vars:Names.Id.Set.t -> inferred_vars:Names.Id.Set.t -> 'a
val map_pguard_error : ('c -> 'd) -> 'c pguard_error -> 'd pguard_error
val map_ptype_error : ('c -> 'd) -> ('c, 'c) ptype_error -> ('d, 'd) ptype_error