package coq
Install
Dune Dependency
Authors
Maintainers
Sources
sha256=b1501d686c21836302191ae30f610cca57fb309214c126518ca009363ad2cd3c
doc/coq-core.engine/Termops/index.html
Module Termops
Source
This file defines various utilities for term manipulation that are not needed in the kernel.
val push_rel_assum :
(Names.Name.t Context.binder_annot * EConstr.types) ->
Environ.env ->
Environ.env
about contexts
val push_rels_assum :
(Names.Name.t Context.binder_annot * Constr.types) list ->
Environ.env ->
Environ.env
val push_named_rec_types :
(Names.Name.t Context.binder_annot array * Constr.types array * 'a) ->
Environ.env ->
Environ.env
Associates the contents of an identifier in a rel_context
. Raise Not_found
if there is no such identifier.
Functions that build argument lists matching a block of binders or a context. rel_vect n m
builds |Rel (n+m);...;Rel(n+1)|
iterators/destructors on terms
val it_mkProd :
EConstr.types ->
(Names.Name.t Context.binder_annot * EConstr.types) list ->
EConstr.types
val it_mkLambda :
EConstr.constr ->
(Names.Name.t Context.binder_annot * EConstr.types) list ->
EConstr.constr
Generic iterators on constr
val map_constr_with_binders_left_to_right :
Environ.env ->
Evd.evar_map ->
(EConstr.rel_declaration -> 'a -> 'a) ->
('a -> EConstr.constr -> EConstr.constr) ->
'a ->
EConstr.constr ->
EConstr.constr
val map_constr_with_full_binders :
Environ.env ->
Evd.evar_map ->
(EConstr.rel_declaration -> 'a -> 'a) ->
('a -> EConstr.constr -> EConstr.constr) ->
'a ->
EConstr.constr ->
EConstr.constr
fold_constr_with_binders g f n acc c
folds f n
on the immediate subterms of c
starting from acc
and proceeding from left to right according to the usual representation of the constructions as fold_constr
but it carries an extra data n
(typically a lift index) which is processed by g
(which typically add 1 to n
) at each binder traversal; it is not recursive
val fold_constr_with_binders :
Evd.evar_map ->
('a -> 'a) ->
('a -> 'b -> EConstr.constr -> 'b) ->
'a ->
'b ->
EConstr.constr ->
'b
val fold_constr_with_full_binders :
Environ.env ->
Evd.evar_map ->
(EConstr.rel_declaration -> 'a -> 'a) ->
('a -> 'b -> EConstr.constr -> 'b) ->
'a ->
'b ->
EConstr.constr ->
'b
occur checks
val occur_var_indirectly :
Environ.env ->
Evd.evar_map ->
Names.Id.t ->
EConstr.constr ->
Names.GlobRef.t option
val occur_var_in_decl :
Environ.env ->
Evd.evar_map ->
Names.Id.t ->
EConstr.named_declaration ->
bool
val occur_vars_in_decl :
Environ.env ->
Evd.evar_map ->
Names.Id.Set.t ->
EConstr.named_declaration ->
bool
As occur_var
but assume the identifier not to be a section variable
dependent m t
tests whether m
is a subterm of t
for visible vars only
Type assignment for metavariables
pop c
lifts by -1 the positive indexes in c
...
Substitution of an arbitrary large term. Uses equality modulo reduction of let
val replace_term_gen :
Evd.evar_map ->
(Evd.evar_map -> int -> EConstr.constr -> bool) ->
int ->
EConstr.constr ->
EConstr.constr ->
EConstr.constr
replace_term_gen eq arity e c
replaces matching subterms according to eq
by e
in c
. If arity
is non-zero applications of larger length are handled atomically.
subst_term d c
replaces d
by Rel 1
in c
val replace_term :
Evd.evar_map ->
EConstr.constr ->
EConstr.constr ->
EConstr.constr ->
EConstr.constr
replace_term d e c
replaces d
by e
in c
Alternative term equalities
val compare_constr_univ :
Evd.evar_map ->
(Reduction.conv_pb -> EConstr.constr -> EConstr.constr -> bool) ->
Reduction.conv_pb ->
EConstr.constr ->
EConstr.constr ->
bool
val constr_cmp :
Evd.evar_map ->
Reduction.conv_pb ->
EConstr.constr ->
EConstr.constr ->
bool
Flattens application lists
prod_applist
forall (x1:B1;...;xn:Bn), B
a1...an
val prod_applist_assum :
Evd.evar_map ->
int ->
EConstr.constr ->
EConstr.constr list ->
EConstr.constr
In prod_applist_assum n c args
, c
is supposed to have the form ∀Γ.c
with Γ
of length m
and possibly with let-ins; it returns c
with the assumptions of Γ
instantiated by args
and the local definitions of Γ
expanded. Note that n
counts both let-ins and prods, while the length of args
only counts prods. In other words, varying n
changes how many trailing let-ins are expanded.
Remove recursively the casts around a term i.e. strip_outer_cast (Cast (Cast ... (Cast c, t) ... ))
is c
.
Lightweight first-order filtering procedure. Unification variables ar represented by (untyped) Evars. filtering c1 c2
returns the substitution n'th evar -> (context,term), or raises CannotFilter
. Warning: Outer-kernel sort subtyping are taken into account: c1 has to be smaller than c2 wrt. sorts.
val filtering :
Evd.evar_map ->
EConstr.rel_context ->
Reduction.conv_pb ->
EConstr.constr ->
EConstr.constr ->
subst
val decompose_prod_letin :
Evd.evar_map ->
EConstr.constr ->
int * EConstr.rel_context * EConstr.constr
val align_prod_letin :
Evd.evar_map ->
EConstr.constr ->
EConstr.constr ->
EConstr.rel_context * EConstr.constr
nb_lam
$
x_1:T_1
...x_n:T_n
c $
where $
c $
is not an abstraction gives $
n $
(casts are ignored)
Similar to nb_lam
, but gives the number of products instead
Similar to nb_prod
, but zeta-contracts let-in on the way
Get the last arg of a constr intended to be an application
val decompose_app_vect :
Evd.evar_map ->
EConstr.constr ->
EConstr.constr * EConstr.constr array
Force the decomposition of a term as an applicative one
val adjust_app_list_size :
EConstr.constr ->
EConstr.constr list ->
EConstr.constr ->
EConstr.constr list ->
EConstr.constr * EConstr.constr list * EConstr.constr * EConstr.constr list
val adjust_app_array_size :
EConstr.constr ->
EConstr.constr array ->
EConstr.constr ->
EConstr.constr array ->
EConstr.constr * EConstr.constr array * EConstr.constr * EConstr.constr array
name contexts
Set of local names
val process_rel_context :
(EConstr.rel_declaration -> Environ.env -> Environ.env) ->
Environ.env ->
Environ.env
other signature iterators
val assums_of_rel_context :
('c, 't) Context.Rel.pt ->
(Names.Name.t Context.binder_annot * 't) list
expand lets in context
val map_rel_context_in_env :
(Environ.env -> Constr.constr -> Constr.constr) ->
Environ.env ->
Constr.rel_context ->
Constr.rel_context
val map_rel_context_with_binders :
(int -> 'c -> 'c) ->
('c, 'c) Context.Rel.pt ->
('c, 'c) Context.Rel.pt
val fold_named_context_both_sides :
('a -> Constr.named_declaration -> Constr.named_declaration list -> 'a) ->
Constr.named_context ->
init:'a ->
'a
val map_rel_decl :
('a -> 'b) ->
('a, 'a) Context.Rel.Declaration.pt ->
('b, 'b) Context.Rel.Declaration.pt
val map_named_decl :
('a -> 'b) ->
('a, 'a) Context.Named.Declaration.pt ->
('b, 'b) Context.Named.Declaration.pt
val global_vars_set_of_decl :
Environ.env ->
Evd.evar_map ->
EConstr.named_declaration ->
Names.Id.Set.t
val global_app_of_constr :
Evd.evar_map ->
EConstr.constr ->
(Names.GlobRef.t * EConstr.EInstance.t) * EConstr.constr option
val dependency_closure :
Environ.env ->
Evd.evar_map ->
EConstr.named_context ->
Names.Id.Set.t ->
Names.Id.t list
Gives an ordered list of hypotheses, closed by dependencies, containing a given set
Test if an identifier is the basename of a global reference
val global_of_constr :
Evd.evar_map ->
EConstr.constr ->
Names.GlobRef.t * EConstr.EInstance.t
Debug pretty-printers
val pr_evar_map_filter :
?with_univs:bool ->
(Evar.t -> Evd.evar_info -> bool) ->
Environ.env ->
Evd.evar_map ->
Pp.t