Module Nativelambda

type prefix = string

This file defines the lambda code generation phase of the native compiler

type lambda =
| Lrel of Names.Name.t * int
| Lvar of Names.Id.t
| Lmeta of Constr.metavariable * lambda
| Levar of Evar.t * lambda array
| Lprod of lambda * lambda
| Llam of Names.Name.t Context.binder_annot array * lambda
| Lrec of Names.Name.t Context.binder_annot * lambda
| Llet of Names.Name.t Context.binder_annot * lambda * lambda
| Lapp of lambda * lambda array
| Lconst of prefix * Constr.pconstant
| Lproj of prefix * Names.inductive * int
| Lprim of prefix * Constr.pconstant * CPrimitives.t * lambda array
| Lcase of Nativevalues.annot_sw * lambda * lambda * lam_branches
| Lif of lambda * lambda * lambda
| Lfix of int array * (string * Names.inductive) array * int * fix_decl
| Lcofix of int * fix_decl
| Lint of int
| Lmakeblock of prefix * Names.inductive * int * lambda array
| Luint of Uint63.t
| Lfloat of Float64.t
| Lval of Nativevalues.t
| Lsort of Sorts.t
| Lind of prefix * Constr.pinductive
| Llazy
| Lforce
and lam_branches = {
constant_branches : lambda array;
nonconstant_branches : (Names.Name.t Context.binder_annot array * lambda) array;
}
and fix_decl = Names.Name.t Context.binder_annot array * lambda array * lambda array
type evars = {
evars_val : Constr.existential -> Constr.constr option;
evars_metas : Constr.metavariable -> Constr.types;
}
val empty_evars : evars
val decompose_Llam : lambda -> Names.Name.t Context.binder_annot array * lambda
val decompose_Llam_Llet : lambda -> (Names.Name.t Context.binder_annot * lambda option) array * lambda
val is_lazy : Constr.constr -> bool
val mk_lazy : lambda -> lambda
val get_mind_prefix : Environ.env -> Names.MutInd.t -> string
val get_alias : Environ.env -> Constr.pconstant -> Constr.pconstant
val lambda_of_constr : Environ.env -> evars -> Constr.constr -> lambda