aboutsummaryrefslogtreecommitdiffhomepage
path: root/kernel
diff options
context:
space:
mode:
Diffstat (limited to 'kernel')
-rw-r--r--kernel/cooking.ml2
-rw-r--r--kernel/cooking.mli3
-rw-r--r--kernel/indtypes.ml15
-rw-r--r--kernel/inductive.ml10
-rw-r--r--kernel/modops.ml4
-rw-r--r--kernel/nativecode.ml14
-rw-r--r--kernel/nativecode.mli2
-rw-r--r--kernel/nativelibrary.ml2
-rw-r--r--kernel/subtyping.ml4
-rw-r--r--kernel/term_typing.ml6
10 files changed, 25 insertions, 37 deletions
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index 521a7d890..094609b96 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -204,7 +204,7 @@ let lift_univs cb subst auctx0 =
let auctx' = Univ.subst_univs_level_abstract_universe_context (Univ.make_instance_subst subst) auctx in
subst, (Polymorphic_const (AUContext.union auctx0 auctx'))
-let cook_constant ~hcons env { from = cb; info } =
+let cook_constant ~hcons { from = cb; info } =
let { Opaqueproof.modlist; abstract } = info in
let cache = RefTable.create 13 in
let abstract, usubst, abs_ctx = abstract in
diff --git a/kernel/cooking.mli b/kernel/cooking.mli
index 4c254d2ea..6ebe691b8 100644
--- a/kernel/cooking.mli
+++ b/kernel/cooking.mli
@@ -10,7 +10,6 @@
open Constr
open Declarations
-open Environ
(** {6 Cooking the constants. } *)
@@ -26,7 +25,7 @@ type result = {
cook_context : Constr.named_context option;
}
-val cook_constant : hcons:bool -> env -> recipe -> result
+val cook_constant : hcons:bool -> recipe -> result
val cook_constr : Opaqueproof.cooking_info -> constr -> constr
(** {6 Utility functions used in module [Discharge]. } *)
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index 61b71df31..5d45c2c1a 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -120,16 +120,6 @@ let mind_check_names mie =
(* Typing the arities and constructor types *)
-(* An inductive definition is a "unit" if it has only one constructor
- and that all arguments expected by this constructor are
- logical, this is the case for equality, conjunction of logical properties
-*)
-let is_unit constrsinfos =
- match constrsinfos with (* One info = One constructor *)
- | [level] -> is_type0m_univ level
- | [] -> (* type without constructors *) true
- | _ -> false
-
let infos_and_sort env t =
let rec aux env t max =
let t = whd_all env t in
@@ -174,10 +164,9 @@ let infer_constructor_packet env_ar_par params lc =
let lc'' = Array.map (fun j -> Term.it_mkProd_or_LetIn j.utj_val params) jlc in
(* compute the max of the sorts of the products of the constructors types *)
let levels = List.map (infos_and_sort env_ar_par) lc in
- let isunit = is_unit levels in
let min = if Array.length jlc > 1 then Universe.type0 else Universe.type0m in
let level = List.fold_left (fun max l -> Universe.sup max l) min levels in
- (lc'', (isunit, level))
+ (lc'', level)
(* If indices matter *)
let cumulate_arity_large_levels env sign =
@@ -354,7 +343,7 @@ let typecheck_inductive env mie =
(* Compute/check the sorts of the inductive types *)
let inds =
- Array.map (fun ((id,full_arity,sign,expltype,def_level,inf_level),cn,lc,(is_unit,clev)) ->
+ Array.map (fun ((id,full_arity,sign,expltype,def_level,inf_level),cn,lc,clev) ->
let infu =
(** Inferred level, with parameters and constructors. *)
match inf_level with
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 584c1af03..88b00600e 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -785,7 +785,7 @@ let rec subterm_specif renv stack t =
| Lambda (x,a,b) ->
let () = assert (List.is_empty l) in
- let spec,stack' = extract_stack renv a stack in
+ let spec,stack' = extract_stack stack in
subterm_specif (push_var renv (x,a,spec)) stack' b
(* Metas and evars are considered OK *)
@@ -817,7 +817,7 @@ and stack_element_specif = function
|SClosure (h_renv,h) -> lazy_subterm_specif h_renv [] h
|SArg x -> x
-and extract_stack renv a = function
+and extract_stack = function
| [] -> Lazy.from_val Not_subterm , []
| h::t -> stack_element_specif h, t
@@ -848,7 +848,7 @@ let error_illegal_rec_call renv fx (arg_renv,arg) =
let error_partial_apply renv fx =
raise (FixGuardError (renv.env,NotEnoughArgumentsForFixCall fx))
-let filter_stack_domain env ci p stack =
+let filter_stack_domain env p stack =
let absctx, ar = dest_lam_assum env p in
(* Optimization: if the predicate is not dependent, no restriction is needed
and we avoid building the recargs tree. *)
@@ -933,7 +933,7 @@ let check_one_fix renv recpos trees def =
let case_spec = branches_specif renv
(lazy_subterm_specif renv [] c_0) ci in
let stack' = push_stack_closures renv l stack in
- let stack' = filter_stack_domain renv.env ci p stack' in
+ let stack' = filter_stack_domain renv.env p stack' in
Array.iteri (fun k br' ->
let stack_br = push_stack_args case_spec.(k) stack' in
check_rec_call renv stack_br br') lrest
@@ -976,7 +976,7 @@ let check_one_fix renv recpos trees def =
| Lambda (x,a,b) ->
let () = assert (List.is_empty l) in
check_rec_call renv [] a ;
- let spec, stack' = extract_stack renv a stack in
+ let spec, stack' = extract_stack stack in
check_rec_call (push_var renv (x,a,spec)) stack' b
| Prod (x,a,b) ->
diff --git a/kernel/modops.ml b/kernel/modops.ml
index 02bab581a..98a997311 100644
--- a/kernel/modops.ml
+++ b/kernel/modops.ml
@@ -265,7 +265,7 @@ let subst_structure subst = subst_structure subst do_delta_codom
(* spiwack: here comes the function which takes care of importing
the retroknowledge declared in the library *)
(* lclrk : retroknowledge_action list, rkaction : retroknowledge action *)
-let add_retroknowledge mp =
+let add_retroknowledge =
let perform rkaction env = match rkaction with
| Retroknowledge.RKRegister (f, e) when (isConst e || isInd e) ->
Environ.register env f e
@@ -309,7 +309,7 @@ and add_module mb linkinfo env =
let env = Environ.shallow_add_module mb env in
match mb.mod_type with
|NoFunctor struc ->
- add_retroknowledge mp mb.mod_retroknowledge
+ add_retroknowledge mb.mod_retroknowledge
(add_structure mp struc mb.mod_delta linkinfo env)
|MoreFunctor _ -> env
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index 1748e98a4..39f7de942 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -1649,15 +1649,15 @@ let pp_mllam fmt l =
and pp_letrec fmt defs =
let len = Array.length defs in
- let pp_one_rec i (fn, argsn, body) =
+ let pp_one_rec (fn, argsn, body) =
Format.fprintf fmt "%a%a =@\n %a"
pp_lname fn
pp_ldecls argsn pp_mllam body in
Format.fprintf fmt "@[let rec ";
- pp_one_rec 0 defs.(0);
+ pp_one_rec defs.(0);
for i = 1 to len - 1 do
Format.fprintf fmt "@\nand ";
- pp_one_rec i defs.(i)
+ pp_one_rec defs.(i)
done;
and pp_blam fmt l =
@@ -1941,7 +1941,7 @@ let is_code_loaded ~interactive name =
let param_name = Name (Id.of_string "params")
let arg_name = Name (Id.of_string "arg")
-let compile_mind prefix ~interactive mb mind stack =
+let compile_mind mb mind stack =
let u = Declareops.inductive_polymorphic_context mb in
(** Generate data for every block *)
let f i stack ob =
@@ -2020,7 +2020,7 @@ let compile_mind_deps env prefix ~interactive
then init
else
let comp_stack =
- compile_mind prefix ~interactive mib mind comp_stack
+ compile_mind mib mind comp_stack
in
let name =
if interactive then LinkedInteractive prefix
@@ -2092,9 +2092,9 @@ let compile_constant_field env prefix con acc cb =
in
gl@acc
-let compile_mind_field prefix mp l acc mb =
+let compile_mind_field mp l acc mb =
let mind = MutInd.make2 mp l in
- compile_mind prefix ~interactive:false mb mind acc
+ compile_mind mb mind acc
let mk_open s = Gopen s
diff --git a/kernel/nativecode.mli b/kernel/nativecode.mli
index 684983a87..96efa7faa 100644
--- a/kernel/nativecode.mli
+++ b/kernel/nativecode.mli
@@ -67,7 +67,7 @@ val register_native_file : string -> unit
val compile_constant_field : env -> string -> Constant.t ->
global list -> constant_body -> global list
-val compile_mind_field : string -> ModPath.t -> Label.t ->
+val compile_mind_field : ModPath.t -> Label.t ->
global list -> mutual_inductive_body -> global list
val mk_conv_code : env -> evars -> string -> constr -> constr -> linkable_code
diff --git a/kernel/nativelibrary.ml b/kernel/nativelibrary.ml
index 8bff43632..edce9367f 100644
--- a/kernel/nativelibrary.ml
+++ b/kernel/nativelibrary.ml
@@ -37,7 +37,7 @@ and translate_field prefix mp env acc (l,x) =
let id = mb.mind_packets.(0).mind_typename in
let msg = Printf.sprintf "Compiling inductive %s..." (Id.to_string id) in
Feedback.msg_debug (Pp.str msg));
- compile_mind_field prefix mp l acc mb
+ compile_mind_field mp l acc mb
| SFBmodule md ->
let mp = md.mod_mp in
(if !Flags.debug then
diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml
index 1e58f5c24..74042f9e0 100644
--- a/kernel/subtyping.ml
+++ b/kernel/subtyping.ml
@@ -224,7 +224,7 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
cst
-let check_constant cst env mp1 l info1 cb2 spec2 subst1 subst2 =
+let check_constant cst env l info1 cb2 spec2 subst1 subst2 =
let error why = error_signature_mismatch l spec2 why in
let check_conv cst poly f = check_conv_error error cst poly f in
let check_type poly cst env t1 t2 =
@@ -292,7 +292,7 @@ and check_signatures cst env mp1 sig1 mp2 sig2 subst1 subst2 reso1 reso2=
let check_one_body cst (l,spec2) =
match spec2 with
| SFBconst cb2 ->
- check_constant cst env mp1 l (get_obj mp1 map1 l)
+ check_constant cst env l (get_obj mp1 map1 l)
cb2 spec2 subst1 subst2
| SFBmind mib2 ->
check_inductive cst env mp1 l (get_obj mp1 map1 l)
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index bad449731..1f7ee145a 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -378,7 +378,7 @@ let build_constant_declaration kn env result =
str "Proof using " ++ declared_vars ++ fnl () ++
str "to" ++ fnl () ++
str "Proof using " ++ inferred_vars) in
- let sort evn l =
+ let sort l =
List.filter (fun decl ->
let id = NamedDecl.get_id decl in
List.exists (NamedDecl.get_id %> Names.Id.equal id) l)
@@ -411,7 +411,7 @@ let build_constant_declaration kn env result =
[], def (* Empty section context: no need to check *)
| Some declared ->
(* We use the declared set and chain a check of correctness *)
- sort env declared,
+ sort declared,
match def with
| Undef _ as x -> x (* nothing to check *)
| Def cs as x ->
@@ -554,7 +554,7 @@ let translate_recipe env kn r =
be useless. It is detected by the dirpath of the constant being empty. *)
let (_, dir, _) = Constant.repr3 kn in
let hcons = DirPath.is_empty dir in
- build_constant_declaration kn env (Cooking.cook_constant ~hcons env r)
+ build_constant_declaration kn env (Cooking.cook_constant ~hcons r)
let translate_local_def env id centry =
let open Cooking in