package coq

  1. Overview
  2. Docs
Legend:
Library
Module
Module type
Parameter
Class
Class type
type !'a puniverses = 'a Univ.puniverses
type pconstant = Names.constant puniverses
type pinductive = Names.inductive puniverses
type pconstructor = Names.constructor puniverses
type existential_key = Evar.t
type metavariable = int
type case_style =
  1. | LetStyle
  2. | IfStyle
  3. | LetPatternStyle
  4. | MatchStyle
  5. | RegularStyle
type case_printing = {
  1. ind_tags : bool list;
  2. cstr_tags : bool list array;
  3. style : case_style;
}
type case_info = {
  1. ci_ind : Names.inductive;
  2. ci_npar : int;
  3. ci_cstr_ndecls : int array;
  4. ci_cstr_nargs : int array;
  5. ci_pp_info : case_printing;
}
type t
type constr = t
type types = constr
val mkRel : int -> constr
val mkVar : Names.Id.t -> constr
val mkMeta : metavariable -> constr
type existential = existential_key * constr array
val mkEvar : existential -> constr
val mkSort : Sorts.t -> types
val mkProp : types
val mkSet : types
val mkType : Univ.universe -> types
type cast_kind =
  1. | VMcast
  2. | NATIVEcast
  3. | DEFAULTcast
  4. | REVERTcast
val mkCast : (constr * cast_kind * constr) -> constr
val mkProd : (Names.Name.t * types * types) -> types
val mkLambda : (Names.Name.t * types * constr) -> constr
val mkLetIn : (Names.Name.t * constr * types * constr) -> constr
val mkApp : (constr * constr array) -> constr
val map_puniverses : ('a -> 'b) -> 'a puniverses -> 'b puniverses
val mkConst : Names.constant -> constr
val mkConstU : pconstant -> constr
val mkProj : (Names.projection * constr) -> constr
val mkInd : Names.inductive -> constr
val mkIndU : pinductive -> constr
val mkConstruct : Names.constructor -> constr
val mkConstructU : pconstructor -> constr
val mkConstructUi : (pinductive * int) -> constr
val mkCase : (case_info * constr * constr * constr array) -> constr
type rec_declaration = Names.Name.t array * types array * constr array
type fixpoint = (int array * int) * rec_declaration
val mkFix : fixpoint -> constr
type cofixpoint = int * rec_declaration
val mkCoFix : cofixpoint -> constr
type !'constr pexistential = existential_key * 'constr array
type (!'constr, !'types) prec_declaration = Names.Name.t array * 'types array * 'constr array
type (!'constr, !'types) pfixpoint = (int array * int) * ('constr, 'types) prec_declaration
type (!'constr, !'types) pcofixpoint = int * ('constr, 'types) prec_declaration
type (!'constr, !'types, !'sort, !'univs) kind_of_term =
  1. | Rel of int
  2. | Var of Names.Id.t
  3. | Meta of metavariable
  4. | Evar of 'constr pexistential
  5. | Sort of 'sort
  6. | Cast of 'constr * cast_kind * 'types
  7. | Prod of Names.Name.t * 'types * 'types
  8. | Lambda of Names.Name.t * 'types * 'constr
  9. | LetIn of Names.Name.t * 'constr * 'types * 'constr
  10. | App of 'constr * 'constr array
  11. | Const of Names.constant * 'univs
  12. | Ind of Names.inductive * 'univs
  13. | Construct of Names.constructor * 'univs
  14. | Case of case_info * 'constr * 'constr * 'constr array
  15. | Fix of ('constr, 'types) pfixpoint
  16. | CoFix of ('constr, 'types) pcofixpoint
  17. | Proj of Names.projection * 'constr
val equal : constr -> constr -> bool
val eq_constr_univs : constr UGraph.check_function
val leq_constr_univs : constr UGraph.check_function
val eq_constr_univs_infer : UGraph.t -> constr -> constr -> bool Univ.constrained
val leq_constr_univs_infer : UGraph.t -> constr -> constr -> bool Univ.constrained
val eq_constr_nounivs : constr -> constr -> bool
val compare : constr -> constr -> int
val fold : ('a -> constr -> 'a) -> 'a -> constr -> 'a
val map : (constr -> constr) -> constr -> constr
val fold_map : ('a -> constr -> 'a * constr) -> 'a -> constr -> 'a * constr
val map_with_binders : ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr
val iter : (constr -> unit) -> constr -> unit
val iter_with_binders : ('a -> 'a) -> ('a -> constr -> unit) -> 'a -> constr -> unit
val compare_head : (constr -> constr -> bool) -> constr -> constr -> bool
val compare_head_gen : (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) -> (Sorts.t -> Sorts.t -> bool) -> (constr -> constr -> bool) -> constr -> constr -> bool
val compare_head_gen_leq_with : (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) -> (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) -> (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) -> (Sorts.t -> Sorts.t -> bool) -> (constr -> constr -> bool) -> (constr -> constr -> bool) -> constr -> constr -> bool
val compare_head_gen_with : (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) -> (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) -> (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) -> (Sorts.t -> Sorts.t -> bool) -> (constr -> constr -> bool) -> constr -> constr -> bool
val compare_head_gen_leq : (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) -> (Sorts.t -> Sorts.t -> bool) -> (constr -> constr -> bool) -> (constr -> constr -> bool) -> constr -> constr -> bool
val hash : constr -> int
val case_info_hash : case_info -> int
val hcons : constr -> constr
type values
OCaml

Innovation. Community. Security.