package coq
sectionYPositions = computeSectionYPositions($el), 10)"
x-init="setTimeout(() => sectionYPositions = computeSectionYPositions($el), 10)"
>
Formal proof management system
Install
dune-project
Dependency
Authors
Maintainers
Sources
coq-8.16.1.tar.gz
sha256=583471c8ed4f227cb374ee8a13a769c46579313d407db67a82d202ee48300e4b
doc/coq-core.kernel/Reduction/index.html
Module ReductionSource
Source
type 'a extended_conversion_function =
?l2r:bool ->
?reds:TransparentState.t ->
Environ.env ->
?evars:(Constr.existential -> Constr.constr option) ->
'a ->
'a ->
unitSource
type 'a universe_compare = {compare_sorts : Environ.env -> conv_pb -> Sorts.t -> Sorts.t -> 'a -> 'a;compare_instances : flex:bool -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a;compare_cumul_instances : conv_pb -> Univ.Variance.t array -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a;
}Source
type ('a, 'b) generic_conversion_function =
Environ.env ->
'b universe_state ->
'a ->
'a ->
'bSource
val get_cumulativity_constraints :
conv_pb ->
Univ.Variance.t array ->
Univ.Instance.t ->
Univ.Instance.t ->
Univ.Constraints.tSource
val constructor_cumulativity_arguments :
(Declarations.mutual_inductive_body * int * int) ->
intSource
val sort_cmp_universes :
Environ.env ->
conv_pb ->
Sorts.t ->
Sorts.t ->
('a * 'a universe_compare) ->
'a * 'a universe_compareSource
val convert_instances :
flex:bool ->
Univ.Instance.t ->
Univ.Instance.t ->
('a * 'a universe_compare) ->
'a * 'a universe_compareThis function never raise UnivInconsistency.
These two functions can only raise NotConvertible
Source
val generic_conv :
conv_pb ->
l2r:bool ->
(Constr.existential -> Constr.constr option) ->
TransparentState.t ->
(Constr.constr, 'a) generic_conversion_functionDepending on the universe state functions, this might raise UniverseInconsistency in addition to NotConvertible (for better error messages).
Builds an application node, reducing beta redexes it may produce.
Builds an application node, reducing beta redexes it may produce.
Builds an application node, reducing beta redexe it may produce.
Pseudo-reduction rule Prod(x,A,B) a --> Bx\a
Source
val hnf_prod_applist_assum :
Environ.env ->
int ->
Constr.types ->
Constr.constr list ->
Constr.typesIn hnf_prod_applist_assum n c args, c is supposed to (whd-)reduce to the form ∀Γ.t with Γ of length n and possibly with let-ins; it returns t with the assumptions of Γ instantiated by args and the local definitions of Γ expanded.
Compatibility alias for Term.lambda_appvect_assum
Source
val dest_lam_n_assum :
Environ.env ->
int ->
Constr.constr ->
Constr.rel_context * Constr.constr sectionYPositions = computeSectionYPositions($el), 10)"
x-init="setTimeout(() => sectionYPositions = computeSectionYPositions($el), 10)"
>