aboutsummaryrefslogtreecommitdiffhomepage
path: root/API/API.mli
diff options
context:
space:
mode:
Diffstat (limited to 'API/API.mli')
-rw-r--r--API/API.mli300
1 files changed, 157 insertions, 143 deletions
diff --git a/API/API.mli b/API/API.mli
index e320e496c..8f46a5832 100644
--- a/API/API.mli
+++ b/API/API.mli
@@ -1293,65 +1293,6 @@ sig
type to_patch_substituted
end
-module Decl_kinds :
-sig
- type polymorphic = bool
- type cumulative_inductive_flag = bool
- type recursivity_kind =
- | Finite
- | CoFinite
- | BiFinite
-
- type discharge =
- | DoDischarge
- | NoDischarge
-
- type locality =
- | Discharge
- | Local
- | Global
-
- type definition_object_kind =
- | Definition
- | Coercion
- | SubClass
- | CanonicalStructure
- | Example
- | Fixpoint
- | CoFixpoint
- | Scheme
- | StructureComponent
- | IdentityCoercion
- | Instance
- | Method
- | Let
- type theorem_kind =
- | Theorem
- | Lemma
- | Fact
- | Remark
- | Property
- | Proposition
- | Corollary
- type goal_object_kind =
- | DefinitionBody of definition_object_kind
- | Proof of theorem_kind
- type goal_kind = locality * polymorphic * goal_object_kind
- type assumption_object_kind =
- | Definitional
- | Logical
- | Conjectural
- type logical_kind =
- | IsAssumption of assumption_object_kind
- | IsDefinition of definition_object_kind
- | IsProof of theorem_kind
- type binding_kind =
- | Explicit
- | Implicit
- type private_flag = bool
- type definition_kind = locality * polymorphic * definition_object_kind
-end
-
module Retroknowledge :
sig
type action
@@ -1501,10 +1442,15 @@ sig
type record_body = (Id.t * Constant.t array * projection_body array) option
+ type recursivity_kind =
+ | Finite
+ | CoFinite
+ | BiFinite
+
type mutual_inductive_body = {
mind_packets : one_inductive_body array;
mind_record : record_body option;
- mind_finite : Decl_kinds.recursivity_kind;
+ mind_finite : recursivity_kind;
mind_ntypes : int;
mind_hyps : Context.Named.t;
mind_nparams : int;
@@ -1578,7 +1524,7 @@ sig
(** Some (Some id): primitive record with id the binder name of the record
in projections.
Some None: non-primitive record *)
- mind_entry_finite : Decl_kinds.recursivity_kind;
+ mind_entry_finite : Declarations.recursivity_kind;
mind_entry_params : (Id.t * local_entry) list;
mind_entry_inds : one_inductive_entry list;
mind_entry_universes : inductive_universes;
@@ -1861,6 +1807,60 @@ end
(* Modules from intf/ *)
(************************************************************************)
+module Libnames :
+sig
+
+ open Util
+ open Names
+
+ type full_path
+ val pr_path : full_path -> Pp.t
+ val make_path : Names.DirPath.t -> Names.Id.t -> full_path
+ val eq_full_path : full_path -> full_path -> bool
+ val repr_path : full_path -> Names.DirPath.t * Names.Id.t
+ val dirpath : full_path -> Names.DirPath.t
+ val path_of_string : string -> full_path
+
+ type qualid
+ val make_qualid : Names.DirPath.t -> Names.Id.t -> qualid
+ val qualid_eq : qualid -> qualid -> bool
+ val repr_qualid : qualid -> Names.DirPath.t * Names.Id.t
+ val pr_qualid : qualid -> Pp.t
+ val string_of_qualid : qualid -> string
+ val qualid_of_string : string -> qualid
+ val qualid_of_path : full_path -> qualid
+ val qualid_of_dirpath : Names.DirPath.t -> qualid
+ val qualid_of_ident : Names.Id.t -> qualid
+
+ type reference =
+ | Qualid of qualid Loc.located
+ | Ident of Names.Id.t Loc.located
+ val loc_of_reference : reference -> Loc.t option
+ val qualid_of_reference : reference -> qualid Loc.located
+ val pr_reference : reference -> Pp.t
+
+ val is_dirpath_prefix_of : Names.DirPath.t -> Names.DirPath.t -> bool
+ val split_dirpath : Names.DirPath.t -> Names.DirPath.t * Names.Id.t
+ val dirpath_of_string : string -> Names.DirPath.t
+ val pr_dirpath : Names.DirPath.t -> Pp.t
+ [@@ocaml.deprecated "Alias for DirPath.print"]
+
+ val string_of_path : full_path -> string
+
+ val basename : full_path -> Names.Id.t
+
+ type object_name = full_path * Names.KerName.t
+ type object_prefix = {
+ obj_dir : DirPath.t;
+ obj_mp : ModPath.t;
+ obj_sec : DirPath.t;
+ }
+
+ module Dirset : Set.S with type elt = DirPath.t
+ module Dirmap : Map.ExtS with type key = DirPath.t and module Set := Dirset
+ module Spmap : CSig.MapS with type key = full_path
+end
+
module Misctypes :
sig
type evars_flag = bool
@@ -1883,10 +1883,15 @@ sig
| GSet (** representation of [Set] literal *)
| GType of 'a (** representation of [Type] literal *)
- type level_info = Names.Name.t Loc.located option
+ type 'a universe_kind =
+ | UAnonymous
+ | UUnknown
+ | UNamed of 'a
+
+ type level_info = Libnames.reference universe_kind
type glob_level = level_info glob_sort_gen
- type sort_info = Names.Name.t Loc.located list
+ type sort_info = (Libnames.reference * int) option list
type glob_sort = sort_info glob_sort_gen
type ('a, 'b) gen_universe_decl = {
@@ -1998,7 +2003,7 @@ end
module Univops :
sig
- val universes_of_constr : Constr.constr -> Univ.LSet.t
+ val universes_of_constr : Environ.env -> Constr.constr -> Univ.LSet.t
val restrict_universe_context : Univ.ContextSet.t -> Univ.LSet.t -> Univ.ContextSet.t
end
@@ -2039,60 +2044,6 @@ sig
[@@ocaml.deprecated "alias of API.Names"]
end
-module Libnames :
-sig
-
- open Util
- open Names
-
- type full_path
- val pr_path : full_path -> Pp.t
- val make_path : Names.DirPath.t -> Names.Id.t -> full_path
- val eq_full_path : full_path -> full_path -> bool
- val repr_path : full_path -> Names.DirPath.t * Names.Id.t
- val dirpath : full_path -> Names.DirPath.t
- val path_of_string : string -> full_path
-
- type qualid
- val make_qualid : Names.DirPath.t -> Names.Id.t -> qualid
- val qualid_eq : qualid -> qualid -> bool
- val repr_qualid : qualid -> Names.DirPath.t * Names.Id.t
- val pr_qualid : qualid -> Pp.t
- val string_of_qualid : qualid -> string
- val qualid_of_string : string -> qualid
- val qualid_of_path : full_path -> qualid
- val qualid_of_dirpath : Names.DirPath.t -> qualid
- val qualid_of_ident : Names.Id.t -> qualid
-
- type reference =
- | Qualid of qualid Loc.located
- | Ident of Names.Id.t Loc.located
- val loc_of_reference : reference -> Loc.t option
- val qualid_of_reference : reference -> qualid Loc.located
- val pr_reference : reference -> Pp.t
-
- val is_dirpath_prefix_of : Names.DirPath.t -> Names.DirPath.t -> bool
- val split_dirpath : Names.DirPath.t -> Names.DirPath.t * Names.Id.t
- val dirpath_of_string : string -> Names.DirPath.t
- val pr_dirpath : Names.DirPath.t -> Pp.t
- [@@ocaml.deprecated "Alias for DirPath.print"]
-
- val string_of_path : full_path -> string
-
- val basename : full_path -> Names.Id.t
-
- type object_name = full_path * Names.KerName.t
- type object_prefix = {
- obj_dir : DirPath.t;
- obj_mp : ModPath.t;
- obj_sec : DirPath.t;
- }
-
- module Dirset : Set.S with type elt = DirPath.t
- module Dirmap : Map.ExtS with type key = DirPath.t and module Set := Dirset
- module Spmap : CSig.MapS with type key = full_path
-end
-
module Globnames :
sig
@@ -2205,6 +2156,66 @@ sig
| SubEvar of Evar.t
end
+module Decl_kinds :
+sig
+ type polymorphic = bool
+ type cumulative_inductive_flag = bool
+ type recursivity_kind = Declarations.recursivity_kind =
+ | Finite
+ | CoFinite
+ | BiFinite
+ [@@ocaml.deprecated "Please use [Declarations.recursivity_kind"]
+
+ type discharge =
+ | DoDischarge
+ | NoDischarge
+
+ type locality =
+ | Discharge
+ | Local
+ | Global
+
+ type definition_object_kind =
+ | Definition
+ | Coercion
+ | SubClass
+ | CanonicalStructure
+ | Example
+ | Fixpoint
+ | CoFixpoint
+ | Scheme
+ | StructureComponent
+ | IdentityCoercion
+ | Instance
+ | Method
+ | Let
+ type theorem_kind =
+ | Theorem
+ | Lemma
+ | Fact
+ | Remark
+ | Property
+ | Proposition
+ | Corollary
+ type goal_object_kind =
+ | DefinitionBody of definition_object_kind
+ | Proof of theorem_kind
+ type goal_kind = locality * polymorphic * goal_object_kind
+ type assumption_object_kind =
+ | Definitional
+ | Logical
+ | Conjectural
+ type logical_kind =
+ | IsAssumption of assumption_object_kind
+ | IsDefinition of definition_object_kind
+ | IsProof of theorem_kind
+ type binding_kind =
+ | Explicit
+ | Implicit
+ type private_flag = bool
+ type definition_kind = locality * polymorphic * definition_object_kind
+end
+
module Glob_term :
sig
type 'a cases_pattern_r =
@@ -2383,7 +2394,7 @@ sig
and case_expr = constr_expr * Names.Name.t Loc.located option * cases_pattern_expr option
and branch_expr =
- (cases_pattern_expr list Loc.located list * constr_expr) Loc.located
+ (cases_pattern_expr list list * constr_expr) Loc.located
and binder_expr =
Names.Name.t Loc.located list * binder_kind * constr_expr
@@ -2754,10 +2765,10 @@ sig
type universe_binders
type universe_opt_subst
val fresh_inductive_instance : Environ.env -> Names.inductive -> Constr.pinductive Univ.in_universe_context_set
- val new_Type : Names.DirPath.t -> Constr.types
+ val new_Type : unit -> Constr.types
val type_of_global : Globnames.global_reference -> Constr.types Univ.in_universe_context_set
val constr_of_global : Globnames.global_reference -> Constr.t
- val new_univ_level : Names.DirPath.t -> Univ.Level.t
+ val new_univ_level : unit -> Univ.Level.t
val new_sort_in_family : Sorts.family -> Sorts.t
val pr_with_global_universes : Univ.Level.t -> Pp.t
val pr_universe_opt_subst : universe_opt_subst -> Pp.t
@@ -2770,6 +2781,8 @@ sig
end
type universe_constraints = Constraints.t
+ [@@ocaml.deprecated "Use Constraints.t"]
+
end
module UState :
@@ -3106,7 +3119,7 @@ sig
val fold : Evd.evar_map -> ('a -> constr -> 'a) -> 'a -> constr -> 'a
val existential_type : Evd.evar_map -> existential -> types
val iter : Evd.evar_map -> (constr -> unit) -> constr -> unit
- val eq_constr_universes : Evd.evar_map -> constr -> constr -> Universes.universe_constraints option
+ val eq_constr_universes : Evd.evar_map -> constr -> constr -> Universes.Constraints.t option
val eq_constr_nounivs : Evd.evar_map -> constr -> constr -> bool
val compare_constr : Evd.evar_map -> (constr -> constr -> bool) -> constr -> constr -> bool
val isApp : Evd.evar_map -> constr -> bool
@@ -3840,10 +3853,9 @@ sig
type matching_result =
{ m_sub : bound_ident_map * Ltac_pretype.patvar_map;
m_ctx : EConstr.constr }
- val match_subterm_gen : Environ.env -> Evd.evar_map ->
- bool ->
- binding_bound_vars * Pattern.constr_pattern -> EConstr.constr ->
- matching_result IStream.t
+ val match_subterm : Environ.env -> Evd.evar_map ->
+ binding_bound_vars * Pattern.constr_pattern -> EConstr.constr ->
+ matching_result IStream.t
val matches : Environ.env -> Evd.evar_map -> Pattern.constr_pattern -> EConstr.constr -> Ltac_pretype.patvar_map
end
@@ -3994,7 +4006,7 @@ sig
type instance_flag = bool option
type coercion_flag = bool
- type inductive_flag = Decl_kinds.recursivity_kind
+ type inductive_flag = Declarations.recursivity_kind
type lname = Names.Name.t Loc.located
type lident = Names.Id.t Loc.located
type opacity_flag =
@@ -4334,6 +4346,7 @@ sig
| Later : [ `thunk ] delay
val print_universes : bool ref
val print_evar_arguments : bool ref
+ val print_allow_match_default_clause : bool ref
val detype : 'a delay -> ?lax:bool -> bool -> Names.Id.Set.t -> Environ.env -> Evd.evar_map -> EConstr.constr -> 'a Glob_term.glob_constr_g
val subst_glob_constr : Mod_subst.substitution -> Glob_term.glob_constr -> Glob_term.glob_constr
val set_detype_anonymous : (?loc:Loc.t -> int -> Names.Id.t) -> unit
@@ -4788,7 +4801,7 @@ sig
| IntroNeedsProduct
| DoesNotOccurIn of Constr.t * Names.Id.t
| NoSuchHyp of Names.Id.t
- exception RefinerError of refiner_error
+ exception RefinerError of Environ.env * Evd.evar_map * refiner_error
val catchable_exception : exn -> bool
end
@@ -4842,17 +4855,16 @@ sig
set : unit -> unit ;
reset : unit -> unit
}
- type proof_universes = UState.t * Universes.universe_binders option
type proof_object = {
id : Names.Id.t;
entries : Safe_typing.private_constants Entries.definition_entry list;
persistence : Decl_kinds.goal_kind;
- universes: proof_universes;
+ universes: UState.t;
}
type proof_ending =
| Admitted of Names.Id.t * Decl_kinds.goal_kind * Entries.parameter_entry *
- proof_universes
+ UState.t
| Proved of Vernacexpr.opacity_flag *
Vernacexpr.lident option *
proof_object
@@ -4960,8 +4972,6 @@ sig
val pf_conv_x : Goal.goal Evd.sigma -> EConstr.constr -> EConstr.constr -> bool
- val pf_is_matching : Goal.goal Evd.sigma -> Pattern.constr_pattern -> EConstr.constr -> bool
-
val pf_hyps_types : Goal.goal Evd.sigma -> (Names.Id.t * EConstr.types) list
val pr_gls : Goal.goal Evd.sigma -> Pp.t
@@ -5006,7 +5016,7 @@ sig
Vernacexpr.goal_selector -> int option -> unit Proofview.tactic ->
Proof.t -> Proof.t * bool
val cook_proof :
- unit -> (Names.Id.t * (Safe_typing.private_constants Entries.definition_entry * Proof_global.proof_universes * Decl_kinds.goal_kind))
+ unit -> (Names.Id.t * (Safe_typing.private_constants Entries.definition_entry * UState.t * Decl_kinds.goal_kind))
val get_current_context : unit -> Evd.evar_map * Environ.env
val current_proof_statement : unit -> Names.Id.t * Decl_kinds.goal_kind * EConstr.types
@@ -5243,16 +5253,20 @@ end
module Genprint :
sig
- type printer_with_level =
+ type 'a with_level =
{ default_already_surrounded : Notation_term.tolerability;
default_ensure_surrounded : Notation_term.tolerability;
- printer : Environ.env -> Evd.evar_map -> Notation_term.tolerability -> Pp.t }
+ printer : 'a }
type printer_result =
- | PrinterBasic of (unit -> Pp.t)
- | PrinterNeedsContext of (Environ.env -> Evd.evar_map -> Pp.t)
- | PrinterNeedsContextAndLevel of printer_with_level
- type 'a printer = 'a -> Pp.t
- type 'a top_printer = 'a -> printer_result
+| PrinterBasic of (unit -> Pp.t)
+| PrinterNeedsLevel of (Notation_term.tolerability -> Pp.t) with_level
+ type printer_fun_with_level = Environ.env -> Evd.evar_map -> Notation_term.tolerability -> Pp.t
+ type top_printer_result =
+ | TopPrinterBasic of (unit -> Pp.t)
+ | TopPrinterNeedsContext of (Environ.env -> Evd.evar_map -> Pp.t)
+ | TopPrinterNeedsContextAndLevel of printer_fun_with_level with_level
+ type 'a printer = 'a -> printer_result
+ type 'a top_printer = 'a -> top_printer_result
val register_print0 : ('raw, 'glb, 'top) Genarg.genarg_type ->
'raw printer -> 'glb printer -> 'top top_printer -> unit
val register_vernac_print0 : ('raw, 'glb, 'top) Genarg.genarg_type ->
@@ -6062,7 +6076,7 @@ sig
val do_mutual_inductive :
(Vernacexpr.one_inductive_expr * Vernacexpr.decl_notation list) list -> Decl_kinds.cumulative_inductive_flag -> Decl_kinds.polymorphic ->
- Decl_kinds.private_flag -> Decl_kinds.recursivity_kind -> unit
+ Decl_kinds.private_flag -> Declarations.recursivity_kind -> unit
val do_definition : Names.Id.t -> Decl_kinds.definition_kind -> Vernacexpr.universe_decl_expr option ->
Constrexpr.local_binder_expr list -> Redexpr.red_expr option -> Constrexpr.constr_expr ->
@@ -6088,7 +6102,7 @@ sig
structured_inductive_expr -> Vernacexpr.decl_notation list ->
Decl_kinds.cumulative_inductive_flag ->
Decl_kinds.polymorphic ->
- Decl_kinds.private_flag -> Decl_kinds.recursivity_kind ->
+ Decl_kinds.private_flag -> Declarations.recursivity_kind ->
Entries.mutual_inductive_entry * Universes.universe_binders * one_inductive_impls list
val declare_mutual_inductive_with_eliminations :