aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--.gitignore4
-rw-r--r--Makefile.build2
-rw-r--r--Makefile.checker45
-rw-r--r--checker/check.ml84
-rw-r--r--checker/check.mli5
-rw-r--r--checker/check.mllib64
-rw-r--r--checker/checkInductive.ml271
-rw-r--r--checker/checkInductive.mli (renamed from checker/subtyping.mli)9
-rw-r--r--checker/checkTypes.ml36
-rw-r--r--checker/checkTypes.mli (renamed from checker/typeops.mli)6
-rw-r--r--checker/check_stat.ml32
-rw-r--r--checker/check_stat.mli3
-rw-r--r--checker/checker.ml113
-rw-r--r--checker/cic.mli471
-rw-r--r--checker/closure.ml791
-rw-r--r--checker/closure.mli174
-rw-r--r--checker/declarations.ml606
-rw-r--r--checker/declarations.mli52
-rw-r--r--checker/dune29
-rw-r--r--checker/environ.ml250
-rw-r--r--checker/environ.mli79
-rw-r--r--checker/indtypes.ml652
-rw-r--r--checker/indtypes.mli39
-rw-r--r--checker/inductive.ml1205
-rw-r--r--checker/inductive.mli85
-rw-r--r--checker/mod_checking.ml56
-rw-r--r--checker/mod_checking.mli2
-rw-r--r--checker/modops.ml151
-rw-r--r--checker/modops.mli49
-rw-r--r--checker/print.ml96
-rw-r--r--checker/reduction.ml627
-rw-r--r--checker/reduction.mli55
-rw-r--r--checker/safe_checking.ml (renamed from checker/safe_typing.mli)19
-rw-r--r--checker/safe_checking.mli (renamed from checker/print.mli)7
-rw-r--r--checker/safe_typing.ml96
-rw-r--r--checker/subtyping.ml344
-rw-r--r--checker/term.ml447
-rw-r--r--checker/term.mli59
-rw-r--r--checker/type_errors.ml115
-rw-r--r--checker/type_errors.mli106
-rw-r--r--checker/typeops.ml382
-rw-r--r--checker/univ.ml1127
-rw-r--r--checker/univ.mli301
-rw-r--r--dev/checker_printers.ml4
-rw-r--r--dev/checker_printers.mli4
-rw-r--r--dev/ocamldebug-coq.run58
-rw-r--r--engine/termops.ml85
-rw-r--r--engine/termops.mli3
-rw-r--r--kernel/constr.ml67
-rw-r--r--kernel/constr.mli3
-rw-r--r--kernel/indtypes.mli13
-rw-r--r--kernel/safe_typing.ml2
-rw-r--r--kernel/safe_typing.mli4
-rw-r--r--kernel/sorts.ml10
-rw-r--r--kernel/sorts.mli4
-rw-r--r--plugins/funind/g_indfun.mlg2
-rw-r--r--pretyping/indrec.ml2
-rw-r--r--pretyping/recordops.ml2
-rw-r--r--pretyping/reductionops.ml2
-rw-r--r--vernac/himsg.ml6
-rw-r--r--vernac/ppvernac.ml4
61 files changed, 661 insertions, 8760 deletions
diff --git a/.gitignore b/.gitignore
index f9e43a0eb7..e513837445 100644
--- a/.gitignore
+++ b/.gitignore
@@ -150,10 +150,6 @@ kernel/byterun/coq_jumptbl.h
kernel/copcodes.ml
ide/index_urls.txt
.lia.cache
-checker/names.ml
-checker/names.mli
-checker/esubst.ml
-checker/esubst.mli
# emacs save files
*~
diff --git a/Makefile.build b/Makefile.build
index fb84a131c7..ee856aae8e 100644
--- a/Makefile.build
+++ b/Makefile.build
@@ -612,7 +612,7 @@ VALIDOPTS=$(if $(VERBOSE),,-silent) -o -m
validate: $(CHICKEN) | $(ALLVO)
$(SHOW)'COQCHK <theories & plugins>'
- $(HIDE)$(CHICKEN) -boot $(VALIDOPTS) $(ALLMODS)
+ $(HIDE)$(CHICKEN) -boot -debug $(VALIDOPTS) $(ALLMODS)
$(ALLSTDLIB).v:
$(SHOW)'MAKE $(notdir $@)'
diff --git a/Makefile.checker b/Makefile.checker
index 3d1d251701..a0f0778d49 100644
--- a/Makefile.checker
+++ b/Makefile.checker
@@ -22,7 +22,7 @@ CHICKEN:=bin/coqchk$(EXE)
# The sources
-CHKLIBS:= -I config -I clib -I lib -I checker
+CHKLIBS:= -I config -I clib -I lib -I checker -I kernel -I kernel/byterun
## NB: currently, both $(OPTFLAGS) and $(BYTEFLAGS) contain -thread
@@ -34,22 +34,9 @@ CHECKMLLIBFILE := checker/.mllibfiles
CHECKERDEPS := $(addsuffix .d, $(CHECKMLDFILE) $(CHECKMLLIBFILE))
-include $(CHECKERDEPS)
-# Copied files
-checker/esubst.mli: kernel/esubst.mli
- cp -a $< $@
- sed -i.bak '1i(* AUTOGENERATED FILE: DO NOT EDIT *)\n\n\n\n\n\n\n\n' $@ && rm $@.bak
-checker/esubst.ml: kernel/esubst.ml
- cp -a $< $@
- sed -i.bak '1i(* AUTOGENERATED FILE: DO NOT EDIT *)\n\n\n\n\n\n\n\n' $@ && rm $@.bak
-checker/names.mli: kernel/names.mli
- cp -a $< $@
- sed -i.bak '1i(* AUTOGENERATED FILE: DO NOT EDIT *)\n\n\n\n\n\n\n\n' $@ && rm $@.bak
-checker/names.ml: kernel/names.ml
- cp -a $< $@
- sed -i.bak '1i(* AUTOGENERATED FILE: DO NOT EDIT *)\n\n\n\n\n\n\n\n' $@ && rm $@.bak
-
ifeq ($(BEST),opt)
-$(CHICKEN): config/config.cmxa checker/check.cmxa checker/coqchk.mli checker/coqchk.ml
+$(CHICKEN): config/config.cmxa clib/clib.cmxa lib/lib.cmxa kernel/kernel.cmxa \
+checker/check.cmxa $(LIBCOQRUN) checker/coqchk.mli checker/coqchk.ml
$(SHOW)'OCAMLOPT -o $@'
$(HIDE)$(OCAMLOPT) -linkpkg $(SYSMOD) $(CHKLIBS) $(OPTFLAGS) $(LINKMETADATA) -o $@ $^
$(STRIP_HIDE) $@
@@ -59,30 +46,29 @@ $(CHICKEN): $(CHICKENBYTE)
cp $< $@
endif
-$(CHICKENBYTE): config/config.cma checker/check.cma checker/coqchk.mli checker/coqchk.ml
+$(CHICKENBYTE): config/config.cma clib/clib.cma lib/lib.cma kernel/kernel.cma \
+checker/check.cma $(LIBCOQRUN) checker/coqchk.mli checker/coqchk.ml
$(SHOW)'OCAMLC -o $@'
- $(HIDE)$(OCAMLC) -linkpkg $(SYSMOD) $(CHKLIBS) $(BYTEFLAGS) $(CUSTOM) -o $@ $^
+ $(HIDE)$(OCAMLC) -linkpkg $(SYSMOD) $(CHKLIBS) $(BYTEFLAGS) -cclib -lcoqrun $(VMBYTEFLAGS) -o $@ $^
-checker/check.cma: checker/check.mllib | md5chk
+checker/check.cma: checker/check.mllib
$(SHOW)'OCAMLC -a -o $@'
$(HIDE)$(OCAMLC) $(CHKLIBS) $(BYTEFLAGS) -a -o $@ $(filter-out %.mllib, $^)
-checker/check.cmxa: checker/check.mllib | md5chk
+checker/check.cmxa: checker/check.mllib
$(SHOW)'OCAMLOPT -a -o $@'
$(HIDE)$(OCAMLOPT) $(CHKLIBS) $(OPTFLAGS) -a -o $@ $(filter-out %.mllib, $^)
-CHECKGENFILES:=$(addprefix checker/, names.mli names.ml esubst.mli esubst.ml)
-
-CHECKMLFILES:=$(filter checker/%, $(MLFILES) $(MLIFILES)) $(CHECKGENFILES) \
+CHECKMLFILES:=$(filter checker/%, $(MLFILES) $(MLIFILES)) \
$(filter dev/checker_%, $(MLFILES) $(MLIFILES))
-$(CHECKMLDFILE).d: $(filter checker/%, $(MLFILES) $(MLIFILES) $(CHECKGENFILES))
+$(CHECKMLDFILE).d: $(filter checker/%, $(MLFILES) $(MLIFILES))
$(SHOW)'OCAMLDEP checker/MLFILES checker/MLIFILES'
$(HIDE)$(OCAMLFIND) ocamldep -slash $(CHKLIBS) $(CHECKMLFILES) $(TOTARGET)
-$(CHECKMLLIBFILE).d: $(filter checker/%, $(MLLIBFILES) $(MLPACKFILES) $(CHECKGENFILES)) | $(OCAMLLIBDEP)
+$(CHECKMLLIBFILE).d: $(filter checker/%, $(MLLIBFILES) $(MLPACKFILES)) | $(OCAMLLIBDEP)
$(SHOW)'OCAMLLIBDEP checker/MLLIBFILES checker/MLPACKFILES'
- $(HIDE)$(OCAMLLIBDEP) $(CHKLIBS) $(filter checker/%, $(MLLIBFILES) $(MLPACKFILES) $(CHECKGENFILES)) $(TOTARGET)
+ $(HIDE)$(OCAMLLIBDEP) $(CHKLIBS) $(filter checker/%, $(MLLIBFILES) $(MLPACKFILES)) $(TOTARGET)
checker/%.cmi: checker/%.mli
$(SHOW)'OCAMLC $<'
@@ -104,13 +90,6 @@ dev/checker_%.cmi: dev/checker_%.mli
$(SHOW)'OCAMLC $<'
$(HIDE)$(OCAMLC) $(CHKLIBS) $(BYTEFLAGS) -I dev/ -c $<
-md5chk:
- $(SHOW)'MD5SUM cic.mli'
- $(HIDE)if grep -q "^MD5 $$($(OCAML) tools/md5sum.ml checker/cic.mli)$$" checker/values.ml; \
- then true; else echo "Error: outdated checker/values.ml" >&2; false; fi
-
-.PHONY: md5chk
-
# For emacs:
# Local Variables:
# mode: makefile
diff --git a/checker/check.ml b/checker/check.ml
index 4bb485d29e..e3a4bda8ec 100644
--- a/checker/check.ml
+++ b/checker/check.ml
@@ -13,7 +13,7 @@ open CErrors
open Util
open Names
-let chk_pp = Pp.pp_with Format.std_formatter
+let chk_pp = Feedback.msg_notice
let pr_dirpath dp = str (DirPath.to_string dp)
let default_root_prefix = DirPath.empty
@@ -48,13 +48,17 @@ let pr_path sp =
(*s Modules loaded in memory contain the following informations. They are
kept in the global table [libraries_table]. *)
+type compilation_unit_name = DirPath.t
+
+type seg_proofs = Constr.constr Future.computation array
+
type library_t = {
- library_name : Cic.compilation_unit_name;
+ library_name : compilation_unit_name;
library_filename : CUnix.physical_path;
- library_compiled : Cic.compiled_library;
- library_opaques : Cic.opaque_table;
- library_deps : Cic.library_deps;
- library_digest : Cic.vodigest;
+ library_compiled : Safe_typing.compiled_library;
+ library_opaques : seg_proofs;
+ library_deps : (compilation_unit_name * Safe_typing.vodigest) array;
+ library_digest : Safe_typing.vodigest;
library_extra_univs : Univ.ContextSet.t }
module LibraryOrdered =
@@ -94,35 +98,36 @@ let access_opaque_table dp i =
with Not_found -> assert false
in
assert (i < Array.length t);
- Future.force t.(i)
+ t.(i)
let access_opaque_univ_table dp i =
try
let t = LibraryMap.find dp !opaque_univ_tables in
assert (i < Array.length t);
- Future.force t.(i)
- with Not_found -> Univ.ContextSet.empty
+ Some t.(i)
+ with Not_found -> None
+let () =
+ Opaqueproof.set_indirect_opaque_accessor access_opaque_table;
+ Opaqueproof.set_indirect_univ_accessor access_opaque_univ_table
-let _ = Declarations.indirect_opaque_access := access_opaque_table
-let _ = Declarations.indirect_opaque_univ_access := access_opaque_univ_table
-
-let check_one_lib admit (dir,m) =
- let file = m.library_filename in
+let check_one_lib admit senv (dir,m) =
let md = m.library_compiled in
let dig = m.library_digest in
(* Look up if the library is to be admitted correct. We could
also check if it carries a validation certificate (yet to
be implemented). *)
- if LibrarySet.mem dir admit then
- (Flags.if_verbose Feedback.msg_notice
- (str "Admitting library: " ++ pr_dirpath dir);
- Safe_typing.unsafe_import file md m.library_extra_univs dig)
- else
- (Flags.if_verbose Feedback.msg_notice
- (str "Checking library: " ++ pr_dirpath dir);
- Safe_typing.import file md m.library_extra_univs dig);
- register_loaded_library m
+ let senv =
+ if LibrarySet.mem dir admit then
+ (Flags.if_verbose Feedback.msg_notice
+ (str "Admitting library: " ++ pr_dirpath dir);
+ Safe_checking.unsafe_import senv md m.library_extra_univs dig)
+ else
+ (Flags.if_verbose Feedback.msg_notice
+ (str "Checking library: " ++ pr_dirpath dir);
+ Safe_checking.import senv md m.library_extra_univs dig)
+ in
+ register_loaded_library m; senv
(*************************************************************************)
(*s Load path. Mapping from physical to logical paths etc.*)
@@ -284,7 +289,21 @@ let raw_intern_library f =
(************************************************************************)
(* Internalise libraries *)
-open Cic
+type summary_disk = {
+ md_name : compilation_unit_name;
+ md_imports : compilation_unit_name array;
+ md_deps : (compilation_unit_name * Safe_typing.vodigest) array;
+}
+
+module Dyn = Dyn.Make ()
+type obj = Dyn.t (* persistent dynamic objects *)
+type lib_objects = (Id.t * obj) list
+type library_objects = lib_objects * lib_objects
+
+type library_disk = {
+ md_compiled : Safe_typing.compiled_library;
+ md_objects : library_objects;
+}
let mk_library sd md f table digest cst = {
library_name = sd.md_name;
@@ -317,12 +336,12 @@ let intern_from_file (dir, f) =
let (sd,md,table,opaque_csts,digest) =
try
let ch = System.with_magic_number_check raw_intern_library f in
- let (sd:Cic.summary_disk), _, digest = marshal_in_segment f ch in
- let (md:Cic.library_disk), _, digest = marshal_in_segment f ch in
+ let (sd:summary_disk), _, digest = marshal_in_segment f ch in
+ let (md:library_disk), _, digest = marshal_in_segment f ch in
let (opaque_csts:'a option), _, udg = marshal_in_segment f ch in
let (discharging:'a option), _, _ = marshal_in_segment f ch in
let (tasks:'a option), _, _ = marshal_in_segment f ch in
- let (table:Cic.opaque_table), pos, checksum =
+ let (table:seg_proofs), pos, checksum =
marshal_in_segment f ch in
(* Verification of the final checksum *)
let () = close_in ch in
@@ -350,8 +369,8 @@ let intern_from_file (dir, f) =
Validate.validate !Flags.debug Values.v_opaques table;
Flags.if_verbose chk_pp (str" done]" ++ fnl ());
let digest =
- if opaque_csts <> None then Cic.Dviovo (digest,udg)
- else (Cic.Dvo digest) in
+ if opaque_csts <> None then Safe_typing.Dvivo (digest,udg)
+ else (Safe_typing.Dvo_or_vi digest) in
sd,md,table,opaque_csts,digest
with e -> Flags.if_verbose chk_pp (str" failed!]" ++ fnl ()); raise e in
depgraph := LibraryMap.add sd.md_name sd.md_deps !depgraph;
@@ -406,7 +425,7 @@ and fold_deps_list seen ff modl needed =
let fold_deps_list ff modl acc =
snd (fold_deps_list LibrarySet.empty ff modl (LibrarySet.empty,acc))
-let recheck_library ~norec ~admit ~check =
+let recheck_library senv ~norec ~admit ~check =
let ml = List.map try_locate_qualified_library check in
let nrl = List.map try_locate_qualified_library norec in
let al = List.map try_locate_qualified_library admit in
@@ -424,5 +443,6 @@ let recheck_library ~norec ~admit ~check =
Flags.if_verbose Feedback.msg_notice (fnl()++hv 2 (str "Ordered list:" ++ fnl() ++
prlist
(fun (dir,_) -> pr_dirpath dir ++ fnl()) needed));
- List.iter (check_one_lib nochk) needed;
- Flags.if_verbose Feedback.msg_notice (str"Modules were successfully checked")
+ let senv = List.fold_left (check_one_lib nochk) senv needed in
+ Flags.if_verbose Feedback.msg_notice (str"Modules were successfully checked");
+ senv
diff --git a/checker/check.mli b/checker/check.mli
index eb6404a172..39cc93c060 100644
--- a/checker/check.mli
+++ b/checker/check.mli
@@ -10,6 +10,7 @@
open CUnix
open Names
+open Safe_typing
type section_path = {
dirpath : string list;
@@ -26,7 +27,7 @@ val default_root_prefix : DirPath.t
val add_load_path : physical_path * logical_path -> unit
-val recheck_library :
+val recheck_library : safe_environment ->
norec:object_file list ->
admit:object_file list ->
- check:object_file list -> unit
+ check:object_file list -> safe_environment
diff --git a/checker/check.mllib b/checker/check.mllib
index 173ad1e325..d47a93c70d 100644
--- a/checker/check.mllib
+++ b/checker/check.mllib
@@ -1,67 +1,9 @@
Analyze
-Hook
-Terminal
-Hashset
-Hashcons
-CSet
-CMap
-Int
-Dyn
-HMap
-Option
-Store
-Exninfo
-Backtrace
-Flags
-Control
-Pp_control
-Loc
-CList
-CString
-Serialize
-Stateid
-CObj
-CArray
-CStack
-Util
-Pp
-Ppstyle
-Xml_datatype
-Richpp
-Feedback
-Segmenttree
-Unicodetable
-Unicode
-CErrors
-CWarnings
-CEphemeron
-Future
-CUnix
-Minisys
-System
-Profile
-RemoteCounter
-Envars
-Predicate
-Rtree
-Names
-Univ
-Esubst
-Term
-Print
-Declarations
-Environ
-Closure
-Reduction
-Type_errors
-Modops
-Inductive
-Typeops
-Indtypes
-Subtyping
+CheckInductive
Mod_checking
-Safe_typing
+CheckTypes
+Safe_checking
Values
Validate
Check
diff --git a/checker/checkInductive.ml b/checker/checkInductive.ml
new file mode 100644
index 0000000000..4e026d6f60
--- /dev/null
+++ b/checker/checkInductive.ml
@@ -0,0 +1,271 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Sorts
+open Pp
+open Declarations
+open Environ
+open Names
+open CErrors
+open Univ
+open Util
+open Constr
+
+let check_kind env ar u =
+ match Constr.kind (snd (Reduction.dest_prod env ar)) with
+ | Sort (Sorts.Type u') when Univ.Universe.equal u' (Univ.Universe.make u) -> ()
+ | _ -> failwith "not the correct sort"
+
+let check_polymorphic_arity env params par =
+ let pl = par.template_param_levels in
+ let rec check_p env pl params =
+ let open Context.Rel.Declaration in
+ match pl, params with
+ Some u::pl, LocalAssum (na,ty)::params ->
+ check_kind env ty u;
+ check_p (push_rel (LocalAssum (na,ty)) env) pl params
+ | None::pl,d::params -> check_p (push_rel d env) pl params
+ | [], _ -> ()
+ | _ -> failwith "check_poly: not the right number of params" in
+ check_p env pl (List.rev params)
+
+let conv_ctxt_prefix env (ctx1:rel_context) ctx2 =
+ let rec chk env rctx1 rctx2 =
+ let open Context.Rel.Declaration in
+ match rctx1, rctx2 with
+ (LocalAssum (_,ty1) as d1)::rctx1', LocalAssum (_,ty2)::rctx2' ->
+ Reduction.conv env ty1 ty2;
+ chk (push_rel d1 env) rctx1' rctx2'
+ | (LocalDef (_,bd1,ty1) as d1)::rctx1', LocalDef (_,bd2,ty2)::rctx2' ->
+ Reduction.conv env ty1 ty2;
+ Reduction.conv env bd1 bd2;
+ chk (push_rel d1 env) rctx1' rctx2'
+ | [],_ -> ()
+ | _ -> failwith "non convertible contexts" in
+ chk env (List.rev ctx1) (List.rev ctx2)
+
+(* check information related to inductive arity *)
+let typecheck_arity env params inds =
+ let nparamargs = Context.Rel.nhyps params in
+ let nparamdecls = Context.Rel.length params in
+ let check_arity arctxt = function
+ | RegularArity mar ->
+ let ar = mar.mind_user_arity in
+ let _ = Typeops.infer_type env ar in
+ Reduction.conv env (Term.it_mkProd_or_LetIn (Constr.mkSort mar.mind_sort) arctxt) ar;
+ ar
+ | TemplateArity par ->
+ check_polymorphic_arity env params par;
+ Term.it_mkProd_or_LetIn (Constr.mkSort(Sorts.Type par.template_level)) arctxt
+ in
+ let env_arities =
+ Array.fold_left
+ (fun env_ar ind ->
+ let ar_ctxt = ind.mind_arity_ctxt in
+ let _ = Typeops.check_context env ar_ctxt in
+ conv_ctxt_prefix env params ar_ctxt;
+ (* Arities (with params) are typed-checked here *)
+ let arity = check_arity ar_ctxt ind.mind_arity in
+ (* mind_nrealargs *)
+ let nrealargs = Context.Rel.nhyps ar_ctxt - nparamargs in
+ if ind.mind_nrealargs <> nrealargs then
+ failwith "bad number of real inductive arguments";
+ let nrealargs_ctxt = Context.Rel.length ar_ctxt - nparamdecls in
+ if ind.mind_nrealdecls <> nrealargs_ctxt then
+ failwith "bad length of real inductive arguments signature";
+ (* We do not need to generate the universe of full_arity; if
+ later, after the validation of the inductive definition,
+ full_arity is used as argument or subject to cast, an
+ upper universe will be generated *)
+ let id = ind.mind_typename in
+ let env_ar' = push_rel (Context.Rel.Declaration.LocalAssum (Name id, arity)) env_ar in
+ env_ar')
+ env
+ inds in
+ let env_ar_par = push_rel_context params env_arities in
+ env_arities, env_ar_par
+
+(* Check that the subtyping information inferred for inductive types in the block is correct. *)
+(* This check produces a value of the unit type if successful or raises an anomaly if check fails. *)
+let check_subtyping cumi paramsctxt env arities =
+ let numparams = Context.Rel.nhyps paramsctxt in
+ (** In [env] we already have [ Var(0) ... Var(n-1) |- cst ] available.
+ We must produce the substitution σ : [ Var(i) -> Var (i + n) | 0 <= i < n]
+ and push the constraints [ Var(n) ... Var(2n - 1) |- cst{σ} ], together
+ with the cumulativity constraints [ cumul_cst ]. *)
+ let uctx = ACumulativityInfo.univ_context cumi in
+ let len = AUContext.size uctx in
+ let inst = Instance.of_array @@ Array.init len (fun i -> Level.var (i + len)) in
+
+ let other_context = ACumulativityInfo.univ_context cumi in
+ let uctx_other = UContext.make (inst, AUContext.instantiate inst other_context) in
+ let cumul_cst =
+ Array.fold_left_i (fun i csts var ->
+ match var with
+ | Variance.Irrelevant -> csts
+ | Variance.Covariant -> Constraint.add (Level.var i,Le,Level.var (i+len)) csts
+ | Variance.Invariant -> Constraint.add (Level.var i,Eq,Level.var (i+len)) csts)
+ Constraint.empty (ACumulativityInfo.variance cumi)
+ in
+ let env = Environ.push_context uctx_other env in
+ let env = Environ.add_constraints cumul_cst env in
+ let dosubst = Vars.subst_instance_constr inst in
+ (* process individual inductive types: *)
+ Array.iter (fun { mind_user_lc = lc; mind_arity = arity } ->
+ match arity with
+ | RegularArity { mind_user_arity = full_arity} ->
+ Indtypes.check_subtyping_arity_constructor env dosubst full_arity numparams true;
+ Array.iter (fun cnt -> Indtypes.check_subtyping_arity_constructor env dosubst cnt numparams false) lc
+ | TemplateArity _ ->
+ anomaly ~label:"check_subtyping"
+ Pp.(str "template polymorphism and cumulative polymorphism are not compatible")
+ ) arities
+
+(* 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 *)
+ | [|constrinfos|] -> Univ.is_type0m_univ constrinfos
+ | [||] -> (* type without constructors *) true
+ | _ -> false
+
+let small_unit constrsinfos =
+ let issmall = Array.for_all Univ.is_small_univ constrsinfos
+ and isunit = is_unit constrsinfos in
+ issmall, isunit
+
+let all_sorts = [InProp;InSet;InType]
+let small_sorts = [InProp;InSet]
+let logical_sorts = [InProp]
+
+let allowed_sorts issmall isunit s =
+ match Sorts.family s with
+ (* Type: all elimination allowed *)
+ | InType -> all_sorts
+
+ (* Small Set is predicative: all elimination allowed *)
+ | InSet when issmall -> all_sorts
+
+ (* Large Set is necessarily impredicative: forbids large elimination *)
+ | InSet -> small_sorts
+
+ (* Unitary/empty Prop: elimination to all sorts are realizable *)
+ (* unless the type is large. If it is large, forbids large elimination *)
+ (* which otherwise allows simulating the inconsistent system Type:Type *)
+ | InProp when isunit -> if issmall then all_sorts else small_sorts
+
+ (* Other propositions: elimination only to Prop *)
+ | InProp -> logical_sorts
+
+let check_predicativity env s small level =
+ match s, engagement env with
+ Type u, _ ->
+ (* let u' = fresh_local_univ () in *)
+ (* let cst = *)
+ (* merge_constraints (enforce_leq u u' empty_constraint) *)
+ (* (universes env) in *)
+ if not (UGraph.check_leq (universes env) level u) then
+ failwith "impredicative Type inductive type"
+ | Set, ImpredicativeSet -> ()
+ | Set, _ ->
+ if not small then failwith "impredicative Set inductive type"
+ | Prop,_ -> ()
+
+let sort_of_ind = function
+ | RegularArity mar -> mar.mind_sort
+ | TemplateArity par -> Type par.template_level
+
+let compute_elim_sorts env_ar params arity lc =
+ let inst = Context.Rel.to_extended_list Constr.mkRel 0 params in
+ let env_params = push_rel_context params env_ar in
+ let lc = Array.map
+ (fun c ->
+ Reduction.hnf_prod_applist env_params (Vars.lift (Context.Rel.length params) c) inst)
+ lc in
+ let s = sort_of_ind arity in
+ let infos = Array.map (Indtypes.infos_and_sort env_params) lc in
+ let (small,unit) = small_unit infos in
+ (* We accept recursive unit types... *)
+ (* compute the max of the sorts of the products of the constructor type *)
+ let min = if Array.length lc > 1 then Universe.type0 else Universe.type0m in
+ let level = Array.fold_left (fun max l -> Universe.sup max l) min infos in
+ check_predicativity env_ar s small level;
+ allowed_sorts small unit s
+
+let typecheck_one_inductive env params mip =
+ (* mind_typename and mind_consnames not checked *)
+ (* mind_reloc_tbl, mind_nb_constant, mind_nb_args not checked (VM) *)
+ (* mind_arity_ctxt, mind_arity, mind_nrealargs DONE (typecheck_arity) *)
+ (* mind_user_lc *)
+ let _ = Array.map (Typeops.infer_type env) mip.mind_user_lc in
+ (* mind_nf_lc *)
+ let _ = Array.map (Typeops.infer_type env) mip.mind_nf_lc in
+ Array.iter2 (Reduction.conv env) mip.mind_nf_lc mip.mind_user_lc;
+ (* mind_consnrealdecls *)
+ let check_cons_args c n =
+ let ctx,_ = Term.decompose_prod_assum c in
+ if n <> Context.Rel.length ctx - Context.Rel.length params then
+ failwith "bad number of real constructor arguments" in
+ Array.iter2 check_cons_args mip.mind_nf_lc mip.mind_consnrealdecls;
+ (* mind_kelim: checked by positivity criterion ? *)
+ let sorts =
+ compute_elim_sorts env params mip.mind_arity mip.mind_nf_lc in
+ let reject_sort s = not (List.mem_f Sorts.family_equal s sorts) in
+ if List.exists reject_sort mip.mind_kelim then
+ failwith "elimination not allowed";
+ (* mind_recargs: checked by positivity criterion *)
+ ()
+
+let check_inductive env kn mib =
+ Flags.if_verbose Feedback.msg_notice (str " checking mutind block: " ++ MutInd.print kn);
+ (* check mind_constraints: should be consistent with env *)
+ let env0 =
+ match mib.mind_universes with
+ | Monomorphic_ind _ -> env
+ | Polymorphic_ind auctx ->
+ let uctx = Univ.AUContext.repr auctx in
+ Environ.push_context uctx env
+ | Cumulative_ind cumi ->
+ let uctx = Univ.AUContext.repr (Univ.ACumulativityInfo.univ_context cumi) in
+ Environ.push_context uctx env
+ in
+ (** Locally set the oracle for further typechecking *)
+ let env0 = Environ.set_oracle env0 mib.mind_typing_flags.conv_oracle in
+ (* check mind_record : TODO ? check #constructor = 1 ? *)
+ (* check mind_finite : always OK *)
+ (* check mind_ntypes *)
+ if Array.length mib.mind_packets <> mib.mind_ntypes then
+ user_err Pp.(str "not the right number of packets");
+ (* check mind_params_ctxt *)
+ let params = mib.mind_params_ctxt in
+ let _ = Typeops.check_context env0 params in
+ (* check mind_nparams *)
+ if Context.Rel.nhyps params <> mib.mind_nparams then
+ user_err Pp.(str "number the right number of parameters");
+ (* mind_packets *)
+ (* - check arities *)
+ let env_ar, env_ar_par = typecheck_arity env0 params mib.mind_packets in
+ (* - check constructor types *)
+ Array.iter (typecheck_one_inductive env_ar params) mib.mind_packets;
+ (* check the inferred subtyping relation *)
+ let () =
+ match mib.mind_universes with
+ | Monomorphic_ind _ | Polymorphic_ind _ -> ()
+ | Cumulative_ind acumi ->
+ check_subtyping acumi params env_ar mib.mind_packets
+ in
+ (* check mind_nparams_rec: positivity condition *)
+ let packets = Array.map (fun p -> (p.mind_typename, Array.to_list p.mind_consnames, p.mind_user_lc, (p.mind_arity_ctxt,p.mind_arity))) mib.mind_packets in
+ let _ = Indtypes.check_positivity ~chkpos:true kn env_ar_par mib.mind_params_ctxt mib.mind_finite packets in
+ (* check mind_equiv... *)
+ (* Now we can add the inductive *)
+ add_mind kn mib env
diff --git a/checker/subtyping.mli b/checker/checkInductive.mli
index bb867186ba..17ca0d4583 100644
--- a/checker/subtyping.mli
+++ b/checker/checkInductive.mli
@@ -9,13 +9,10 @@
(************************************************************************)
(*i*)
-open Cic
+open Names
open Environ
(*i*)
-(** Invariant: the first [module_type_body] is now supposed to be
- known by [env] *)
-
-val check_subtypes : env -> module_type_body -> module_type_body -> unit
-
+(*s The following function does checks on inductive declarations. *)
+val check_inductive : env -> MutInd.t -> Declarations.mutual_inductive_body -> env
diff --git a/checker/checkTypes.ml b/checker/checkTypes.ml
new file mode 100644
index 0000000000..7eaa5eedd5
--- /dev/null
+++ b/checker/checkTypes.ml
@@ -0,0 +1,36 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+open Util
+open Term
+open Constr
+open Declarations
+open Reduction
+open Environ
+
+(* Polymorphic arities utils *)
+
+let check_kind env ar u =
+ match Constr.kind (snd (dest_prod env ar)) with
+ | Sort (Type u') when Univ.Universe.equal u' (Univ.Universe.make u) -> ()
+ | _ -> failwith "not the correct sort"
+
+let check_polymorphic_arity env params par =
+ let pl = par.template_param_levels in
+ let rec check_p env pl params =
+ let open Context.Rel.Declaration in
+ match pl, params with
+ Some u::pl, LocalAssum (na,ty)::params ->
+ check_kind env ty u;
+ check_p (push_rel (LocalAssum (na,ty)) env) pl params
+ | None::pl,d::params -> check_p (push_rel d env) pl params
+ | [], _ -> ()
+ | _ -> failwith "check_poly: not the right number of params" in
+ check_p env pl (List.rev params)
diff --git a/checker/typeops.mli b/checker/checkTypes.mli
index c2d7d19ce2..022f9bc603 100644
--- a/checker/typeops.mli
+++ b/checker/checkTypes.mli
@@ -9,14 +9,12 @@
(************************************************************************)
(*i*)
-open Cic
+open Declarations
+open Constr
open Environ
(*i*)
(*s Typing functions (not yet tagged as safe) *)
-val infer : env -> constr -> constr
-val infer_type : env -> constr -> sorts
-val check_ctxt : env -> rel_context -> env
val check_polymorphic_arity :
env -> rel_context -> template_arity -> unit
diff --git a/checker/check_stat.ml b/checker/check_stat.ml
index 3f00f924e9..57adc79475 100644
--- a/checker/check_stat.ml
+++ b/checker/check_stat.ml
@@ -10,7 +10,6 @@
open Pp
open Names
-open Cic
open Declarations
open Environ
@@ -25,43 +24,32 @@ let print_memory_stat () =
let output_context = ref false
-let pr_engagement impr_set =
+let pr_engagement env =
begin
- match impr_set with
+ match engagement env with
| ImpredicativeSet -> str "Theory: Set is impredicative"
| PredicativeSet -> str "Theory: Set is predicative"
end
-let cst_filter f csts =
- Cmap_env.fold
- (fun c ce acc -> if f c ce then c::acc else acc)
- csts []
+let is_ax _ cb = not (Declareops.constant_has_body cb)
-let is_ax _ cb = not (constant_has_body cb)
-
-let pr_ax csts =
- let axs = cst_filter is_ax csts in
+let pr_ax env =
+ let axs = fold_constants (fun c ce acc -> if is_ax c ce then c::acc else acc) env [] in
if axs = [] then
str "Axioms: <none>"
else
- hv 2 (str "Axioms:" ++ fnl() ++ prlist_with_sep fnl Indtypes.prcon axs)
+ hv 2 (str "Axioms:" ++ fnl() ++ prlist_with_sep fnl Constant.print axs)
let print_context env =
if !output_context then begin
- let
- {env_globals=
- {env_constants=csts; env_inductives=inds;
- env_modules=mods; env_modtypes=mtys};
- env_stratification=
- {env_universes=univ; env_engagement=engt}} = env in
Feedback.msg_notice
(hov 0
(fnl() ++ str"CONTEXT SUMMARY" ++ fnl() ++
str"===============" ++ fnl() ++ fnl() ++
- str "* " ++ hov 0 (pr_engagement engt ++ fnl()) ++ fnl() ++
- str "* " ++ hov 0 (pr_ax csts)));
+ str "* " ++ hov 0 (pr_engagement env ++ fnl()) ++ fnl() ++
+ str "* " ++ hov 0 (pr_ax env)));
end
-let stats () =
- print_context (Safe_typing.get_env());
+let stats env =
+ print_context env;
print_memory_stat ()
diff --git a/checker/check_stat.mli b/checker/check_stat.mli
index 823b107f5b..b094da1c44 100644
--- a/checker/check_stat.mli
+++ b/checker/check_stat.mli
@@ -7,8 +7,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-
val memory_stat : bool ref
val output_context : bool ref
-val stats : unit -> unit
+val stats : Environ.env -> unit
diff --git a/checker/checker.ml b/checker/checker.ml
index 8a5220c9ca..346ae5fffb 100644
--- a/checker/checker.ml
+++ b/checker/checker.ml
@@ -14,11 +14,10 @@ open Util
open System
open Names
open Check
+open Environ
let () = at_exit flush_all
-let pp_arrayi pp fmt a = Array.iteri (fun i x -> pp fmt (i,x)) a
-
let fatal_error info anomaly =
flush_all (); Format.eprintf "@[Fatal Error: @[%a@]@]@\n%!" Pp.pp_with info; flush_all ();
exit (if anomaly then 129 else 1)
@@ -137,9 +136,13 @@ let init_load_path () =
let set_debug () = Flags.debug := true
-let impredicative_set = ref Cic.PredicativeSet
-let set_impredicative_set () = impredicative_set := Cic.ImpredicativeSet
-let engage () = Safe_typing.set_engagement (!impredicative_set)
+let impredicative_set = ref Declarations.PredicativeSet
+let set_impredicative_set () = impredicative_set := Declarations.ImpredicativeSet
+let engage = Safe_typing.set_engagement (!impredicative_set)
+
+let disable_compilers senv =
+ let senv = Safe_typing.set_VM false senv in
+ Safe_typing.set_native_compiler false senv
let admit_list = ref ([] : object_file list)
@@ -158,8 +161,8 @@ let add_compile s =
We no longer use [Arg.parse], in order to use share [Usage.print_usage]
between coqtop and coqc. *)
-let compile_files () =
- Check.recheck_library
+let compile_files senv =
+ Check.recheck_library senv
~norec:(List.rev !norec_list)
~admit:(List.rev !admit_list)
~check:(List.rev !compile_list)
@@ -244,16 +247,13 @@ let explain_exn = function
hov 0 (anomaly_string () ++ str "uncaught exception Invalid_argument " ++ guill s ++ report ())
| Sys.Break ->
hov 0 (fnl () ++ str "User interrupt.")
- | Univ.AlreadyDeclared ->
- hov 0 (str "Error: Multiply declared universe.")
- | Univ.UniverseInconsistency (o,u,v) ->
- let msg =
- if !Flags.debug (*!Constrextern.print_universes*) then
- spc() ++ str "(cannot enforce" ++ spc() ++ Univ.pr_uni u ++ spc() ++
- str (match o with Univ.Lt -> "<" | Univ.Le -> "<=" | Univ.Eq -> "=")
- ++ spc() ++ Univ.pr_uni v ++ str")"
- else
- mt() in
+ | Univ.UniverseInconsistency i ->
+ let msg =
+ if !Flags.debug then
+ str "." ++ spc() ++
+ Univ.explain_universe_inconsistency Univ.Level.pr i
+ else
+ mt() in
hov 0 (str "Error: Universe inconsistency" ++ msg ++ str ".")
| TypeError(ctx,te) ->
hov 0 (str "Type error: " ++
@@ -270,31 +270,29 @@ let explain_exn = function
| IllFormedBranch _ -> str"IllFormedBranch"
| Generalization _ -> str"Generalization"
| ActualType _ -> str"ActualType"
- | CantApplyBadType ((n,a,b),(hd,hdty),args) ->
- (* This mix of printf / pp was here before... *)
- let fmt = Format.std_formatter in
- let open Format in
- let open Print in
- fprintf fmt "====== ill-typed term ====@\n";
- fprintf fmt "@[<hov 2>application head=@ %a@]@\n" print_pure_constr hd;
- fprintf fmt "@[<hov 2>head type=@ %a@]@\n" print_pure_constr hdty;
- let pp_arg fmt (i,(t,ty)) = fprintf fmt "@[<hv>@[<1>arg %d=@ @[%a@]@]@,@[<1>type=@ @[%a@]@]@]@\n@," (i+1)
- print_pure_constr t print_pure_constr ty
+ | CantApplyBadType ((n,a,b),{uj_val = hd; uj_type = hdty},args) ->
+ let pp_arg i judge =
+ hv 1 (str"arg " ++ int (i+1) ++ str"= " ++
+ Constr.debug_print judge.uj_val ++
+ str ",type= " ++ Constr.debug_print judge.uj_type) ++ fnl ()
in
- fprintf fmt "arguments:@\n@[<hv>%a@]@\n" (pp_arrayi pp_arg) args;
- fprintf fmt "====== type error ====@\n";
- fprintf fmt "%a@\n" print_pure_constr b;
- fprintf fmt "is not convertible with@\n";
- fprintf fmt "%a@\n" print_pure_constr a;
- fprintf fmt "====== universes ====@\n";
- fprintf fmt "%a@\n%!" Pp.pp_with
- (Univ.pr_universes
- (ctx.Environ.env_stratification.Environ.env_universes));
+ Feedback.msg_notice (str"====== ill-typed term ====" ++ fnl () ++
+ hov 2 (str"application head= " ++ Constr.debug_print hd) ++ fnl () ++
+ hov 2 (str"head type= " ++ Constr.debug_print hdty) ++ fnl () ++
+ str"arguments:" ++ fnl () ++ hv 1 (prvecti pp_arg args));
+ Feedback.msg_notice (str"====== type error ====@" ++ fnl () ++
+ Constr.debug_print b ++ fnl () ++
+ str"is not convertible with" ++ fnl () ++
+ Constr.debug_print a ++ fnl ());
+ Feedback.msg_notice (str"====== universes ====" ++ fnl () ++
+ (UGraph.pr_universes Univ.Level.pr
+ (ctx.Environ.env_stratification.Environ.env_universes)));
str "CantApplyBadType at argument " ++ int n
| CantApplyNonFunctional _ -> str"CantApplyNonFunctional"
| IllFormedRecBody _ -> str"IllFormedRecBody"
| IllTypedRecBody _ -> str"IllTypedRecBody"
- | UnsatisfiedConstraints _ -> str"UnsatisfiedConstraints"))
+ | UnsatisfiedConstraints _ -> str"UnsatisfiedConstraints"
+ | UndeclaredUniverse _ -> str"UndeclaredUniverse"))
| Indtypes.InductiveError e ->
hov 0 (str "Error related to inductive types")
@@ -368,35 +366,34 @@ let parse_args argv =
parse (List.tl (Array.to_list argv))
-(* To prevent from doing the initialization twice *)
-let initialized = ref false
-
(* XXX: At some point we need to either port the checker to use the
feedback system or to remove its use completely. *)
let init_with_argv argv =
- if not !initialized then begin
- initialized := true;
- Sys.catch_break false; (* Ctrl-C is fatal during the initialisation *)
- let _fhandle = Feedback.(add_feeder (console_feedback_listener Format.err_formatter)) in
- try
- parse_args argv;
- if !Flags.debug then Printexc.record_backtrace true;
- Envars.set_coqlib ~fail:(fun x -> CErrors.user_err Pp.(str x));
- Flags.if_verbose print_header ();
- init_load_path ();
- engage ();
- with e ->
- fatal_error (str "Error during initialization :" ++ (explain_exn e)) (is_anomaly e)
- end
+ Sys.catch_break false; (* Ctrl-C is fatal during the initialisation *)
+ let _fhandle = Feedback.(add_feeder (console_feedback_listener Format.err_formatter)) in
+ try
+ parse_args argv;
+ if !Flags.debug then Printexc.record_backtrace true;
+ Envars.set_coqlib ~fail:(fun x -> CErrors.user_err Pp.(str x));
+ Flags.if_verbose print_header ();
+ init_load_path ();
+ let senv = Safe_typing.empty_environment in
+ disable_compilers (engage senv)
+ with e ->
+ fatal_error (str "Error during initialization :" ++ (explain_exn e)) (is_anomaly e)
let init() = init_with_argv Sys.argv
-let run () =
+let run senv =
try
- compile_files ();
- flush_all()
+ let senv = compile_files senv in
+ flush_all(); senv
with e ->
if !Flags.debug then Printexc.print_backtrace stderr;
fatal_error (explain_exn e) (is_anomaly e)
-let start () = init(); run(); Check_stat.stats(); exit 0
+let start () =
+ let senv = init() in
+ let senv = run senv in
+ Check_stat.stats (Safe_typing.env_of_safe_env senv);
+ exit 0
diff --git a/checker/cic.mli b/checker/cic.mli
deleted file mode 100644
index 754cc2a096..0000000000
--- a/checker/cic.mli
+++ /dev/null
@@ -1,471 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(** Type definitions for the Calculus of Inductive Constructions *)
-
-(** We regroup here the type definitions for structures of the Coq kernel
- that are present in .vo files. Here is everything the Checker needs
- to know about these structures for verifying a .vo. Note that this
- isn't an exact copy of the kernel code :
-
- - there isn't any abstraction here (see e.g. [constr] or [lazy_constr])
- - some types are left undefined when they aren't used by the Checker
- - some types have less constructors when the final constructors aren't
- supposed to appear in .vo (see [REVERTcast] and [Direct]).
-
- The following types are also described in a reified manner in values.ml,
- for validating the layout of structures after de-marshalling. So:
-
- IF YOU ADAPT THIS FILE, YOU SHOULD MODIFY values.ml ACCORDINGLY !
-*)
-
-open Names
-
-(*************************************************************************)
-(** {4 From term.ml} *)
-
-(** {6 The sorts of CCI. } *)
-
-type sorts =
- | Prop
- | Set
- | Type of Univ.universe
-
-(** {6 The sorts family of CCI. } *)
-
-type sorts_family = InProp | InSet | InType
-
-(** {6 Useful types } *)
-
-(** {6 Existential variables } *)
-type existential_key = int
-
-(** {6 Existential variables } *)
-type metavariable = int
-
-(** {6 Case annotation } *)
-type case_style = LetStyle | IfStyle | LetPatternStyle | MatchStyle
- | RegularStyle (** infer printing form from number of constructor *)
-type case_printing =
- { ind_tags : bool list; (* tell whether letin or lambda in the arity of the inductive type *)
- cstr_tags : bool list array; (* whether each pattern var of each constructor is a let-in (true) or not (false) *)
- style : case_style }
-
-(** the integer is the number of real args, needed for reduction *)
-type case_info =
- { ci_ind : inductive;
- ci_npar : int;
- ci_cstr_ndecls : int array; (* number of pattern vars of each constructor (with let's)*)
- ci_cstr_nargs : int array; (* number of pattern vars of each constructor (w/o let's) *)
- ci_pp_info : case_printing (** not interpreted by the kernel *)
- }
-
-(** This defines the strategy to use for verifiying a Cast. *)
-type cast_kind = VMcast | NATIVEcast | DEFAULTcast (* | REVERTcast *)
-
-(** {6 The type of constructions } *)
-
-(** [constr array] is an instance matching definitional [named_context] in
- the same order (i.e. last argument first) *)
-type 'constr pexistential = existential_key * 'constr array
-type 'constr prec_declaration =
- Name.t array * 'constr array * 'constr array
-type 'constr pfixpoint =
- (int array * int) * 'constr prec_declaration
-type 'constr pcofixpoint =
- int * 'constr prec_declaration
-type 'a puniverses = 'a Univ.puniverses
-type pconstant = Constant.t puniverses
-type pinductive = inductive puniverses
-type pconstructor = constructor puniverses
-
-type constr =
- | Rel of int
- | Var of Id.t (** Shouldn't occur in a .vo *)
- | Meta of metavariable (** Shouldn't occur in a .vo *)
- | Evar of constr pexistential (** Shouldn't occur in a .vo *)
- | Sort of sorts
- | Cast of constr * cast_kind * constr
- | Prod of Name.t * constr * constr
- | Lambda of Name.t * constr * constr
- | LetIn of Name.t * constr * constr * constr
- | App of constr * constr array
- | Const of pconstant
- | Ind of pinductive
- | Construct of pconstructor
- | Case of case_info * constr * constr * constr array
- | Fix of constr pfixpoint
- | CoFix of constr pcofixpoint
- | Proj of Projection.t * constr
-
-type existential = constr pexistential
-type rec_declaration = constr prec_declaration
-type fixpoint = constr pfixpoint
-type cofixpoint = constr pcofixpoint
-
-(** {6 Type of assumptions and contexts} *)
-
-type rel_declaration = LocalAssum of Name.t * constr (* name, type *)
- | LocalDef of Name.t * constr * constr (* name, value, type *)
-type rel_context = rel_declaration list
-
-(** The declarations below in .vo should be outside sections,
- so we expect there a value compatible with an empty list *)
-type section_context = unit
-
-
-(*************************************************************************)
-(** {4 From mod_susbt.ml and lazyconstr.ml} *)
-
-(** {6 Substitutions} *)
-
-type delta_hint =
- | Inline of int * (Univ.AUContext.t * constr) option
- | Equiv of KerName.t
-
-type delta_resolver = ModPath.t MPmap.t * delta_hint KNmap.t
-
-type 'a umap_t = 'a MPmap.t
-type substitution = (ModPath.t * delta_resolver) umap_t
-
-(** {6 Delayed constr} *)
-
-type 'a substituted = {
- mutable subst_value : 'a;
- mutable subst_subst : substitution list;
-}
-
-type constr_substituted = constr substituted
-
-(** Nota : in coqtop, the [lazy_constr] type also have a [Direct]
- constructor, but it shouldn't occur inside a .vo, so we ignore it *)
-
-type lazy_constr =
- | Indirect of substitution list * DirPath.t * int
-(* | Direct of constr_substituted *)
-
-
-(*************************************************************************)
-(** {4 From declarations.mli} *)
-
-(** Some types unused in the checker, hence left undefined *)
-
-(** Bytecode *)
-type reloc_table
-type to_patch_substituted
-(** Native code *)
-type native_name
-(** Retroknowledge *)
-type action
-
-(** Engagements *)
-
-type set_predicativity = ImpredicativeSet | PredicativeSet
-
-type engagement = set_predicativity
-
-(** {6 Conversion oracle} *)
-
-type level = Expand | Level of int | Opaque
-
-type oracle = {
- var_opacity : level Id.Map.t;
- cst_opacity : level Cmap.t;
- var_trstate : Id.Pred.t;
- cst_trstate : Cpred.t;
-}
-
-(** {6 Representation of constants (Definition/Axiom) } *)
-
-
-type template_arity = {
- template_param_levels : Univ.universe_level option list;
- template_level : Univ.universe;
-}
-
-type ('a, 'b) declaration_arity =
- | RegularArity of 'a
- | TemplateArity of 'b
-
-(** Inlining level of parameters at functor applications.
- This is ignored by the checker. *)
-
-type inline = int option
-
-(** A constant can have no body (axiom/parameter), or a
- transparent body, or an opaque one *)
-
-type constant_def =
- | Undef of inline
- | Def of constr_substituted
- | OpaqueDef of lazy_constr
-
-type constant_universes =
- | Monomorphic_const of Univ.ContextSet.t
- | Polymorphic_const of Univ.abstract_universe_context
-
-(** The [typing_flags] are instructions to the type-checker which
- modify its behaviour. The typing flags used in the type-checking
- of a constant are tracked in their {!constant_body} so that they
- can be displayed to the user. *)
-type typing_flags = {
- check_guarded : bool; (** If [false] then fixed points and co-fixed
- points are assumed to be total. *)
- check_universes : bool; (** If [false] universe constraints are not checked *)
- conv_oracle : oracle; (** Unfolding strategies for conversion *)
- share_reduction : bool; (** Use by-need reduction algorithm *)
- enable_VM : bool; (** If [false], all VM conversions fall back to interpreted ones *)
- enable_native_compiler : bool; (** If [false], all native conversions fall back to VM ones *)
-}
-
-type constant_body = {
- const_hyps : section_context; (** New: younger hyp at top *)
- const_body : constant_def;
- const_type : constr;
- const_body_code : to_patch_substituted;
- const_universes : constant_universes;
- const_inline_code : bool;
- const_typing_flags : typing_flags;
-}
-
-(** {6 Representation of mutual inductive types } *)
-
-type recarg =
- | Norec
- | Mrec of inductive
- | Imbr of inductive
-
-type wf_paths = recarg Rtree.t
-
-type record_info =
-| NotRecord
-| FakeRecord
-| PrimRecord of (Id.t * Label.t array * constr array) array
-
-type regular_inductive_arity = {
- mind_user_arity : constr;
- mind_sort : sorts;
-}
-
-type recursivity_kind =
- | Finite (** = inductive *)
- | CoFinite (** = coinductive *)
- | BiFinite (** = non-recursive, like in "Record" definitions *)
-
-type inductive_arity = (regular_inductive_arity, template_arity) declaration_arity
-
-type one_inductive_body = {
-(** {8 Primitive datas } *)
-
- mind_typename : Id.t; (** Name of the type: [Ii] *)
-
- mind_arity_ctxt : rel_context; (** Arity context of [Ii] with parameters: [forall params, Ui] *)
-
- mind_arity : inductive_arity; (** Arity sort and original user arity if monomorphic *)
-
- mind_consnames : Id.t array; (** Names of the constructors: [cij] *)
-
- mind_user_lc : constr array;
- (** Types of the constructors with parameters: [forall params, Tij],
- where the Ik are replaced by de Bruijn index in the
- context I1:forall params, U1 .. In:forall params, Un *)
-
-(** {8 Derived datas } *)
-
- mind_nrealargs : int; (** Number of expected real arguments of the type (no let, no params) *)
-
- mind_nrealdecls : int; (** Length of realargs context (with let, no params) *)
-
- mind_kelim : sorts_family list; (** List of allowed elimination sorts *)
-
- mind_nf_lc : constr array; (** Head normalized constructor types so that their conclusion is atomic *)
-
- mind_consnrealargs : int array;
- (** Length of the signature of the constructors (w/o let, w/o params)
- (not used in the kernel) *)
-
- mind_consnrealdecls : int array;
- (** Length of the signature of the constructors (with let, w/o params)
- (not used in the kernel) *)
-
- mind_recargs : wf_paths; (** Signature of recursive arguments in the constructors *)
-
-(** {8 Datas for bytecode compilation } *)
-
- mind_nb_constant : int; (** number of constant constructor *)
-
- mind_nb_args : int; (** number of no constant constructor *)
-
- mind_reloc_tbl : reloc_table;
- }
-
-type abstract_inductive_universes =
- | Monomorphic_ind of Univ.ContextSet.t
- | Polymorphic_ind of Univ.abstract_universe_context
- | Cumulative_ind of Univ.abstract_cumulativity_info
-
-type mutual_inductive_body = {
-
- mind_packets : one_inductive_body array; (** The component of the mutual inductive block *)
-
- mind_record : record_info; (** Whether the inductive type has been declared as a record. *)
-
- mind_finite : recursivity_kind; (** Whether the type is inductive or coinductive *)
-
- mind_ntypes : int; (** Number of types in the block *)
-
- mind_hyps : section_context; (** Section hypotheses on which the block depends *)
-
- mind_nparams : int; (** Number of expected parameters *)
-
- mind_nparams_rec : int; (** Number of recursively uniform (i.e. ordinary) parameters *)
-
- mind_params_ctxt : rel_context; (** The context of parameters (includes let-in declaration) *)
-
- mind_universes : abstract_inductive_universes; (** Local universe variables and constraints together with subtyping constraints *)
-
- mind_private : bool option; (** allow pattern-matching: Some true ok, Some false blocked *)
-
- mind_typing_flags : typing_flags; (** typing flags at the time of the inductive creation *)
- }
-
-(** {6 Module declarations } *)
-
-(** Functor expressions are forced to be on top of other expressions *)
-
-type ('ty,'a) functorize =
- | NoFunctor of 'a
- | MoreFunctor of MBId.t * 'ty * ('ty,'a) functorize
-
-(** The fully-algebraic module expressions : names, applications, 'with ...'.
- They correspond to the user entries of non-interactive modules.
- They will be later expanded into module structures in [Mod_typing],
- and won't play any role into the kernel after that : they are kept
- only for short module printing and for extraction. *)
-
-type with_declaration
-
-type module_alg_expr =
- | MEident of ModPath.t
- | MEapply of module_alg_expr * ModPath.t
- | MEwith of module_alg_expr * with_declaration
-
-(** A component of a module structure *)
-
-type structure_field_body =
- | SFBconst of constant_body
- | SFBmind of mutual_inductive_body
- | SFBmodule of module_body
- | SFBmodtype of module_type_body
-
-(** A module structure is a list of labeled components.
-
- Note : we may encounter now (at most) twice the same label in
- a [structure_body], once for a module ([SFBmodule] or [SFBmodtype])
- and once for an object ([SFBconst] or [SFBmind]) *)
-
-and structure_body = (Label.t * structure_field_body) list
-
-(** A module signature is a structure, with possibly functors on top of it *)
-
-and module_signature = (module_type_body,structure_body) functorize
-
-(** A module expression is an algebraic expression, possibly functorized. *)
-
-and module_expression = (module_type_body,module_alg_expr) functorize
-
-and module_implementation =
- | Abstract (** no accessible implementation (keep this constructor first!) *)
- | Algebraic of module_expression (** non-interactive algebraic expression *)
- | Struct of module_signature (** interactive body *)
- | FullStruct (** special case of [Struct] : the body is exactly [mod_type] *)
-
-and 'a generic_module_body =
- { mod_mp : ModPath.t; (** absolute path of the module *)
- mod_expr : 'a; (** implementation *)
- mod_type : module_signature; (** expanded type *)
- (** algebraic type, kept if it's relevant for extraction *)
- mod_type_alg : module_expression option;
- (** set of all constraints in the module *)
- mod_constraints : Univ.ContextSet.t;
- (** quotiented set of equivalent constants and inductive names *)
- mod_delta : delta_resolver;
- mod_retroknowledge : 'a module_retroknowledge; }
-
-and module_body = module_implementation generic_module_body
-
-(** A [module_type_body] is just a [module_body] with no
- implementation and also an empty [mod_retroknowledge] *)
-
-and module_type_body = unit generic_module_body
-
-and _ module_retroknowledge =
-| ModBodyRK :
- action list -> module_implementation module_retroknowledge
-| ModTypeRK : unit module_retroknowledge
-
-(*************************************************************************)
-(** {4 From safe_typing.ml} *)
-
-type nativecode_symb_array
-
-type compilation_unit_name = DirPath.t
-
-type vodigest =
- | Dvo of Digest.t (* The digest of the seg_lib part *)
- | Dviovo of Digest.t * Digest.t (* The digest of the seg_lib+seg_univ part *)
-
-type library_info = compilation_unit_name * vodigest
-
-type library_deps = library_info array
-
-type compiled_library = {
- comp_name : compilation_unit_name;
- comp_mod : module_body;
- comp_deps : library_deps;
- comp_enga : engagement;
- comp_natsymbs : nativecode_symb_array
-}
-
-
-(*************************************************************************)
-(** {4 From library.ml} *)
-
-type library_objects
-
-type summary_disk = {
- md_name : compilation_unit_name;
- md_imports : compilation_unit_name array;
- md_deps : library_deps;
-}
-
-type library_disk = {
- md_compiled : compiled_library;
- md_objects : library_objects;
-}
-
-type opaque_table = constr Future.computation array
-type univ_table =
- (Univ.universe_context_set Future.computation array * Univ.universe_context_set * bool) option
-
-(** A .vo file is currently made of :
-
- 1) a magic number (4 bytes, cf output_binary_int)
- 2) a marshalled [library_disk] structure
- 3) a [Digest.t] string (16 bytes)
- 4) a marshalled [univ_table] (* Some if vo was obtained from vi *)
- 5) a [Digest.t] string (16 bytes)
- 6) a marshalled [None] discharge_table (* Some in vi files *)
- 7) a [Digest.t] string (16 bytes)
- 8) a marshalled [None] todo_table (* Some in vi files *)
- 9) a [Digest.t] string (16 bytes)
- 10) a marshalled [opaque_table]
- 11) a [Digest.t] string (16 bytes)
-*)
diff --git a/checker/closure.ml b/checker/closure.ml
deleted file mode 100644
index 138499b0e6..0000000000
--- a/checker/closure.ml
+++ /dev/null
@@ -1,791 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-open Util
-open Pp
-open Names
-open Cic
-open Term
-open Esubst
-open Environ
-
-let stats = ref false
-let share = ref true
-
-(* Profiling *)
-let beta = ref 0
-let delta = ref 0
-let zeta = ref 0
-let evar = ref 0
-let iota = ref 0
-let prune = ref 0
-
-let reset () =
- beta := 0; delta := 0; zeta := 0; evar := 0; iota := 0; prune := 0
-
-let stop() =
- Feedback.msg_debug (str "[Reds: beta=" ++ int !beta ++ str" delta=" ++ int !delta ++
- str" zeta=" ++ int !zeta ++ str" evar=" ++ int !evar ++
- str" iota=" ++ int !iota ++ str" prune=" ++ int !prune ++ str"]")
-
-let incr_cnt red cnt =
- if red then begin
- if !stats then incr cnt;
- true
- end else
- false
-
-let with_stats c =
- if !stats then begin
- reset();
- let r = Lazy.force c in
- stop();
- r
- end else
- Lazy.force c
-
-module type RedFlagsSig = sig
- type reds
- type red_kind
- val fBETA : red_kind
- val fDELTA : red_kind
- val fIOTA : red_kind
- val fZETA : red_kind
- val no_red : reds
- val red_add : reds -> red_kind -> reds
- val mkflags : red_kind list -> reds
- val red_set : reds -> red_kind -> bool
-end
-
-module RedFlags = (struct
-
- (* [r_const=(true,cl)] means all constants but those in [cl] *)
- (* [r_const=(false,cl)] means only those in [cl] *)
- (* [r_delta=true] just mean [r_const=(true,[])] *)
-
- type reds = {
- r_beta : bool;
- r_delta : bool;
- r_zeta : bool;
- r_evar : bool;
- r_iota : bool }
-
- type red_kind = BETA | DELTA | IOTA | ZETA
-
- let fBETA = BETA
- let fDELTA = DELTA
- let fIOTA = IOTA
- let fZETA = ZETA
- let no_red = {
- r_beta = false;
- r_delta = false;
- r_zeta = false;
- r_evar = false;
- r_iota = false }
-
- let red_add red = function
- | BETA -> { red with r_beta = true }
- | DELTA -> { red with r_delta = true}
- | IOTA -> { red with r_iota = true }
- | ZETA -> { red with r_zeta = true }
-
- let mkflags = List.fold_left red_add no_red
-
- let red_set red = function
- | BETA -> incr_cnt red.r_beta beta
- | ZETA -> incr_cnt red.r_zeta zeta
- | IOTA -> incr_cnt red.r_iota iota
- | DELTA -> (* Used for Rel/Var defined in context *)
- incr_cnt red.r_delta delta
-
-end : RedFlagsSig)
-
-open RedFlags
-
-let betadeltaiota = mkflags [fBETA;fDELTA;fZETA;fIOTA]
-let betadeltaiotanolet = mkflags [fBETA;fDELTA;fIOTA]
-let betaiotazeta = mkflags [fBETA;fIOTA;fZETA]
-
-(* specification of the reduction function *)
-
-
-(* Flags of reduction and cache of constants: 'a is a type that may be
- * mapped to constr. 'a infos implements a cache for constants and
- * abstractions, storing a representation (of type 'a) of the body of
- * this constant or abstraction.
- * * i_tab is the cache table of the results
- * * i_rels = (4,[(1,c);(3,d)]) means there are 4 free rel variables
- * and only those with index 1 and 3 have bodies which are c and d resp.
- *
- * ref_value_cache searchs in the tab, otherwise uses i_repr to
- * compute the result and store it in the table. If the constant can't
- * be unfolded, returns None, but does not store this failure. * This
- * doesn't take the RESET into account. You mustn't keep such a table
- * after a Reset. * This type is not exported. Only its two
- * instantiations (cbv or lazy) are.
- *)
-
-type table_key = Constant.t puniverses tableKey
-
-
-let eq_pconstant_key (c,u) (c',u') =
- eq_constant_key c c' && Univ.Instance.equal u u'
-
-module KeyHash =
-struct
- type t = table_key
- let equal = Names.eq_table_key eq_pconstant_key
-
- open Hashset.Combine
-
- let hash = function
- | ConstKey (c,u) -> combinesmall 1 (Constant.UserOrd.hash c)
- | VarKey id -> combinesmall 2 (Id.hash id)
- | RelKey i -> combinesmall 3 (Int.hash i)
-end
-
-module KeyTable = Hashtbl.Make(KeyHash)
-
-let defined_rels flags env =
-(* if red_local_const (snd flags) then*)
- fold_rel_context
- (fun decl (i,subs) ->
- match decl with
- | LocalAssum _ -> (i+1, subs)
- | LocalDef (_,body,_) -> (i+1, (i,body) :: subs))
- (rel_context env) ~init:(0,[])
-(* else (0,[])*)
-
-(**********************************************************************)
-(* Lazy reduction: the one used in kernel operations *)
-
-(* type of shared terms. fconstr and frterm are mutually recursive.
- * Clone of the constr structure, but completely mutable, and
- * annotated with reduction state (reducible or not).
- * - FLIFT is a delayed shift; allows sharing between 2 lifted copies
- * of a given term.
- * - FCLOS is a delayed substitution applied to a constr
- * - FLOCKED is used to erase the content of a reference that must
- * be updated. This is to allow the garbage collector to work
- * before the term is computed.
- *)
-
-(* Norm means the term is fully normalized and cannot create a redex
- when substituted
- Cstr means the term is in head normal form and that it can
- create a redex when substituted (i.e. constructor, fix, lambda)
- Whnf means we reached the head normal form and that it cannot
- create a redex when substituted
- Red is used for terms that might be reduced
-*)
-type red_state = Norm | Cstr | Whnf | Red
-
-let neutr = function
- | (Whnf|Norm) -> Whnf
- | (Red|Cstr) -> Red
-
-type fconstr = {
- mutable norm: red_state;
- mutable term: fterm }
-
-and fterm =
- | FRel of int
- | FAtom of constr (* Metas and Sorts *)
- | FCast of fconstr * cast_kind * fconstr
- | FFlex of table_key
- | FInd of pinductive
- | FConstruct of pconstructor
- | FApp of fconstr * fconstr array
- | FProj of Projection.t * fconstr
- | FFix of fixpoint * fconstr subs
- | FCoFix of cofixpoint * fconstr subs
- | FCaseT of case_info * constr * fconstr * constr array * fconstr subs (* predicate and branches are closures *)
- | FLambda of int * (Name.t * constr) list * constr * fconstr subs
- | FProd of Name.t * fconstr * fconstr
- | FLetIn of Name.t * fconstr * fconstr * constr * fconstr subs
- | FEvar of existential_key * fconstr array (* why diff from kernel/closure? *)
- | FLIFT of int * fconstr
- | FCLOS of constr * fconstr subs
- | FLOCKED
-
-type clos_infos = {
- i_flags : reds;
- i_env : env;
- i_rels : int * (int * constr) list;
- i_tab : fconstr KeyTable.t }
-
-let fterm_of v = v.term
-let set_norm v = v.norm <- Norm
-
-(* Could issue a warning if no is still Red, pointing out that we loose
- sharing. *)
-let update v1 (no,t) =
- if !share then
- (v1.norm <- no;
- v1.term <- t;
- v1)
- else {norm=no;term=t}
-
-(**********************************************************************)
-(* The type of (machine) stacks (= lambda-bar-calculus' contexts) *)
-
-type stack_member =
- | Zapp of fconstr array
- | ZcaseT of case_info * constr * constr array * fconstr subs
- | Zproj of Projection.Repr.t
- | Zfix of fconstr * stack
- | Zshift of int
- | Zupdate of fconstr
-
-and stack = stack_member list
-
-let append_stack v s =
- if Array.length v = 0 then s else
- match s with
- | Zapp l :: s -> Zapp (Array.append v l) :: s
- | _ -> Zapp v :: s
-
-(* Collapse the shifts in the stack *)
-let zshift n s =
- match (n,s) with
- (0,_) -> s
- | (_,Zshift(k)::s) -> Zshift(n+k)::s
- | _ -> Zshift(n)::s
-
-let rec stack_args_size = function
- | Zapp v :: s -> Array.length v + stack_args_size s
- | Zshift(_)::s -> stack_args_size s
- | Zupdate(_)::s -> stack_args_size s
- | _ -> 0
-
-(* Lifting. Preserves sharing (useful only for cell with norm=Red).
- lft_fconstr always create a new cell, while lift_fconstr avoids it
- when the lift is 0. *)
-let rec lft_fconstr n ft =
- match ft.term with
- | (FInd _|FConstruct _|FFlex(ConstKey _|VarKey _)) -> ft
- | FRel i -> {norm=Norm;term=FRel(i+n)}
- | FLambda(k,tys,f,e) -> {norm=Cstr; term=FLambda(k,tys,f,subs_shft(n,e))}
- | FFix(fx,e) -> {norm=Cstr; term=FFix(fx,subs_shft(n,e))}
- | FCoFix(cfx,e) -> {norm=Cstr; term=FCoFix(cfx,subs_shft(n,e))}
- | FLIFT(k,m) -> lft_fconstr (n+k) m
- | FLOCKED -> assert false
- | _ -> {norm=ft.norm; term=FLIFT(n,ft)}
-let lift_fconstr k f =
- if k=0 then f else lft_fconstr k f
-let lift_fconstr_vect k v =
- if k=0 then v else Array.map (fun f -> lft_fconstr k f) v
-
-let clos_rel e i =
- match expand_rel i e with
- | Inl(n,mt) -> lift_fconstr n mt
- | Inr(k,None) -> {norm=Norm; term= FRel k}
- | Inr(k,Some p) ->
- lift_fconstr (k-p) {norm=Red;term=FFlex(RelKey p)}
-
-(* since the head may be reducible, we might introduce lifts of 0 *)
-let compact_stack head stk =
- let rec strip_rec depth = function
- | Zshift(k)::s -> strip_rec (depth+k) s
- | Zupdate(m)::s ->
- (* Be sure to create a new cell otherwise sharing would be
- lost by the update operation *)
- let h' = lft_fconstr depth head in
- let _ = update m (h'.norm,h'.term) in
- strip_rec depth s
- | stk -> zshift depth stk in
- strip_rec 0 stk
-
-(* Put an update mark in the stack, only if needed *)
-let zupdate m s =
- if !share && m.norm = Red
- then
- let s' = compact_stack m s in
- let _ = m.term <- FLOCKED in
- Zupdate(m)::s'
- else s
-
-let mk_lambda env t =
- let (rvars,t') = decompose_lam t in
- FLambda(List.length rvars, List.rev rvars, t', env)
-
-let destFLambda clos_fun t =
- match t.term with
- FLambda(_,[(na,ty)],b,e) -> (na,clos_fun e ty,clos_fun (subs_lift e) b)
- | FLambda(n,(na,ty)::tys,b,e) ->
- (na,clos_fun e ty,{norm=Cstr;term=FLambda(n-1,tys,b,subs_lift e)})
- | _ -> assert false
-
-(* Optimization: do not enclose variables in a closure.
- Makes variable access much faster *)
-let mk_clos e t =
- match t with
- | Rel i -> clos_rel e i
- | Var x -> { norm = Red; term = FFlex (VarKey x) }
- | Const c -> { norm = Red; term = FFlex (ConstKey c) }
- | Meta _ | Sort _ -> { norm = Norm; term = FAtom t }
- | Ind kn -> { norm = Norm; term = FInd kn }
- | Construct kn -> { norm = Cstr; term = FConstruct kn }
- | (CoFix _|Lambda _|Fix _|Prod _|Evar _|App _|Case _|Cast _|LetIn _|Proj _) ->
- {norm = Red; term = FCLOS(t,e)}
-
-let mk_clos_vect env v = Array.map (mk_clos env) v
-
-let inject = mk_clos (subs_id 0)
-
-let ref_value_cache info ref =
- try
- Some (KeyTable.find info.i_tab ref)
- with Not_found ->
- try
- let body =
- match ref with
- | RelKey n ->
- let (s,l) = info.i_rels in lift n (Int.List.assoc (s-n) l)
- | VarKey id -> raise Not_found
- | ConstKey cst -> constant_value info.i_env cst
- in
- let v = inject body in
- KeyTable.add info.i_tab ref v;
- Some v
- with
- | Not_found (* List.assoc *)
- | NotEvaluableConst _ (* Const *)
- -> None
-
-let mind_equiv_infos info = mind_equiv info.i_env
-
-(* Translate the head constructor of t from constr to fconstr. This
- function is parameterized by the function to apply on the direct
- subterms.
- Could be used insted of mk_clos. *)
-let mk_clos_deep clos_fun env t =
- match t with
- | (Rel _|Ind _|Const _|Construct _|Var _|Meta _ | Sort _) ->
- mk_clos env t
- | Cast (a,k,b) ->
- { norm = Red;
- term = FCast (clos_fun env a, k, clos_fun env b)}
- | App (f,v) ->
- { norm = Red;
- term = FApp (clos_fun env f, Array.map (clos_fun env) v) }
- | Proj (p,c) ->
- { norm = Red;
- term = FProj (p, clos_fun env c) }
- | Case (ci,p,c,v) ->
- { norm = Red; term = FCaseT (ci, p, clos_fun env c, v, env) }
- | Fix fx ->
- { norm = Cstr; term = FFix (fx, env) }
- | CoFix cfx ->
- { norm = Cstr; term = FCoFix(cfx,env) }
- | Lambda _ ->
- { norm = Cstr; term = mk_lambda env t }
- | Prod (n,t,c) ->
- { norm = Whnf;
- term = FProd (n, clos_fun env t, clos_fun (subs_lift env) c) }
- | LetIn (n,b,t,c) ->
- { norm = Red;
- term = FLetIn (n, clos_fun env b, clos_fun env t, c, env) }
- | Evar(ev,args) ->
- { norm = Whnf; term = FEvar(ev,Array.map (clos_fun env) args) }
-
-(* A better mk_clos? *)
-let mk_clos2 = mk_clos_deep mk_clos
-
-(* The inverse of mk_clos_deep: move back to constr *)
-let rec to_constr constr_fun lfts v =
- match v.term with
- | FRel i -> Rel (reloc_rel i lfts)
- | FFlex (RelKey p) -> Rel (reloc_rel p lfts)
- | FFlex (VarKey x) -> Var x
- | FAtom c -> exliftn lfts c
- | FCast (a,k,b) ->
- Cast (constr_fun lfts a, k, constr_fun lfts b)
- | FFlex (ConstKey op) -> Const op
- | FInd op -> Ind op
- | FConstruct op -> Construct op
- | FCaseT (ci,p,c,ve,e) ->
- let fp = mk_clos2 e p in
- let fve = mk_clos_vect e ve in
- Case (ci, constr_fun lfts fp, constr_fun lfts c, Array.map (constr_fun lfts) fve)
- | FFix ((op,(lna,tys,bds)),e) ->
- let n = Array.length bds in
- let ftys = Array.map (mk_clos e) tys in
- let fbds = Array.map (mk_clos (subs_liftn n e)) bds in
- let lfts' = el_liftn n lfts in
- Fix (op, (lna, Array.map (constr_fun lfts) ftys,
- Array.map (constr_fun lfts') fbds))
- | FCoFix ((op,(lna,tys,bds)),e) ->
- let n = Array.length bds in
- let ftys = Array.map (mk_clos e) tys in
- let fbds = Array.map (mk_clos (subs_liftn n e)) bds in
- let lfts' = el_liftn (Array.length bds) lfts in
- CoFix (op, (lna, Array.map (constr_fun lfts) ftys,
- Array.map (constr_fun lfts') fbds))
- | FApp (f,ve) ->
- App (constr_fun lfts f,
- Array.map (constr_fun lfts) ve)
- | FProj (p,c) ->
- Proj (p,constr_fun lfts c)
- | FLambda _ ->
- let (na,ty,bd) = destFLambda mk_clos2 v in
- Lambda (na, constr_fun lfts ty,
- constr_fun (el_lift lfts) bd)
- | FProd (n,t,c) ->
- Prod (n, constr_fun lfts t,
- constr_fun (el_lift lfts) c)
- | FLetIn (n,b,t,f,e) ->
- let fc = mk_clos2 (subs_lift e) f in
- LetIn (n, constr_fun lfts b,
- constr_fun lfts t,
- constr_fun (el_lift lfts) fc)
- | FEvar (ev,args) -> Evar(ev,Array.map (constr_fun lfts) args)
- | FLIFT (k,a) -> to_constr constr_fun (el_shft k lfts) a
- | FCLOS (t,env) ->
- let fr = mk_clos2 env t in
- let unfv = update v (fr.norm,fr.term) in
- to_constr constr_fun lfts unfv
- | FLOCKED -> assert false (*mkVar(Id.of_string"_LOCK_")*)
-
-(* This function defines the correspondance between constr and
- fconstr. When we find a closure whose substitution is the identity,
- then we directly return the constr to avoid possibly huge
- reallocation. *)
-let term_of_fconstr =
- let rec term_of_fconstr_lift lfts v =
- match v.term with
- | FCLOS(t,env) when is_subs_id env && is_lift_id lfts -> t
- | FLambda(_,tys,f,e) when is_subs_id e && is_lift_id lfts ->
- compose_lam (List.rev tys) f
- | FCaseT(ci,p,c,b,env) when is_subs_id env && is_lift_id lfts ->
- Case(ci,p,term_of_fconstr_lift lfts c,b)
- | FFix(fx,e) when is_subs_id e && is_lift_id lfts -> Fix fx
- | FCoFix(cfx,e) when is_subs_id e && is_lift_id lfts -> CoFix cfx
- | _ -> to_constr term_of_fconstr_lift lfts v in
- term_of_fconstr_lift el_id
-
-
-
-(* fstrong applies unfreeze_fun recursively on the (freeze) term and
- * yields a term. Assumes that the unfreeze_fun never returns a
- * FCLOS term.
-let rec fstrong unfreeze_fun lfts v =
- to_constr (fstrong unfreeze_fun) lfts (unfreeze_fun v)
-*)
-
-let rec zip m stk =
- match stk with
- | [] -> m
- | Zapp args :: s -> zip {norm=neutr m.norm; term=FApp(m, args)} s
- | ZcaseT(ci,p,br,e)::s ->
- let t = FCaseT(ci, p, m, br, e) in
- zip {norm=neutr m.norm; term=t} s
- | Zproj p :: s ->
- zip {norm=neutr m.norm; term=FProj (Projection.make p true,m)} s
- | Zfix(fx,par)::s ->
- zip fx (par @ append_stack [|m|] s)
- | Zshift(n)::s ->
- zip (lift_fconstr n m) s
- | Zupdate(rf)::s ->
- zip (update rf (m.norm,m.term)) s
-
-let fapp_stack (m,stk) = zip m stk
-
-(*********************************************************************)
-
-(* The assertions in the functions below are granted because they are
- called only when m is a constructor, a cofix
- (strip_update_shift_app), a fix (get_nth_arg) or an abstraction
- (strip_update_shift, through get_arg). *)
-
-(* optimised for the case where there are no shifts... *)
-let strip_update_shift_app head stk =
- assert (head.norm <> Red);
- let rec strip_rec rstk h depth = function
- | Zshift(k) as e :: s ->
- strip_rec (e::rstk) (lift_fconstr k h) (depth+k) s
- | (Zapp args :: s) ->
- strip_rec (Zapp args :: rstk)
- {norm=h.norm;term=FApp(h,args)} depth s
- | Zupdate(m)::s ->
- strip_rec rstk (update m (h.norm,h.term)) depth s
- | stk -> (depth,List.rev rstk, stk) in
- strip_rec [] head 0 stk
-
-
-let get_nth_arg head n stk =
- assert (head.norm <> Red);
- let rec strip_rec rstk h n = function
- | Zshift(k) as e :: s ->
- strip_rec (e::rstk) (lift_fconstr k h) n s
- | Zapp args::s' ->
- let q = Array.length args in
- if n >= q
- then
- strip_rec (Zapp args::rstk)
- {norm=h.norm;term=FApp(h,args)} (n-q) s'
- else
- let bef = Array.sub args 0 n in
- let aft = Array.sub args (n+1) (q-n-1) in
- let stk' =
- List.rev (if n = 0 then rstk else (Zapp bef :: rstk)) in
- (Some (stk', args.(n)), append_stack aft s')
- | Zupdate(m)::s ->
- strip_rec rstk (update m (h.norm,h.term)) n s
- | s -> (None, List.rev rstk @ s) in
- strip_rec [] head n stk
-
-(* Beta reduction: look for an applied argument in the stack.
- Since the encountered update marks are removed, h must be a whnf *)
-let rec get_args n tys f e stk =
- match stk with
- Zupdate r :: s ->
- let _hd = update r (Cstr,FLambda(n,tys,f,e)) in
- get_args n tys f e s
- | Zshift k :: s ->
- get_args n tys f (subs_shft (k,e)) s
- | Zapp l :: s ->
- let na = Array.length l in
- if n == na then (Inl (subs_cons(l,e)),s)
- else if n < na then (* more arguments *)
- let args = Array.sub l 0 n in
- let eargs = Array.sub l n (na-n) in
- (Inl (subs_cons(args,e)), Zapp eargs :: s)
- else (* more lambdas *)
- let etys = List.skipn na tys in
- get_args (n-na) etys f (subs_cons(l,e)) s
- | _ -> (Inr {norm=Cstr;term=FLambda(n,tys,f,e)}, stk)
-
-(* Eta expansion: add a reference to implicit surrounding lambda at end of stack *)
-let rec eta_expand_stack = function
- | (Zapp _ | Zfix _ | ZcaseT _ | Zproj _
- | Zshift _ | Zupdate _ as e) :: s ->
- e :: eta_expand_stack s
- | [] ->
- [Zshift 1; Zapp [|{norm=Norm; term= FRel 1}|]]
-
-(* Iota reduction: extract the arguments to be passed to the Case
- branches *)
-let rec reloc_rargs_rec depth stk =
- match stk with
- Zapp args :: s ->
- Zapp (lift_fconstr_vect depth args) :: reloc_rargs_rec depth s
- | Zshift(k)::s -> if k=depth then s else reloc_rargs_rec (depth-k) s
- | _ -> stk
-
-let reloc_rargs depth stk =
- if depth = 0 then stk else reloc_rargs_rec depth stk
-
-let rec try_drop_parameters depth n argstk =
- match argstk with
- Zapp args::s ->
- let q = Array.length args in
- if n > q then try_drop_parameters depth (n-q) s
- else if Int.equal n q then reloc_rargs depth s
- else
- let aft = Array.sub args n (q-n) in
- reloc_rargs depth (append_stack aft s)
- | Zshift(k)::s -> try_drop_parameters (depth-k) n s
- | [] ->
- if Int.equal n 0 then []
- else raise Not_found
- | _ -> assert false
- (* strip_update_shift_app only produces Zapp and Zshift items *)
-
-let drop_parameters depth n argstk =
- try try_drop_parameters depth n argstk
- with Not_found -> assert false
- (* we know that n < stack_args_size(argstk) (if well-typed term) *)
-
-(** Projections and eta expansion *)
-
-let eta_expand_ind_stack env ind m s (f, s') =
- let mib = lookup_mind (fst ind) env in
- (* disallow eta-exp for non-primitive records *)
- if not (mib.mind_finite == BiFinite) then raise Not_found;
- match Declarations.inductive_make_projections ind mib with
- | Some projs ->
- (* (Construct, pars1 .. parsm :: arg1...argn :: []) ~= (f, s') ->
- arg1..argn ~= (proj1 t...projn t) where t = zip (f,s') *)
- let pars = mib.mind_nparams in
- let right = fapp_stack (f, s') in
- let (depth, args, s) = strip_update_shift_app m s in
- (** Try to drop the params, might fail on partially applied constructors. *)
- let argss = try_drop_parameters depth pars args in
- let hstack =
- Array.map (fun p ->
- { norm = Red; (* right can't be a constructor though *)
- term = FProj (Projection.make p false, right) })
- projs
- in
- argss, [Zapp hstack]
- | None -> raise Not_found (* disallow eta-exp for non-primitive records *)
-
-let rec project_nth_arg n argstk =
- match argstk with
- | Zapp args :: s ->
- let q = Array.length args in
- if n >= q then project_nth_arg (n - q) s
- else (* n < q *) args.(n)
- | _ -> assert false
- (* After drop_parameters we have a purely applicative stack *)
-
-
-(* Iota reduction: expansion of a fixpoint.
- * Given a fixpoint and a substitution, returns the corresponding
- * fixpoint body, and the substitution in which it should be
- * evaluated: its first variables are the fixpoint bodies
- *
- * FCLOS(fix Fi {F0 := T0 .. Fn-1 := Tn-1}, S)
- * -> (S. FCLOS(F0,S) . ... . FCLOS(Fn-1,S), Ti)
- *)
-(* does not deal with FLIFT *)
-let contract_fix_vect fix =
- let (thisbody, make_body, env, nfix) =
- match fix with
- | FFix (((reci,i),(_,_,bds as rdcl)),env) ->
- (bds.(i),
- (fun j -> { norm = Cstr; term = FFix (((reci,j),rdcl),env) }),
- env, Array.length bds)
- | FCoFix ((i,(_,_,bds as rdcl)),env) ->
- (bds.(i),
- (fun j -> { norm = Cstr; term = FCoFix ((j,rdcl),env) }),
- env, Array.length bds)
- | _ -> assert false
- in
- (subs_cons(Array.init nfix make_body, env), thisbody)
-
-let unfold_projection env p =
- Zproj (Projection.repr p)
-
-(*********************************************************************)
-(* A machine that inspects the head of a term until it finds an
- atom or a subterm that may produce a redex (abstraction,
- constructor, cofix, letin, constant), or a neutral term (product,
- inductive) *)
-let rec knh info m stk =
- match m.term with
- | FLIFT(k,a) -> knh info a (zshift k stk)
- | FCLOS(t,e) -> knht info e t (zupdate m stk)
- | FLOCKED -> assert false
- | FApp(a,b) -> knh info a (append_stack b (zupdate m stk))
- | FCaseT(ci,p,t,br,env) -> knh info t (ZcaseT(ci,p,br,env)::zupdate m stk)
- | FFix(((ri,n),(_,_,_)),_) ->
- (match get_nth_arg m ri.(n) stk with
- (Some(pars,arg),stk') -> knh info arg (Zfix(m,pars)::stk')
- | (None, stk') -> (m,stk'))
- | FCast(t,_,_) -> knh info t stk
-
- | FProj (p,c) ->
- if red_set info.i_flags fDELTA then
- let s = unfold_projection info.i_env p in
- knh info c (s :: zupdate m stk)
- else (m,stk)
-
-(* cases where knh stops *)
- | (FFlex _|FLetIn _|FConstruct _|FEvar _|
- FCoFix _|FLambda _|FRel _|FAtom _|FInd _|FProd _) ->
- (m, stk)
-
-(* The same for pure terms *)
-and knht info e t stk =
- match t with
- | App(a,b) ->
- knht info e a (append_stack (mk_clos_vect e b) stk)
- | Case(ci,p,t,br) -> knht info e t (ZcaseT(ci, p, br, e)::stk)
- | Fix _ -> knh info (mk_clos2 e t) stk (* laziness *)
- | Cast(a,_,_) -> knht info e a stk
- | Rel n -> knh info (clos_rel e n) stk
- | Proj (p,c) -> knh info (mk_clos2 e t) stk (* laziness *)
- | (Lambda _|Prod _|Construct _|CoFix _|Ind _|
- LetIn _|Const _|Var _|Evar _|Meta _|Sort _) ->
- (mk_clos2 e t, stk)
-
-
-(************************************************************************)
-
-(* Computes a weak head normal form from the result of knh. *)
-let rec knr info m stk =
- match m.term with
- | FLambda(n,tys,f,e) when red_set info.i_flags fBETA ->
- (match get_args n tys f e stk with
- Inl e', s -> knit info e' f s
- | Inr lam, s -> (lam,s))
- | FFlex(ConstKey kn) when red_set info.i_flags fDELTA ->
- (match ref_value_cache info (ConstKey kn) with
- Some v -> kni info v stk
- | None -> (set_norm m; (m,stk)))
- | FFlex(VarKey id) when red_set info.i_flags fDELTA ->
- (match ref_value_cache info (VarKey id) with
- Some v -> kni info v stk
- | None -> (set_norm m; (m,stk)))
- | FFlex(RelKey k) when red_set info.i_flags fDELTA ->
- (match ref_value_cache info (RelKey k) with
- Some v -> kni info v stk
- | None -> (set_norm m; (m,stk)))
- | FConstruct((ind,c),u) when red_set info.i_flags fIOTA ->
- (match strip_update_shift_app m stk with
- | (depth, args, ZcaseT(ci,_,br,env)::s) ->
- assert (ci.ci_npar>=0);
- let rargs = drop_parameters depth ci.ci_npar args in
- knit info env br.(c-1) (rargs@s)
- | (_, cargs, Zfix(fx,par)::s) ->
- let rarg = fapp_stack(m,cargs) in
- let stk' = par @ append_stack [|rarg|] s in
- let (fxe,fxbd) = contract_fix_vect fx.term in
- knit info fxe fxbd stk'
- | (depth, args, Zproj p::s) ->
- let rargs = drop_parameters depth (Projection.Repr.npars p) args in
- let rarg = project_nth_arg (Projection.Repr.arg p) rargs in
- kni info rarg s
- | (_,args,s) -> (m,args@s))
- | FCoFix _ when red_set info.i_flags fIOTA ->
- (match strip_update_shift_app m stk with
- (_, args, (((ZcaseT _|Zproj _)::_) as stk')) ->
- let (fxe,fxbd) = contract_fix_vect m.term in
- knit info fxe fxbd (args@stk')
- | (_,args,s) -> (m,args@s))
- | FLetIn (_,v,_,bd,e) when red_set info.i_flags fZETA ->
- knit info (subs_cons([|v|],e)) bd stk
- | _ -> (m,stk)
-
-(* Computes the weak head normal form of a term *)
-and kni info m stk =
- let (hm,s) = knh info m stk in
- knr info hm s
-and knit info e t stk =
- let (ht,s) = knht info e t stk in
- knr info ht s
-
-let kh info v stk = fapp_stack(kni info v stk)
-
-(************************************************************************)
-(* Initialization and then normalization *)
-
-(* weak reduction *)
-let whd_val info v =
- with_stats (lazy (term_of_fconstr (kh info v [])))
-
-let whd_stack infos m stk =
- let k = kni infos m stk in
- let _ = fapp_stack k in (* to unlock Zupdates! *)
- k
-
-let infos_env x = x.i_env
-let infos_flags x = x.i_flags
-let oracle_of_infos x = x.i_env.env_conv_oracle
-
-let create_clos_infos flgs env =
- { i_flags = flgs;
- i_env = env;
- i_rels = defined_rels flgs env;
- i_tab = KeyTable.create 17 }
-
-let unfold_reference = ref_value_cache
diff --git a/checker/closure.mli b/checker/closure.mli
deleted file mode 100644
index 4c6643754b..0000000000
--- a/checker/closure.mli
+++ /dev/null
@@ -1,174 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(*i*)
-open Names
-open Cic
-open Esubst
-open Environ
-(*i*)
-
-(* Flags for profiling reductions. *)
-val stats : bool ref
-val share : bool ref
-
-val with_stats: 'a Lazy.t -> 'a
-
-(*s Delta implies all consts (both global (= by
- [kernel_name]) and local (= by [Rel] or [Var])), all evars, and letin's.
- Rem: reduction of a Rel/Var bound to a term is Delta, but reduction of
- a LetIn expression is Letin reduction *)
-
-(* Sets of reduction kinds. *)
-module type RedFlagsSig = sig
- type reds
- type red_kind
-
- (* The different kinds of reduction *)
- val fBETA : red_kind
- val fDELTA : red_kind
- val fIOTA : red_kind
- val fZETA : red_kind
-
- (* No reduction at all *)
- val no_red : reds
-
- (* Adds a reduction kind to a set *)
- val red_add : reds -> red_kind -> reds
-
- (* Build a reduction set from scratch = iter [red_add] on [no_red] *)
- val mkflags : red_kind list -> reds
-
- (* Tests if a reduction kind is set *)
- val red_set : reds -> red_kind -> bool
-end
-
-module RedFlags : RedFlagsSig
-open RedFlags
-
-val betadeltaiota : reds
-val betaiotazeta : reds
-val betadeltaiotanolet : reds
-
-(***********************************************************************)
-
-type table_key = Constant.t puniverses tableKey
-
-(************************************************************************)
-(*s Lazy reduction. *)
-
-(* [fconstr] is the type of frozen constr *)
-
-type fconstr
-
-(* [fconstr] can be accessed by using the function [fterm_of] and by
- matching on type [fterm] *)
-
-type fterm =
- | FRel of int
- | FAtom of constr (* Metas and Sorts *)
- | FCast of fconstr * cast_kind * fconstr
- | FFlex of table_key
- | FInd of pinductive
- | FConstruct of pconstructor
- | FApp of fconstr * fconstr array
- | FProj of Projection.t * fconstr
- | FFix of fixpoint * fconstr subs
- | FCoFix of cofixpoint * fconstr subs
- | FCaseT of case_info * constr * fconstr * constr array * fconstr subs (* predicate and branches are closures *)
- | FLambda of int * (Name.t * constr) list * constr * fconstr subs
- | FProd of Name.t * fconstr * fconstr
- | FLetIn of Name.t * fconstr * fconstr * constr * fconstr subs
- | FEvar of existential_key * fconstr array
- | FLIFT of int * fconstr
- | FCLOS of constr * fconstr subs
- | FLOCKED
-
-(************************************************************************)
-(*s A [stack] is a context of arguments, arguments are pushed by
- [append_stack] one array at a time but popped with [decomp_stack]
- one by one *)
-
-type stack_member =
- | Zapp of fconstr array
- | ZcaseT of case_info * constr * constr array * fconstr subs
- | Zproj of Projection.Repr.t
- | Zfix of fconstr * stack
- | Zshift of int
- | Zupdate of fconstr
-
-and stack = stack_member list
-
-val append_stack : fconstr array -> stack -> stack
-
-val stack_args_size : stack -> int
-
-val eta_expand_stack : stack -> stack
-
-val eta_expand_ind_stack : env -> inductive -> fconstr -> stack ->
- (fconstr * stack) -> stack * stack
-
-val unfold_projection : env -> Projection.t -> stack_member
-
-(* To lazy reduce a constr, create a [clos_infos] with
- [create_clos_infos], inject the term to reduce with [inject]; then use
- a reduction function *)
-
-val inject : constr -> fconstr
-
-val fterm_of : fconstr -> fterm
-val term_of_fconstr : fconstr -> constr
-val destFLambda :
- (fconstr subs -> constr -> fconstr) -> fconstr -> Name.t * fconstr * fconstr
-
-(* Global and local constant cache *)
-type clos_infos
-val create_clos_infos : reds -> env -> clos_infos
-val infos_env : clos_infos -> env
-val infos_flags : clos_infos -> reds
-val oracle_of_infos : clos_infos -> oracle
-
-
-(* Reduction function *)
-
-(* [whd_val] is for weak head normalization *)
-val whd_val : clos_infos -> fconstr -> constr
-
-(* [whd_stack] performs weak head normalization in a given stack. It
- stops whenever a reduction is blocked. *)
-val whd_stack :
- clos_infos -> fconstr -> stack -> fconstr * stack
-
-(* Conversion auxiliary functions to do step by step normalisation *)
-
-(* [unfold_reference] unfolds references in a [fconstr] *)
-val unfold_reference : clos_infos -> table_key -> fconstr option
-
-(* [mind_equiv] checks whether two inductive types are intentionally equal *)
-val mind_equiv_infos : clos_infos -> inductive -> inductive -> bool
-
-(************************************************************************)
-(*i This is for lazy debug *)
-
-val lift_fconstr : int -> fconstr -> fconstr
-val lift_fconstr_vect : int -> fconstr array -> fconstr array
-
-val mk_clos : fconstr subs -> constr -> fconstr
-val mk_clos_vect : fconstr subs -> constr array -> fconstr array
-val mk_clos_deep :
- (fconstr subs -> constr -> fconstr) ->
- fconstr subs -> constr -> fconstr
-
-val kni: clos_infos -> fconstr -> stack -> fconstr * stack
-val knr: clos_infos -> fconstr -> stack -> fconstr * stack
-
-val to_constr : (lift -> fconstr -> constr) -> lift -> fconstr -> constr
-
-(* End of cbn debug section i*)
diff --git a/checker/declarations.ml b/checker/declarations.ml
deleted file mode 100644
index 93d5f8bfa2..0000000000
--- a/checker/declarations.ml
+++ /dev/null
@@ -1,606 +0,0 @@
-open Util
-open Names
-open Cic
-open Term
-
-(** Substitutions, code imported from kernel/mod_subst *)
-
-module Deltamap = struct
- [@@@ocaml.warning "-32-34"]
- type t = delta_resolver
- let empty = MPmap.empty, KNmap.empty
- let is_empty (mm, km) = MPmap.is_empty mm && KNmap.is_empty km
- let add_kn kn hint (mm,km) = (mm,KNmap.add kn hint km)
- let add_mp mp mp' (mm,km) = (MPmap.add mp mp' mm, km)
- let remove_mp mp (mm,km) = (MPmap.remove mp mm, km)
- let find_mp mp map = MPmap.find mp (fst map)
- let find_kn kn map = KNmap.find kn (snd map)
- let mem_mp mp map = MPmap.mem mp (fst map)
- let mem_kn kn map = KNmap.mem kn (snd map)
- let fold_kn f map i = KNmap.fold f (snd map) i
- let fold fmp fkn (mm,km) i =
- MPmap.fold fmp mm (KNmap.fold fkn km i)
- let join map1 map2 = fold add_mp add_kn map1 map2
-end
-
-let empty_delta_resolver = Deltamap.empty
-
-module Umap = struct
- [@@@ocaml.warning "-32-34"]
- type 'a t = 'a umap_t
- let empty = MPmap.empty
- let is_empty m = MPmap.is_empty m
- let add_mbi mbi x m = MPmap.add (MPbound mbi) x m
- let add_mp mp x m = MPmap.add mp x m
- let find = MPmap.find
- let fold = MPmap.fold
- let join map1 map2 = fold add_mp map1 map2
-end
-
-type 'a subst_fun = substitution -> 'a -> 'a
-
-let empty_subst = Umap.empty
-
-let is_empty_subst = Umap.is_empty
-
-let add_mbid mbid mp = Umap.add_mbi mbid (mp,empty_delta_resolver)
-let add_mp mp1 mp2 = Umap.add_mp mp1 (mp2,empty_delta_resolver)
-
-let map_mbid mbid mp = add_mbid mbid mp empty_subst
-let map_mp mp1 mp2 = add_mp mp1 mp2 empty_subst
-
-let mp_in_delta mp =
- Deltamap.mem_mp mp
-
-let find_prefix resolve mp =
- let rec sub_mp = function
- | MPdot(mp,l) as mp_sup ->
- (try Deltamap.find_mp mp_sup resolve
- with Not_found -> MPdot(sub_mp mp,l))
- | p -> Deltamap.find_mp p resolve
- in
- try sub_mp mp with Not_found -> mp
-
-(** Nota: the following function is slightly different in mod_subst
- PL: Is it on purpose ? *)
-
-let solve_delta_kn resolve kn =
- try
- match Deltamap.find_kn kn resolve with
- | Equiv kn1 -> kn1
- | Inline _ -> raise Not_found
- with Not_found ->
- let mp,l = KerName.repr kn in
- let new_mp = find_prefix resolve mp in
- if mp == new_mp then
- kn
- else
- KerName.make new_mp l
-
-let gen_of_delta resolve x kn fix_can =
- let new_kn = solve_delta_kn resolve kn in
- if kn == new_kn then x else fix_can new_kn
-
-let constant_of_delta resolve con =
- let kn = Constant.user con in
- gen_of_delta resolve con kn (Constant.make kn)
-
-let constant_of_delta2 resolve con =
- let kn, kn' = Constant.canonical con, Constant.user con in
- gen_of_delta resolve con kn (Constant.make kn')
-
-let mind_of_delta resolve mind =
- let kn = MutInd.user mind in
- gen_of_delta resolve mind kn (MutInd.make kn)
-
-let mind_of_delta2 resolve mind =
- let kn, kn' = MutInd.canonical mind, MutInd.user mind in
- gen_of_delta resolve mind kn (MutInd.make kn')
-
-let find_inline_of_delta kn resolve =
- match Deltamap.find_kn kn resolve with
- | Inline (_,o) -> o
- | _ -> raise Not_found
-
-let constant_of_delta_with_inline resolve con =
- let kn1,kn2 = Constant.canonical con, Constant.user con in
- try find_inline_of_delta kn2 resolve
- with Not_found ->
- try find_inline_of_delta kn1 resolve
- with Not_found -> None
-
-let subst_mp0 sub mp = (* 's like subst *)
- let rec aux mp =
- match mp with
- | MPfile _ | MPbound _ -> Umap.find mp sub
- | MPdot (mp1,l) as mp2 ->
- begin
- try Umap.find mp2 sub
- with Not_found ->
- let mp1',resolve = aux mp1 in
- MPdot (mp1',l),resolve
- end
- in
- try Some (aux mp) with Not_found -> None
-
-let subst_mp sub mp =
- match subst_mp0 sub mp with
- None -> mp
- | Some (mp',_) -> mp'
-
-let subst_kn_delta sub kn =
- let mp,l = KerName.repr kn in
- match subst_mp0 sub mp with
- Some (mp',resolve) ->
- solve_delta_kn resolve (KerName.make mp' l)
- | None -> kn
-
-let subst_kn sub kn =
- let mp,l = KerName.repr kn in
- match subst_mp0 sub mp with
- Some (mp',_) ->
- KerName.make mp' l
- | None -> kn
-
-exception No_subst
-
-type sideconstantsubst =
- | User
- | Canonical
-
-
-let gen_subst_mp f sub mp1 mp2 =
- match subst_mp0 sub mp1, subst_mp0 sub mp2 with
- | None, None -> raise No_subst
- | Some (mp',resolve), None -> User, (f mp' mp2), resolve
- | None, Some (mp',resolve) -> Canonical, (f mp1 mp'), resolve
- | Some (mp1',_), Some (mp2',resolve2) -> Canonical, (f mp1' mp2'), resolve2
-
-let make_mind_equiv mpu mpc l =
- let knu = KerName.make mpu l in
- if mpu == mpc then MutInd.make1 knu
- else MutInd.make knu (KerName.make mpc l)
-
-let subst_ind sub mind =
- let kn1,kn2 = MutInd.user mind, MutInd.canonical mind in
- let mp1,l = KerName.repr kn1 in
- let mp2,_ = KerName.repr kn2 in
- let rebuild_mind mp1 mp2 = make_mind_equiv mp1 mp2 l in
- try
- let side,mind',resolve = gen_subst_mp rebuild_mind sub mp1 mp2 in
- match side with
- | User -> mind_of_delta resolve mind'
- | Canonical -> mind_of_delta2 resolve mind'
- with No_subst -> mind
-
-let make_con_equiv mpu mpc l =
- let knu = KerName.make mpu l in
- if mpu == mpc then Constant.make1 knu
- else Constant.make knu (KerName.make mpc l)
-
-let subst_con0 sub con u =
- let kn1,kn2 = Constant.user con, Constant.canonical con in
- let mp1,l = KerName.repr kn1 in
- let mp2,_ = KerName.repr kn2 in
- let rebuild_con mp1 mp2 = make_con_equiv mp1 mp2 l in
- let dup con = con, Const (con, u) in
- let side,con',resolve = gen_subst_mp rebuild_con sub mp1 mp2 in
- match constant_of_delta_with_inline resolve con' with
- | Some (ctx, t) ->
- (** FIXME: we never typecheck the inlined term, so that it could well
- be garbage. What environment do we type it in though? The substitution
- code should be moot in the checker but it **is** used nonetheless. *)
- let () = assert (Univ.AUContext.size ctx == Univ.Instance.length u) in
- con', subst_instance_constr u t
- | None ->
- let con'' = match side with
- | User -> constant_of_delta resolve con'
- | Canonical -> constant_of_delta2 resolve con'
- in
- if con'' == con then raise No_subst else dup con''
-
-let rec map_kn f f' c =
- let func = map_kn f f' in
- match c with
- | Const (kn, u) -> (try snd (f' kn u) with No_subst -> c)
- | Proj (p,t) ->
- let p' = Projection.map f p in
- let t' = func t in
- if p' == p && t' == t then c
- else Proj (p', t')
- | Ind ((kn,i),u) ->
- let kn' = f kn in
- if kn'==kn then c else Ind ((kn',i),u)
- | Construct (((kn,i),j),u) ->
- let kn' = f kn in
- if kn'==kn then c else Construct (((kn',i),j),u)
- | Case (ci,p,ct,l) ->
- let ci_ind =
- let (kn,i) = ci.ci_ind in
- let kn' = f kn in
- if kn'==kn then ci.ci_ind else kn',i
- in
- let p' = func p in
- let ct' = func ct in
- let l' = Array.Smart.map func l in
- if (ci.ci_ind==ci_ind && p'==p
- && l'==l && ct'==ct)then c
- else
- Case ({ci with ci_ind = ci_ind},
- p',ct', l')
- | Cast (ct,k,t) ->
- let ct' = func ct in
- let t'= func t in
- if (t'==t && ct'==ct) then c
- else Cast (ct', k, t')
- | Prod (na,t,ct) ->
- let ct' = func ct in
- let t'= func t in
- if (t'==t && ct'==ct) then c
- else Prod (na, t', ct')
- | Lambda (na,t,ct) ->
- let ct' = func ct in
- let t'= func t in
- if (t'==t && ct'==ct) then c
- else Lambda (na, t', ct')
- | LetIn (na,b,t,ct) ->
- let ct' = func ct in
- let t'= func t in
- let b'= func b in
- if (t'==t && ct'==ct && b==b') then c
- else LetIn (na, b', t', ct')
- | App (ct,l) ->
- let ct' = func ct in
- let l' = Array.Smart.map func l in
- if (ct'== ct && l'==l) then c
- else App (ct',l')
- | Evar (e,l) ->
- let l' = Array.Smart.map func l in
- if (l'==l) then c
- else Evar (e,l')
- | Fix (ln,(lna,tl,bl)) ->
- let tl' = Array.Smart.map func tl in
- let bl' = Array.Smart.map func bl in
- if (bl == bl'&& tl == tl') then c
- else Fix (ln,(lna,tl',bl'))
- | CoFix(ln,(lna,tl,bl)) ->
- let tl' = Array.Smart.map func tl in
- let bl' = Array.Smart.map func bl in
- if (bl == bl'&& tl == tl') then c
- else CoFix (ln,(lna,tl',bl'))
- | _ -> c
-
-let subst_mps sub c =
- if is_empty_subst sub then c
- else map_kn (subst_ind sub) (subst_con0 sub) c
-
-let rec replace_mp_in_mp mpfrom mpto mp =
- match mp with
- | _ when ModPath.equal mp mpfrom -> mpto
- | MPdot (mp1,l) ->
- let mp1' = replace_mp_in_mp mpfrom mpto mp1 in
- if mp1==mp1' then mp
- else MPdot (mp1',l)
- | _ -> mp
-
-let rec mp_in_mp mp mp1 =
- match mp1 with
- | _ when ModPath.equal mp1 mp -> true
- | MPdot (mp2,l) -> mp_in_mp mp mp2
- | _ -> false
-
-let subset_prefixed_by mp resolver =
- let mp_prefix mkey mequ rslv =
- if mp_in_mp mp mkey then Deltamap.add_mp mkey mequ rslv else rslv
- in
- let kn_prefix kn hint rslv =
- match hint with
- | Inline _ -> rslv
- | Equiv _ ->
- if mp_in_mp mp (KerName.modpath kn)
- then Deltamap.add_kn kn hint rslv
- else rslv
- in
- Deltamap.fold mp_prefix kn_prefix resolver empty_delta_resolver
-
-let subst_dom_delta_resolver subst resolver =
- let mp_apply_subst mkey mequ rslv =
- Deltamap.add_mp (subst_mp subst mkey) mequ rslv
- in
- let kn_apply_subst kkey hint rslv =
- Deltamap.add_kn (subst_kn subst kkey) hint rslv
- in
- Deltamap.fold mp_apply_subst kn_apply_subst resolver empty_delta_resolver
-
-let subst_mp_delta sub mp mkey =
- match subst_mp0 sub mp with
- None -> empty_delta_resolver,mp
- | Some (mp',resolve) ->
- let mp1 = find_prefix resolve mp' in
- let resolve1 = subset_prefixed_by mp1 resolve in
- (subst_dom_delta_resolver
- (map_mp mp1 mkey) resolve1),mp1
-
-let gen_subst_delta_resolver dom subst resolver =
- let mp_apply_subst mkey mequ rslv =
- let mkey' = if dom then subst_mp subst mkey else mkey in
- let rslv',mequ' = subst_mp_delta subst mequ mkey in
- Deltamap.join rslv' (Deltamap.add_mp mkey' mequ' rslv)
- in
- let kn_apply_subst kkey hint rslv =
- let kkey' = if dom then subst_kn subst kkey else kkey in
- let hint' = match hint with
- | Equiv kequ -> Equiv (subst_kn_delta subst kequ)
- | Inline (lev,Some (ctx, t)) -> Inline (lev,Some (ctx, subst_mps subst t))
- | Inline (_,None) -> hint
- in
- Deltamap.add_kn kkey' hint' rslv
- in
- Deltamap.fold mp_apply_subst kn_apply_subst resolver empty_delta_resolver
-
-let subst_codom_delta_resolver = gen_subst_delta_resolver false
-let subst_dom_codom_delta_resolver = gen_subst_delta_resolver true
-
-let update_delta_resolver resolver1 resolver2 =
- let mp_apply_rslv mkey mequ rslv =
- if Deltamap.mem_mp mkey resolver2 then rslv
- else Deltamap.add_mp mkey (find_prefix resolver2 mequ) rslv
- in
- let kn_apply_rslv kkey hint rslv =
- if Deltamap.mem_kn kkey resolver2 then rslv
- else
- let hint' = match hint with
- | Equiv kequ -> Equiv (solve_delta_kn resolver2 kequ)
- | _ -> hint
- in
- Deltamap.add_kn kkey hint' rslv
- in
- Deltamap.fold mp_apply_rslv kn_apply_rslv resolver1 empty_delta_resolver
-
-let add_delta_resolver resolver1 resolver2 =
- if resolver1 == resolver2 then
- resolver2
- else if Deltamap.is_empty resolver2 then
- resolver1
- else
- Deltamap.join (update_delta_resolver resolver1 resolver2) resolver2
-
-let substition_prefixed_by k mp subst =
- let mp_prefixmp kmp (mp_to,reso) sub =
- if mp_in_mp mp kmp && not (ModPath.equal mp kmp) then
- let new_key = replace_mp_in_mp mp k kmp in
- Umap.add_mp new_key (mp_to,reso) sub
- else sub
- in
- Umap.fold mp_prefixmp subst empty_subst
-
-let join subst1 subst2 =
- let apply_subst mpk add (mp,resolve) res =
- let mp',resolve' =
- match subst_mp0 subst2 mp with
- | None -> mp, None
- | Some (mp',resolve') -> mp', Some resolve' in
- let resolve'' =
- match resolve' with
- | Some res ->
- add_delta_resolver
- (subst_dom_codom_delta_resolver subst2 resolve) res
- | None ->
- subst_codom_delta_resolver subst2 resolve
- in
- let prefixed_subst = substition_prefixed_by mpk mp' subst2 in
- Umap.join prefixed_subst (add (mp',resolve'') res)
- in
- let mp_apply_subst mp = apply_subst mp (Umap.add_mp mp) in
- let subst = Umap.fold mp_apply_subst subst1 empty_subst in
- Umap.join subst2 subst
-
-let from_val x = { subst_value = x; subst_subst = []; }
-
-let force fsubst r = match r.subst_subst with
-| [] -> r.subst_value
-| s ->
- let subst = List.fold_left join empty_subst (List.rev s) in
- let x = fsubst subst r.subst_value in
- let () = r.subst_subst <- [] in
- let () = r.subst_value <- x in
- x
-
-let subst_substituted s r = { r with subst_subst = s :: r.subst_subst; }
-
-let force_constr = force subst_mps
-
-let subst_constr_subst = subst_substituted
-
-let subst_lazy_constr sub = function
- | Indirect (l,dp,i) -> Indirect (sub::l,dp,i)
-
-let indirect_opaque_access =
- ref ((fun dp i -> assert false) : DirPath.t -> int -> constr)
-let indirect_opaque_univ_access =
- ref ((fun dp i -> assert false) : DirPath.t -> int -> Univ.ContextSet.t)
-
-let force_lazy_constr = function
- | Indirect (l,dp,i) ->
- let c = !indirect_opaque_access dp i in
- force_constr (List.fold_right subst_constr_subst l (from_val c))
-
-let force_lazy_constr_univs = function
- | OpaqueDef (Indirect (l,dp,i)) -> !indirect_opaque_univ_access dp i
- | _ -> Univ.ContextSet.empty
-
-let subst_constant_def sub = function
- | Undef inl -> Undef inl
- | Def c -> Def (subst_constr_subst sub c)
- | OpaqueDef lc -> OpaqueDef (subst_lazy_constr sub lc)
-
-(** Local variables and graph *)
-
-let body_of_constant cb = match cb.const_body with
- | Undef _ -> None
- | Def c -> Some (force_constr c)
- | OpaqueDef c -> Some (force_lazy_constr c)
-
-let constant_has_body cb = match cb.const_body with
- | Undef _ -> false
- | Def _ | OpaqueDef _ -> true
-
-let is_opaque cb = match cb.const_body with
- | OpaqueDef _ -> true
- | Def _ | Undef _ -> false
-
-let opaque_univ_context cb = force_lazy_constr_univs cb.const_body
-
-let subst_recarg sub r = match r with
- | Norec -> r
- | (Mrec(kn,i)|Imbr (kn,i)) -> let kn' = subst_ind sub kn in
- if kn==kn' then r else Imbr (kn',i)
-
-let mk_norec = Rtree.mk_node Norec [||]
-
-let mk_paths r recargs =
- Rtree.mk_node r
- (Array.map (fun l -> Rtree.mk_node Norec (Array.of_list l)) recargs)
-
-let dest_recarg p = fst (Rtree.dest_node p)
-
-let dest_subterms p =
- let (_,cstrs) = Rtree.dest_node p in
- Array.map (fun t -> Array.to_list (snd (Rtree.dest_node t))) cstrs
-
-let subst_wf_paths sub p = Rtree.Smart.map (subst_recarg sub) p
-
-let eq_recarg r1 r2 = match r1, r2 with
- | Norec, Norec -> true
- | Mrec i1, Mrec i2 -> Names.eq_ind_chk i1 i2
- | Imbr i1, Imbr i2 -> Names.eq_ind_chk i1 i2
- | _ -> false
-
-let eq_wf_paths = Rtree.equal eq_recarg
-
-let inductive_make_projections ind mib =
- match mib.mind_record with
- | NotRecord | FakeRecord -> None
- | PrimRecord infos ->
- let projs = Array.mapi (fun proj_arg lab ->
- Names.Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg lab)
- (pi2 infos.(snd ind))
- in
- Some projs
-
-(**********************************************************************)
-(* Representation of mutual inductive types in the kernel *)
-(*
- Inductive I1 (params) : U1 := c11 : T11 | ... | c1p1 : T1p1
- ...
- with In (params) : Un := cn1 : Tn1 | ... | cnpn : Tnpn
-*)
-
-
-let subst_decl_arity f g sub ar =
- match ar with
- | RegularArity x ->
- let x' = f sub x in
- if x' == x then ar
- else RegularArity x'
- | TemplateArity x ->
- let x' = g sub x in
- if x' == x then ar
- else TemplateArity x'
-
-let subst_rel_declaration sub =
- Term.map_rel_decl (subst_mps sub)
-
-let subst_rel_context sub = List.Smart.map (subst_rel_declaration sub)
-
-let constant_is_polymorphic cb =
- match cb.const_universes with
- | Monomorphic_const _ -> false
- | Polymorphic_const _ -> true
-
-(* TODO: should be changed to non-coping after Term.subst_mps *)
-(* NB: we leave bytecode and native code fields untouched *)
-let subst_const_body sub cb =
- { cb with
- const_body = subst_constant_def sub cb.const_body;
- const_type = subst_mps sub cb.const_type }
-
-
-let subst_regular_ind_arity sub s =
- let uar' = subst_mps sub s.mind_user_arity in
- if uar' == s.mind_user_arity then s
- else { mind_user_arity = uar'; mind_sort = s.mind_sort }
-
-let subst_template_ind_arity sub s = s
-
-(* FIXME records *)
-let subst_ind_arity =
- subst_decl_arity subst_regular_ind_arity subst_template_ind_arity
-
-let subst_mind_packet sub mbp =
- { mind_consnames = mbp.mind_consnames;
- mind_consnrealdecls = mbp.mind_consnrealdecls;
- mind_consnrealargs = mbp.mind_consnrealargs;
- mind_typename = mbp.mind_typename;
- mind_nf_lc = Array.Smart.map (subst_mps sub) mbp.mind_nf_lc;
- mind_arity_ctxt = subst_rel_context sub mbp.mind_arity_ctxt;
- mind_arity = subst_ind_arity sub mbp.mind_arity;
- mind_user_lc = Array.Smart.map (subst_mps sub) mbp.mind_user_lc;
- mind_nrealargs = mbp.mind_nrealargs;
- mind_nrealdecls = mbp.mind_nrealdecls;
- mind_kelim = mbp.mind_kelim;
- mind_recargs = subst_wf_paths sub mbp.mind_recargs (*wf_paths*);
- mind_nb_constant = mbp.mind_nb_constant;
- mind_nb_args = mbp.mind_nb_args;
- mind_reloc_tbl = mbp.mind_reloc_tbl }
-
-
-let subst_mind sub mib =
- { mib with
- mind_params_ctxt = map_rel_context (subst_mps sub) mib.mind_params_ctxt;
- mind_packets = Array.Smart.map (subst_mind_packet sub) mib.mind_packets }
-
-(* Modules *)
-
-let rec functor_map fty f0 = function
- | NoFunctor a -> NoFunctor (f0 a)
- | MoreFunctor (mbid,ty,e) -> MoreFunctor(mbid,fty ty,functor_map fty f0 e)
-
-let implem_map fs fa = function
- | Struct s -> Struct (fs s)
- | Algebraic a -> Algebraic (fa a)
- | impl -> impl
-
-let rec subst_expr sub = function
- | MEident mp -> MEident (subst_mp sub mp)
- | MEapply (me1,mp2)-> MEapply (subst_expr sub me1, subst_mp sub mp2)
- | MEwith (me,wd)-> MEwith (subst_expr sub me, wd)
-
-let rec subst_expression sub me =
- functor_map (subst_module_type sub) (subst_expr sub) me
-
-and subst_signature sub sign =
- functor_map (subst_module_type sub) (subst_structure sub) sign
-
-and subst_structure sub struc =
- let subst_body = function
- | SFBconst cb -> SFBconst (subst_const_body sub cb)
- | SFBmind mib -> SFBmind (subst_mind sub mib)
- | SFBmodule mb -> SFBmodule (subst_module sub mb)
- | SFBmodtype mtb -> SFBmodtype (subst_module_type sub mtb)
- in
- List.map (fun (l,b) -> (l,subst_body b)) struc
-
-and subst_body : 'a. (_ -> 'a -> 'a) -> _ -> 'a generic_module_body -> 'a generic_module_body =
- fun subst_impl sub mb ->
- { mb with
- mod_mp = subst_mp sub mb.mod_mp;
- mod_expr = subst_impl sub mb.mod_expr;
- mod_type = subst_signature sub mb.mod_type;
- mod_type_alg = Option.Smart.map (subst_expression sub) mb.mod_type_alg }
-
-and subst_module sub mb =
- subst_body (fun sub e -> implem_map (subst_signature sub) (subst_expression sub) e) sub mb
-
-and subst_module_type sub mb =
- subst_body (fun _ () -> ()) sub mb
diff --git a/checker/declarations.mli b/checker/declarations.mli
deleted file mode 100644
index ce852644ef..0000000000
--- a/checker/declarations.mli
+++ /dev/null
@@ -1,52 +0,0 @@
-open Names
-open Cic
-
-val force_constr : constr_substituted -> constr
-val force_lazy_constr_univs : Cic.constant_def -> Univ.ContextSet.t
-val from_val : constr -> constr_substituted
-
-val indirect_opaque_access : (DirPath.t -> int -> constr) ref
-val indirect_opaque_univ_access : (DirPath.t -> int -> Univ.ContextSet.t) ref
-
-(** Constant_body *)
-
-val body_of_constant : constant_body -> constr option
-val constant_has_body : constant_body -> bool
-val is_opaque : constant_body -> bool
-val opaque_univ_context : constant_body -> Univ.ContextSet.t
-val constant_is_polymorphic : constant_body -> bool
-
-(* Mutual inductives *)
-
-val mk_norec : wf_paths
-val mk_paths : recarg -> wf_paths list array -> wf_paths
-val dest_recarg : wf_paths -> recarg
-val dest_subterms : wf_paths -> wf_paths list array
-val eq_recarg : recarg -> recarg -> bool
-val eq_wf_paths : wf_paths -> wf_paths -> bool
-
-val inductive_make_projections : Names.inductive -> mutual_inductive_body ->
- Names.Projection.Repr.t array option
-
-(* Modules *)
-
-val empty_delta_resolver : delta_resolver
-
-(* Substitutions *)
-
-type 'a subst_fun = substitution -> 'a -> 'a
-
-val empty_subst : substitution
-val add_mbid : MBId.t -> ModPath.t -> substitution -> substitution
-val add_mp : ModPath.t -> ModPath.t -> substitution -> substitution
-val map_mbid : MBId.t -> ModPath.t -> substitution
-val map_mp : ModPath.t -> ModPath.t -> substitution
-val mp_in_delta : ModPath.t -> delta_resolver -> bool
-val mind_of_delta : delta_resolver -> MutInd.t -> MutInd.t
-
-val subst_const_body : constant_body subst_fun
-val subst_mind : mutual_inductive_body subst_fun
-val subst_signature : substitution -> module_signature -> module_signature
-val subst_module : substitution -> module_body -> module_body
-
-val join : substitution -> substitution -> substitution
diff --git a/checker/dune b/checker/dune
index 5aff9211d7..35a35a1f82 100644
--- a/checker/dune
+++ b/checker/dune
@@ -1,7 +1,25 @@
-(rule (copy %{project_root}/kernel/names.ml names.ml))
-(rule (copy %{project_root}/kernel/names.mli names.mli))
-(rule (copy %{project_root}/kernel/esubst.ml esubst.ml))
-(rule (copy %{project_root}/kernel/esubst.mli esubst.mli))
+(copy_files#
+ %{project_root}/kernel/{names,esubst,declarations,environ,constr,term,univ,evar,sorts,uGraph,context}.ml{,i})
+
+(copy_files#
+ %{project_root}/kernel/{mod_subst,vars,opaqueproof,conv_oracle,reduction,typeops,inductive,indtypes,declareops,type_errors}.ml{,i})
+
+(copy_files#
+ %{project_root}/kernel/{modops,mod_typing,}.ml{,i})
+
+(copy_files#
+ %{project_root}/kernel/{cClosure,cPrimitives,csymtable,vconv,vm,uint31,cemitcodes,vmvalues,cbytecodes,cinstr,retroknowledge,copcodes}.ml{,i})
+
+(copy_files#
+ %{project_root}/kernel/{cbytegen,clambda,nativeinstr,nativevalues,nativeconv,nativecode,nativelib,nativelibrary,nativelambda}.ml{,i})
+
+(copy_files#
+ %{project_root}/kernel/{subtyping,term_typing,safe_typing,entries,cooking}.ml{,i})
+
+; VM stuff
+
+(copy_files#
+ %{project_root}/kernel/byterun/{*.c,*.h})
; Careful with bug https://github.com/ocaml/odoc/issues/148
;
@@ -12,7 +30,8 @@
(public_name coq.checklib)
(synopsis "Coq's Standalone Proof Checker")
(modules :standard \ coqchk votour)
- (modules_without_implementation cic)
+ (modules_without_implementation cinstr nativeinstr)
+ (c_names coq_fix_code coq_memory coq_values coq_interp)
(wrapped true)
(libraries coq.lib))
diff --git a/checker/environ.ml b/checker/environ.ml
deleted file mode 100644
index b172acb126..0000000000
--- a/checker/environ.ml
+++ /dev/null
@@ -1,250 +0,0 @@
-open CErrors
-open Util
-open Names
-open Cic
-open Term
-open Declarations
-
-type globals = {
- env_constants : constant_body Cmap_env.t;
- env_inductives : mutual_inductive_body Mindmap_env.t;
- env_inductives_eq : KerName.t KNmap.t;
- env_modules : module_body MPmap.t;
- env_modtypes : module_type_body MPmap.t}
-
-type stratification = {
- env_universes : Univ.universes;
- env_engagement : engagement
-}
-
-type env = {
- env_globals : globals;
- env_rel_context : rel_context;
- env_stratification : stratification;
- env_imports : Cic.vodigest MPmap.t;
- env_conv_oracle : oracle; }
-
-let empty_oracle = {
- var_opacity = Id.Map.empty;
- cst_opacity = Cmap.empty;
- var_trstate = Id.Pred.empty;
- cst_trstate = Cpred.empty;
-}
-
-let empty_env = {
- env_globals =
- { env_constants = Cmap_env.empty;
- env_inductives = Mindmap_env.empty;
- env_inductives_eq = KNmap.empty;
- env_modules = MPmap.empty;
- env_modtypes = MPmap.empty};
- env_rel_context = [];
- env_stratification =
- { env_universes = Univ.initial_universes;
- env_engagement = PredicativeSet };
- env_imports = MPmap.empty;
- env_conv_oracle = empty_oracle }
-
-let engagement env = env.env_stratification.env_engagement
-let universes env = env.env_stratification.env_universes
-let rel_context env = env.env_rel_context
-
-let set_engagement (impr_set as c) env =
- let expected_impr_set =
- env.env_stratification.env_engagement in
- begin
- match impr_set,expected_impr_set with
- | PredicativeSet, ImpredicativeSet -> user_err Pp.(str "Incompatible engagement")
- | _ -> ()
- end;
- { env with env_stratification =
- { env.env_stratification with env_engagement = c } }
-
-let set_oracle env oracle = { env with env_conv_oracle = oracle }
-
-(* Digests *)
-
-let add_digest env dp digest =
- { env with env_imports = MPmap.add (MPfile dp) digest env.env_imports }
-
-let lookup_digest env dp =
- MPmap.find (MPfile dp) env.env_imports
-
-(* Rel context *)
-let lookup_rel n env =
- let rec lookup_rel n sign =
- match n, sign with
- | 1, decl :: _ -> decl
- | n, _ :: sign -> lookup_rel (n-1) sign
- | _, [] -> raise Not_found in
- lookup_rel n env.env_rel_context
-
-let push_rel d env =
- { env with
- env_rel_context = d :: env.env_rel_context }
-
-let push_rel_context ctxt x = fold_rel_context push_rel ctxt ~init:x
-
-let push_rec_types (lna,typarray,_) env =
- let ctxt = Array.map2_i (fun i na t -> LocalAssum (na, lift i t)) lna typarray in
- Array.fold_left (fun e assum -> push_rel assum e) env ctxt
-
-(* Universe constraints *)
-let map_universes f env =
- let s = env.env_stratification in
- { env with env_stratification =
- { s with env_universes = f s.env_universes } }
-
-let add_constraints c env =
- if c == Univ.Constraint.empty then env
- else map_universes (Univ.merge_constraints c) env
-
-let push_context ?(strict=false) ctx env =
- map_universes (Univ.merge_context strict ctx) env
-
-let push_context_set ?(strict=false) ctx env =
- map_universes (Univ.merge_context_set strict ctx) env
-
-let check_constraints cst env =
- Univ.check_constraints cst env.env_stratification.env_universes
-
-(* Global constants *)
-
-let lookup_constant kn env =
- Cmap_env.find kn env.env_globals.env_constants
-
-let anomaly s = anomaly (Pp.str s)
-
-let add_constant kn cs env =
- if Cmap_env.mem kn env.env_globals.env_constants then
- Printf.ksprintf anomaly ("Constant %s is already defined.")
- (Constant.to_string kn);
- let new_constants =
- Cmap_env.add kn cs env.env_globals.env_constants in
- let new_globals =
- { env.env_globals with
- env_constants = new_constants } in
- { env with env_globals = new_globals }
-
-type const_evaluation_result = NoBody | Opaque
-
-(* Constant types *)
-
-let constraints_of cb u =
- match cb.const_universes with
- | Monomorphic_const _ -> Univ.Constraint.empty
- | Polymorphic_const ctx -> Univ.AUContext.instantiate u ctx
-
-(* constant_type gives the type of a constant *)
-let constant_type env (kn,u) =
- let cb = lookup_constant kn env in
- match cb.const_universes with
- | Monomorphic_const _ -> cb.const_type, Univ.Constraint.empty
- | Polymorphic_const ctx ->
- let csts = constraints_of cb u in
- (subst_instance_constr u cb.const_type, csts)
-
-exception NotEvaluableConst of const_evaluation_result
-
-let constant_value env (kn,u) =
- let cb = lookup_constant kn env in
- match cb.const_body with
- | Def l_body ->
- let b = force_constr l_body in
- begin
- match cb.const_universes with
- | Monomorphic_const _ -> b
- | Polymorphic_const _ -> subst_instance_constr u (force_constr l_body)
- end
- | OpaqueDef _ -> raise (NotEvaluableConst Opaque)
- | Undef _ -> raise (NotEvaluableConst NoBody)
-
-(* A global const is evaluable if it is defined and not opaque *)
-let evaluable_constant cst env =
- try let _ = constant_value env (cst, Univ.Instance.empty) in true
- with Not_found | NotEvaluableConst _ -> false
-
-(* Mutual Inductives *)
-let scrape_mind env kn=
- try
- KNmap.find kn env.env_globals.env_inductives_eq
- with
- Not_found -> kn
-
-let mind_equiv env (kn1,i1) (kn2,i2) =
- Int.equal i1 i2 &&
- KerName.equal
- (scrape_mind env (MutInd.user kn1))
- (scrape_mind env (MutInd.user kn2))
-
-
-let lookup_mind kn env =
- Mindmap_env.find kn env.env_globals.env_inductives
-
-let add_mind kn mib env =
- if Mindmap_env.mem kn env.env_globals.env_inductives then
- Printf.ksprintf anomaly ("Mutual inductive block %s is already defined.")
- (MutInd.to_string kn);
- let new_inds = Mindmap_env.add kn mib env.env_globals.env_inductives in
- let kn1,kn2 = MutInd.user kn, MutInd.canonical kn in
- let new_inds_eq = if KerName.equal kn1 kn2 then
- env.env_globals.env_inductives_eq
- else
- KNmap.add kn1 kn2 env.env_globals.env_inductives_eq in
- let new_globals =
- { env.env_globals with
- env_inductives = new_inds;
- env_inductives_eq = new_inds_eq} in
- { env with env_globals = new_globals }
-
-let lookup_projection p env =
- let mind,i = Projection.inductive p in
- let mib = lookup_mind mind env in
- match mib.mind_record with
- | NotRecord | FakeRecord -> CErrors.anomaly ~label:"lookup_projection" Pp.(str "not a projection")
- | PrimRecord infos ->
- let _,labs,typs = infos.(i) in
- let parg = Projection.arg p in
- if not (Label.equal labs.(parg) (Projection.label p))
- then CErrors.anomaly ~label:"lookup_projection" Pp.(str "incorrect label on projection")
- else if not (Int.equal mib.mind_nparams (Projection.npars p))
- then CErrors.anomaly ~label:"lookup_projection" Pp.(str "incorrect param number on projection")
- else typs.(parg)
-
-(* Modules *)
-
-let add_modtype ln mtb env =
- if MPmap.mem ln env.env_globals.env_modtypes then
- Printf.ksprintf anomaly ("Module type %s is already defined.")
- (ModPath.to_string ln);
- let new_modtypes = MPmap.add ln mtb env.env_globals.env_modtypes in
- let new_globals =
- { env.env_globals with
- env_modtypes = new_modtypes } in
- { env with env_globals = new_globals }
-
-let shallow_add_module mp mb env =
- if MPmap.mem mp env.env_globals.env_modules then
- Printf.ksprintf anomaly ("Module %s is already defined.")
- (ModPath.to_string mp);
- let new_mods = MPmap.add mp mb env.env_globals.env_modules in
- let new_globals =
- { env.env_globals with
- env_modules = new_mods } in
- { env with env_globals = new_globals }
-
-let shallow_remove_module mp env =
- if not (MPmap.mem mp env.env_globals.env_modules) then
- Printf.ksprintf anomaly ("Module %s is unknown.")
- (ModPath.to_string mp);
- let new_mods = MPmap.remove mp env.env_globals.env_modules in
- let new_globals =
- { env.env_globals with
- env_modules = new_mods } in
- { env with env_globals = new_globals }
-
-let lookup_module mp env =
- MPmap.find mp env.env_globals.env_modules
-
-let lookup_modtype ln env =
- MPmap.find ln env.env_globals.env_modtypes
diff --git a/checker/environ.mli b/checker/environ.mli
deleted file mode 100644
index af1b2a6228..0000000000
--- a/checker/environ.mli
+++ /dev/null
@@ -1,79 +0,0 @@
-open Names
-open Cic
-
-(* Environments *)
-
-type globals = {
- env_constants : constant_body Cmap_env.t;
- env_inductives : mutual_inductive_body Mindmap_env.t;
- env_inductives_eq : KerName.t KNmap.t;
- env_modules : module_body MPmap.t;
- env_modtypes : module_type_body MPmap.t}
-type stratification = {
- env_universes : Univ.universes;
- env_engagement : engagement;
-}
-type env = {
- env_globals : globals;
- env_rel_context : rel_context;
- env_stratification : stratification;
- env_imports : Cic.vodigest MPmap.t;
- env_conv_oracle : Cic.oracle;
-}
-val empty_env : env
-
-(* Engagement *)
-val engagement : env -> Cic.engagement
-val set_engagement : Cic.engagement -> env -> env
-
-(** Oracle *)
-
-val set_oracle : env -> Cic.oracle -> env
-
-(* Digests *)
-val add_digest : env -> DirPath.t -> Cic.vodigest -> env
-val lookup_digest : env -> DirPath.t -> Cic.vodigest
-
-(* de Bruijn variables *)
-val rel_context : env -> rel_context
-val lookup_rel : int -> env -> rel_declaration
-val push_rel : rel_declaration -> env -> env
-val push_rel_context : rel_context -> env -> env
-val push_rec_types : Name.t array * constr array * 'a -> env -> env
-
-(* Universes *)
-val universes : env -> Univ.universes
-val add_constraints : Univ.constraints -> env -> env
-val push_context : ?strict:bool -> Univ.universe_context -> env -> env
-val push_context_set : ?strict:bool -> Univ.universe_context_set -> env -> env
-val check_constraints : Univ.constraints -> env -> bool
-
-(* Constants *)
-val lookup_constant : Constant.t -> env -> Cic.constant_body
-val add_constant : Constant.t -> Cic.constant_body -> env -> env
-val constant_type : env -> Constant.t puniverses -> constr Univ.constrained
-type const_evaluation_result = NoBody | Opaque
-exception NotEvaluableConst of const_evaluation_result
-val constant_value : env -> Constant.t puniverses -> constr
-val evaluable_constant : Constant.t -> env -> bool
-
-(** NB: here in the checker we check the inferred info (npars, label) *)
-val lookup_projection : Projection.t -> env -> constr
-
-(* Inductives *)
-val mind_equiv : env -> inductive -> inductive -> bool
-
-val lookup_mind :
- MutInd.t -> env -> Cic.mutual_inductive_body
-
-val add_mind :
- MutInd.t -> Cic.mutual_inductive_body -> env -> env
-
-(* Modules *)
-val add_modtype :
- ModPath.t -> Cic.module_type_body -> env -> env
-val shallow_add_module :
- ModPath.t -> Cic.module_body -> env -> env
-val shallow_remove_module : ModPath.t -> env -> env
-val lookup_module : ModPath.t -> env -> Cic.module_body
-val lookup_modtype : ModPath.t -> env -> Cic.module_type_body
diff --git a/checker/indtypes.ml b/checker/indtypes.ml
deleted file mode 100644
index f6c510ee1c..0000000000
--- a/checker/indtypes.ml
+++ /dev/null
@@ -1,652 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-open CErrors
-open Util
-open Names
-open Cic
-open Term
-open Inductive
-open Reduction
-open Typeops
-open Pp
-open Declarations
-open Environ
-
-let rec debug_string_of_mp = function
- | MPfile sl -> DirPath.to_string sl
- | MPbound uid -> "bound("^MBId.to_string uid^")"
- | MPdot (mp,l) -> debug_string_of_mp mp ^ "." ^ Label.to_string l
-
-let rec string_of_mp = function
- | MPfile sl -> DirPath.to_string sl
- | MPbound uid -> MBId.to_string uid
- | MPdot (mp,l) -> string_of_mp mp ^ "." ^ Label.to_string l
-
-let string_of_mp mp =
- if !Flags.debug then debug_string_of_mp mp else string_of_mp mp
-
-let prkn kn =
- let (mp,l) = KerName.repr kn in
- str(string_of_mp mp ^ "." ^ Label.to_string l)
-let prcon c =
- let ck = Constant.canonical c in
- let uk = Constant.user c in
- if KerName.equal ck uk then prkn uk else (prkn uk ++str"(="++prkn ck++str")")
-
-(* Same as noccur_between but may perform reductions.
- Could be refined more... *)
-let weaker_noccur_between env x nvars t =
- if noccur_between x nvars t then Some t
- else
- let t' = whd_all env t in
- if noccur_between x nvars t' then Some t'
- else None
-
-let is_constructor_head t =
- match fst(decompose_app t) with
- | Rel _ -> true
- | _ -> false
-
-let conv_ctxt_prefix env (ctx1:rel_context) ctx2 =
- let rec chk env rctx1 rctx2 =
- match rctx1, rctx2 with
- (LocalAssum (_,ty1) as d1)::rctx1', LocalAssum (_,ty2)::rctx2' ->
- conv env ty1 ty2;
- chk (push_rel d1 env) rctx1' rctx2'
- | (LocalDef (_,bd1,ty1) as d1)::rctx1', LocalDef (_,bd2,ty2)::rctx2' ->
- conv env ty1 ty2;
- conv env bd1 bd2;
- chk (push_rel d1 env) rctx1' rctx2'
- | [],_ -> ()
- | _ -> failwith "non convertible contexts" in
- chk env (List.rev ctx1) (List.rev ctx2)
-
-(************************************************************************)
-(* Various well-formedness check for inductive declarations *)
-
-(* Errors related to inductive constructions *)
-type inductive_error =
- | NonPos of env * constr * constr
- | NotEnoughArgs of env * constr * constr
- | NotConstructor of env * constr * constr
- | NonPar of env * constr * int * constr * constr
- | SameNamesTypes of Id.t
- | SameNamesConstructors of Id.t
- | SameNamesOverlap of Id.t list
- | NotAnArity of Id.t
- | BadEntry
-
-exception InductiveError of inductive_error
-
-(************************************************************************)
-(************************************************************************)
-
-(* Typing the arities and constructor types *)
-
-let rec sorts_of_constr_args env t =
- let t = whd_allnolet env t in
- match t with
- | Prod (name,c1,c2) ->
- let varj = infer_type env c1 in
- let env1 = push_rel (LocalAssum (name,c1)) env in
- varj :: sorts_of_constr_args env1 c2
- | LetIn (name,def,ty,c) ->
- let env1 = push_rel (LocalDef (name,def,ty)) env in
- sorts_of_constr_args env1 c
- | _ when is_constructor_head t -> []
- | _ -> anomaly ~label:"infos_and_sort" (Pp.str "not a positive constructor.")
-
-
-(* Prop and Set are small *)
-let is_small_sort = function
- | Prop | Set -> true
- | _ -> false
-
-let is_logic_sort = function
-| Prop -> true
-| _ -> false
-
-(* [infos] is a sequence of pair [islogic,issmall] for each type in
- the product of a constructor or arity *)
-
-let is_small_constr infos = List.for_all (fun s -> is_small_sort s) infos
-let is_logic_constr infos = List.for_all (fun s -> is_logic_sort s) infos
-
-(* 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 *)
- | [|constrinfos|] -> is_logic_constr constrinfos
- | [||] -> (* type without constructors *) true
- | _ -> false
-
-let small_unit constrsinfos =
- let issmall = Array.for_all is_small_constr constrsinfos
- and isunit = is_unit constrsinfos in
- issmall, isunit
-
-(* check information related to inductive arity *)
-let typecheck_arity env params inds =
- let nparamargs = rel_context_nhyps params in
- let nparamdecls = rel_context_length params in
- let check_arity arctxt = function
- | RegularArity mar ->
- let ar = mar.mind_user_arity in
- let _ = infer_type env ar in
- conv env (it_mkProd_or_LetIn (Sort mar.mind_sort) arctxt) ar;
- ar
- | TemplateArity par ->
- check_polymorphic_arity env params par;
- it_mkProd_or_LetIn (Sort(Type par.template_level)) arctxt
- in
- let env_arities =
- Array.fold_left
- (fun env_ar ind ->
- let ar_ctxt = ind.mind_arity_ctxt in
- let _ = check_ctxt env ar_ctxt in
- conv_ctxt_prefix env params ar_ctxt;
- (* Arities (with params) are typed-checked here *)
- let arity = check_arity ar_ctxt ind.mind_arity in
- (* mind_nrealargs *)
- let nrealargs = rel_context_nhyps ar_ctxt - nparamargs in
- if ind.mind_nrealargs <> nrealargs then
- failwith "bad number of real inductive arguments";
- let nrealargs_ctxt = rel_context_length ar_ctxt - nparamdecls in
- if ind.mind_nrealdecls <> nrealargs_ctxt then
- failwith "bad length of real inductive arguments signature";
- (* We do not need to generate the universe of full_arity; if
- later, after the validation of the inductive definition,
- full_arity is used as argument or subject to cast, an
- upper universe will be generated *)
- let id = ind.mind_typename in
- let env_ar' = push_rel (LocalAssum (Name id, arity)) env_ar in
- env_ar')
- env
- inds in
- env_arities
-
-(* Allowed eliminations *)
-
-let check_predicativity env s small level =
- match s, engagement env with
- Type u, _ ->
- (* let u' = fresh_local_univ () in *)
- (* let cst = *)
- (* merge_constraints (enforce_leq u u' empty_constraint) *)
- (* (universes env) in *)
- if not (Univ.check_leq (universes env) level u) then
- failwith "impredicative Type inductive type"
- | Set, ImpredicativeSet -> ()
- | Set, _ ->
- if not small then failwith "impredicative Set inductive type"
- | Prop,_ -> ()
-
-
-let sort_of_ind = function
- | RegularArity mar -> mar.mind_sort
- | TemplateArity par -> Type par.template_level
-
-let all_sorts = [InProp;InSet;InType]
-let small_sorts = [InProp;InSet]
-let logical_sorts = [InProp]
-
-let allowed_sorts issmall isunit s =
- match family_of_sort s with
- (* Type: all elimination allowed *)
- | InType -> all_sorts
-
- (* Small Set is predicative: all elimination allowed *)
- | InSet when issmall -> all_sorts
-
- (* Large Set is necessarily impredicative: forbids large elimination *)
- | InSet -> small_sorts
-
- (* Unitary/empty Prop: elimination to all sorts are realizable *)
- (* unless the type is large. If it is large, forbids large elimination *)
- (* which otherwise allows simulating the inconsistent system Type:Type *)
- | InProp when isunit -> if issmall then all_sorts else small_sorts
-
- (* Other propositions: elimination only to Prop *)
- | InProp -> logical_sorts
-
-
-
-let compute_elim_sorts env_ar params arity lc =
- let inst = extended_rel_list 0 params in
- let env_params = push_rel_context params env_ar in
- let lc = Array.map
- (fun c ->
- hnf_prod_applist env_params (lift (rel_context_length params) c) inst)
- lc in
- let s = sort_of_ind arity in
- let infos = Array.map (sorts_of_constr_args env_params) lc in
- let (small,unit) = small_unit infos in
- (* We accept recursive unit types... *)
- (* compute the max of the sorts of the products of the constructor type *)
- let level = max_inductive_sort
- (Array.concat (Array.to_list (Array.map Array.of_list infos))) in
- check_predicativity env_ar s small level;
- allowed_sorts small unit s
-
-
-let typecheck_one_inductive env params mip =
- (* mind_typename and mind_consnames not checked *)
- (* mind_reloc_tbl, mind_nb_constant, mind_nb_args not checked (VM) *)
- (* mind_arity_ctxt, mind_arity, mind_nrealargs DONE (typecheck_arity) *)
- (* mind_user_lc *)
- let _ = Array.map (infer_type env) mip.mind_user_lc in
- (* mind_nf_lc *)
- let _ = Array.map (infer_type env) mip.mind_nf_lc in
- Array.iter2 (conv env) mip.mind_nf_lc mip.mind_user_lc;
- (* mind_consnrealdecls *)
- let check_cons_args c n =
- let ctx,_ = decompose_prod_assum c in
- if n <> rel_context_length ctx - rel_context_length params then
- failwith "bad number of real constructor arguments" in
- Array.iter2 check_cons_args mip.mind_nf_lc mip.mind_consnrealdecls;
- (* mind_kelim: checked by positivity criterion ? *)
- let sorts =
- compute_elim_sorts env params mip.mind_arity mip.mind_nf_lc in
- let reject_sort s = not (List.mem_f family_equal s sorts) in
- if List.exists reject_sort mip.mind_kelim then
- failwith "elimination not allowed";
- (* mind_recargs: checked by positivity criterion *)
- ()
-
-(************************************************************************)
-(************************************************************************)
-(* Positivity *)
-
-type ill_formed_ind =
- | LocalNonPos of int
- | LocalNotEnoughArgs of int
- | LocalNotConstructor
- | LocalNonPar of int * int * int
-
-exception IllFormedInd of ill_formed_ind
-
-(* [mind_extract_params mie] extracts the params from an inductive types
- declaration, and checks that they are all present (and all the same)
- for all the given types. *)
-
-let mind_extract_params = decompose_prod_n_assum
-
-let explain_ind_err ntyp env0 nbpar c err =
- let (lpar,c') = mind_extract_params nbpar c in
- let env = push_rel_context lpar env0 in
- match err with
- | LocalNonPos kt ->
- raise (InductiveError (NonPos (env,c',Rel (kt+nbpar))))
- | LocalNotEnoughArgs kt ->
- raise (InductiveError
- (NotEnoughArgs (env,c',Rel (kt+nbpar))))
- | LocalNotConstructor ->
- raise (InductiveError
- (NotConstructor (env,c',Rel (ntyp+nbpar))))
- | LocalNonPar (n,i,l) ->
- raise (InductiveError
- (NonPar (env,c',n,Rel i,Rel (l+nbpar))))
-
-let failwith_non_pos n ntypes c =
- for k = n to n + ntypes - 1 do
- if not (noccurn k c) then raise (IllFormedInd (LocalNonPos (k-n+1)))
- done
-
-let failwith_non_pos_vect n ntypes v =
- Array.iter (failwith_non_pos n ntypes) v;
- anomaly ~label:"failwith_non_pos_vect" (Pp.str "some k in [n;n+ntypes-1] should occur.")
-
-let failwith_non_pos_list n ntypes l =
- List.iter (failwith_non_pos n ntypes) l;
- anomaly ~label:"failwith_non_pos_list" (Pp.str "some k in [n;n+ntypes-1] should occur.")
-
-(* Check the inductive type is called with the expected parameters *)
-(* [n] is the index of the last inductive type in [env] *)
-let check_correct_par (env,n,ntypes,_) paramdecls ind_index args =
- let nparams = rel_context_nhyps paramdecls in
- let args = Array.of_list args in
- if Array.length args < nparams then
- raise (IllFormedInd (LocalNotEnoughArgs ind_index));
- let (params,realargs) = Array.chop nparams args in
- let nparamdecls = List.length paramdecls in
- let rec check param_index paramdecl_index = function
- | [] -> ()
- | LocalDef _ :: paramdecls ->
- check param_index (paramdecl_index+1) paramdecls
- | _::paramdecls ->
- match whd_all env params.(param_index) with
- | Rel w when Int.equal w paramdecl_index ->
- check (param_index-1) (paramdecl_index+1) paramdecls
- | _ ->
- let paramdecl_index_in_env = paramdecl_index-n+nparamdecls+1 in
- let err =
- LocalNonPar (param_index+1, paramdecl_index_in_env, ind_index) in
- raise (IllFormedInd err)
- in check (nparams-1) (n-nparamdecls) paramdecls;
- if not (Array.for_all (noccur_between n ntypes) realargs) then
- failwith_non_pos_vect n ntypes realargs
-
-(* Arguments of constructor: check the number of recursive parameters nrecp.
- the first parameters which are constant in recursive arguments
- n is the current depth, nmr is the maximum number of possible
- recursive parameters *)
-
-let check_rec_par (env,n,_,_) hyps nrecp largs =
- let (lpar,_) = List.chop nrecp largs in
- let rec find index =
- function
- | ([],_) -> ()
- | (_,[]) ->
- failwith "number of recursive parameters cannot be greater than the number of parameters."
- | (lp,LocalDef _ :: hyps) -> find (index-1) (lp,hyps)
- | (p::lp,_::hyps) ->
- (match whd_all env p with
- | Rel w when w = index -> find (index-1) (lp,hyps)
- | _ -> failwith "bad number of recursive parameters")
- in find (n-1) (lpar,List.rev hyps)
-
-let lambda_implicit_lift n a =
- let lambda_implicit a = Lambda(Anonymous,Evar(0,[||]),a) in
- iterate lambda_implicit n (lift n a)
-
-(* This removes global parameters of the inductive types in lc (for
- nested inductive types only ) *)
-let abstract_mind_lc ntyps npars lc =
- if npars = 0 then
- lc
- else
- let make_abs =
- List.init ntyps
- (function i -> lambda_implicit_lift npars (Rel (i+1)))
- in
- Array.map (substl make_abs) lc
-
-(* [env] is the typing environment
- [n] is the dB of the last inductive type
- [ntypes] is the number of inductive types in the definition
- (i.e. range of inductives is [n; n+ntypes-1])
- [lra] is the list of recursive tree of each variable
- *)
-let ienv_push_var (env, n, ntypes, lra) (x,a,ra) =
- (push_rel (LocalAssum (x,a)) env, n+1, ntypes, (Norec,ra)::lra)
-
-let ienv_push_inductive (env, n, ntypes, ra_env) ((mi,u),lpar) =
- let auxntyp = 1 in
- let specif = lookup_mind_specif env mi in
- let env' =
- let decl = LocalAssum (Anonymous,
- hnf_prod_applist env (type_of_inductive env (specif,u)) lpar) in
- push_rel decl env in
- let ra_env' =
- (Imbr mi,(Rtree.mk_rec_calls 1).(0)) ::
- List.map (fun (r,t) -> (r,Rtree.lift 1 t)) ra_env in
- (* New index of the inductive types *)
- let newidx = n + auxntyp in
- (env', newidx, ntypes, ra_env')
-
-let rec ienv_decompose_prod (env,_,_,_ as ienv) n c =
- if n=0 then (ienv,c) else
- let c' = whd_all env c in
- match c' with
- Prod(na,a,b) ->
- let ienv' = ienv_push_var ienv (na,a,mk_norec) in
- ienv_decompose_prod ienv' (n-1) b
- | _ -> assert false
-
-(* The recursive function that checks positivity and builds the list
- of recursive arguments *)
-let check_positivity_one (env, _,ntypes,_ as ienv) hyps nrecp (_,i as ind) indlc =
- let lparams = rel_context_length hyps in
- (* check the inductive types occur positively in [c] *)
- let rec check_pos (env, n, ntypes, ra_env as ienv) c =
- let x,largs = decompose_app (whd_all env c) in
- match x with
- | Prod (na,b,d) ->
- assert (List.is_empty largs);
- (match weaker_noccur_between env n ntypes b with
- None -> failwith_non_pos_list n ntypes [b]
- | Some b ->
- check_pos (ienv_push_var ienv (na, b, mk_norec)) d)
- | Rel k ->
- (try let (ra,rarg) = List.nth ra_env (k-1) in
- let largs = List.map (whd_all env) largs in
- (match ra with
- Mrec _ -> check_rec_par ienv hyps nrecp largs
- | _ -> ());
- if not (List.for_all (noccur_between n ntypes) largs)
- then failwith_non_pos_list n ntypes largs
- else rarg
- with Failure _ | Invalid_argument _ -> mk_norec)
- | Ind ind_kn ->
- (* If the inductive type being defined appears in a
- parameter, then we have an imbricated type *)
- if List.for_all (noccur_between n ntypes) largs then mk_norec
- else check_positive_imbr ienv (ind_kn, largs)
- | err ->
- if noccur_between n ntypes x &&
- List.for_all (noccur_between n ntypes) largs
- then mk_norec
- else failwith_non_pos_list n ntypes (x::largs)
-
- (* accesses to the environment are not factorised, but is it worth it? *)
- and check_positive_imbr (env,n,ntypes,ra_env as ienv) ((mi,u), largs) =
- let (mib,mip) = lookup_mind_specif env mi in
- let auxnpar = mib.mind_nparams_rec in
- let nonrecpar = mib.mind_nparams - auxnpar in
- let (lpar,auxlargs) =
- try List.chop auxnpar largs
- with Failure _ -> raise (IllFormedInd (LocalNonPos n)) in
- (* If the inductive appears in the args (non params) then the
- definition is not positive. *)
- if not (List.for_all (noccur_between n ntypes) auxlargs) then
- raise (IllFormedInd (LocalNonPos n));
- (* We do not deal with imbricated mutual inductive types *)
- let auxntyp = mib.mind_ntypes in
- if auxntyp <> 1 then raise (IllFormedInd (LocalNonPos n));
- (* The nested inductive type with parameters removed *)
- let auxlcvect = abstract_mind_lc auxntyp auxnpar mip.mind_nf_lc in
- (* Extends the environment with a variable corresponding to
- the inductive def *)
- let (env',_,_,_ as ienv') = ienv_push_inductive ienv ((mi,u),lpar) in
- (* Parameters expressed in env' *)
- let lpar' = List.map (lift auxntyp) lpar in
- let irecargs =
- (* fails if the inductive type occurs non positively *)
- (* with recursive parameters substituted *)
- Array.map
- (function c ->
- let c' = hnf_prod_applist env' c lpar' in
- (* skip non-recursive parameters *)
- let (ienv',c') = ienv_decompose_prod ienv' nonrecpar c' in
- check_constructors ienv' false c')
- auxlcvect in
- (Rtree.mk_rec [|mk_paths (Imbr mi) irecargs|]).(0)
-
- (* check the inductive types occur positively in the products of C, if
- check_head=true, also check the head corresponds to a constructor of
- the ith type *)
-
- and check_constructors ienv check_head c =
- let rec check_constr_rec (env,n,ntypes,ra_env as ienv) lrec c =
- let x,largs = decompose_app (whd_all env c) in
- match x with
- | Prod (na,b,d) ->
- assert (List.is_empty largs);
- let recarg = check_pos ienv b in
- let ienv' = ienv_push_var ienv (na,b,mk_norec) in
- check_constr_rec ienv' (recarg::lrec) d
-
- | hd ->
- if check_head then
- match hd with
- | Rel j when j = (n + ntypes - i - 1) ->
- check_correct_par ienv hyps (ntypes-i) largs
- | _ ->
- raise (IllFormedInd LocalNotConstructor)
- else
- if not (List.for_all (noccur_between n ntypes) largs)
- then raise (IllFormedInd (LocalNonPos n));
- List.rev lrec
- in check_constr_rec ienv [] c
- in
- let irecargs =
- Array.map
- (fun c ->
- let _,rawc = mind_extract_params lparams c in
- try
- check_constructors ienv true rawc
- with IllFormedInd err ->
- explain_ind_err (ntypes-i) env lparams c err)
- indlc
- in mk_paths (Mrec ind) irecargs
-
-let prrecarg = function
- | Norec -> str "Norec"
- | Mrec (mind,i) ->
- str "Mrec[" ++ MutInd.debug_print mind ++ pr_comma () ++ int i ++ str "]"
- | Imbr (mind,i) ->
- str "Imbr[" ++ MutInd.debug_print mind ++ pr_comma () ++ int i ++ str "]"
-
-let check_subtree t1 t2 =
- let cmp_labels l1 l2 = l1 == Norec || eq_recarg l1 l2 in
- if not (Rtree.equiv eq_recarg cmp_labels t1 t2)
- then user_err Pp.(str "Bad recursive tree: found " ++ fnl ()
- ++ Rtree.pp_tree prrecarg t1 ++ fnl () ++ str " when expected " ++ fnl ()
- ++ Rtree.pp_tree prrecarg t2)
-(* if t1=t2 then () else msg_warning (str"TODO: check recursive positions")*)
-
-let check_positivity env_ar mind params nrecp inds =
- let ntypes = Array.length inds in
- let rc =
- Array.mapi (fun j t -> (Mrec(mind,j),t)) (Rtree.mk_rec_calls ntypes) in
- let lra_ind = List.rev (Array.to_list rc) in
- let lparams = rel_context_length params in
- let ra_env =
- List.init lparams (fun _ -> (Norec,mk_norec)) @ lra_ind in
- let env_ar_par = push_rel_context params env_ar in
- let check_one i mip =
- let ienv = (env_ar_par, 1+lparams, ntypes, ra_env) in
- check_positivity_one ienv params nrecp (mind,i) mip.mind_nf_lc
- in
- let irecargs = Array.mapi check_one inds in
- let wfp = Rtree.mk_rec irecargs in
- Array.iter2 (fun ind wfpi -> check_subtree ind.mind_recargs wfpi) inds wfp
-
-(* Check arities and constructors *)
-let check_subtyping_arity_constructor env (subst : Univ.Instance.t) (arcn : constr) numparams is_arity =
- let numchecked = ref 0 in
- let basic_check ev tp =
- if !numchecked < numparams then () else conv_leq ev tp (Term.subst_instance_constr subst tp);
- numchecked := !numchecked + 1
- in
- let check_typ typ typ_env =
- match typ with
- | LocalAssum (_, typ') ->
- begin
- try
- basic_check typ_env typ'; Environ.push_rel typ typ_env
- with NotConvertible ->
- anomaly ~label:"bad inductive subtyping relation" (Pp.str "Invalid subtyping relation")
- end
- | _ -> anomaly (Pp.str "")
- in
- let typs, codom = dest_prod env arcn in
- let last_env = fold_rel_context_outside check_typ typs ~init:env in
- if not is_arity then basic_check last_env codom else ()
-
-(* Check that the subtyping information inferred for inductive types in the block is correct. *)
-(* This check produces a value of the unit type if successful or raises an anomaly if check fails. *)
-let check_subtyping cumi paramsctxt env inds =
- let open Univ in
- let numparams = rel_context_nhyps paramsctxt in
- (** In [env] we already have [ Var(0) ... Var(n-1) |- cst ] available.
- We must produce the substitution σ : [ Var(i) -> Var (i + n) | 0 <= i < n]
- and push the constraints [ Var(n) ... Var(2n - 1) |- cst{σ} ], together
- with the cumulativity constraints [ cumul_cst ]. *)
- let uctx = ACumulativityInfo.univ_context cumi in
- let len = AUContext.size uctx in
- let inst = Instance.of_array @@ Array.init len (fun i -> Level.var (i + len)) in
-
- let other_context = ACumulativityInfo.univ_context cumi in
- let uctx_other = UContext.make (inst, AUContext.instantiate inst other_context) in
- let cumul_cst =
- Array.fold_left_i (fun i csts var ->
- match var with
- | Variance.Irrelevant -> csts
- | Variance.Covariant -> Constraint.add (Level.var i,Le,Level.var (i+len)) csts
- | Variance.Invariant -> Constraint.add (Level.var i,Eq,Level.var (i+len)) csts)
- Constraint.empty (ACumulativityInfo.variance cumi)
- in
- let env = Environ.push_context uctx_other env in
- let env = Environ.add_constraints cumul_cst env in
-
- (* process individual inductive types: *)
- Array.iter (fun { mind_user_lc = lc; mind_arity = arity } ->
- match arity with
- | RegularArity { mind_user_arity = full_arity} ->
- check_subtyping_arity_constructor env inst full_arity numparams true;
- Array.iter (fun cnt -> check_subtyping_arity_constructor env inst cnt numparams false) lc
- | TemplateArity _ -> ()
- ) inds
-
-(************************************************************************)
-(************************************************************************)
-
-let print_mutind ind =
- let kn = MutInd.user ind in
- str (ModPath.to_string (KerName.modpath kn) ^ "." ^ Label.to_string (KerName.label kn))
-
-let check_inductive env kn mib =
- Flags.if_verbose Feedback.msg_notice (str " checking mutind block: " ++ print_mutind kn);
- (* check mind_constraints: should be consistent with env *)
- let env0 =
- match mib.mind_universes with
- | Monomorphic_ind _ -> env
- | Polymorphic_ind auctx ->
- let uctx = Univ.AUContext.repr auctx in
- Environ.push_context uctx env
- | Cumulative_ind cumi ->
- let uctx = Univ.AUContext.repr (Univ.ACumulativityInfo.univ_context cumi) in
- Environ.push_context uctx env
- in
- (** Locally set the oracle for further typechecking *)
- let env0 = Environ.set_oracle env0 mib.mind_typing_flags.conv_oracle in
- (* check mind_record : TODO ? check #constructor = 1 ? *)
- (* check mind_finite : always OK *)
- (* check mind_ntypes *)
- if Array.length mib.mind_packets <> mib.mind_ntypes then
- user_err Pp.(str "not the right number of packets");
- (* check mind_params_ctxt *)
- let params = mib.mind_params_ctxt in
- let _ = check_ctxt env0 params in
- (* check mind_nparams *)
- if rel_context_nhyps params <> mib.mind_nparams then
- user_err Pp.(str "number the right number of parameters");
- (* mind_packets *)
- (* - check arities *)
- let env_ar = typecheck_arity env0 params mib.mind_packets in
- (* - check constructor types *)
- Array.iter (typecheck_one_inductive env_ar params) mib.mind_packets;
- (* check the inferred subtyping relation *)
- let () =
- match mib.mind_universes with
- | Monomorphic_ind _ | Polymorphic_ind _ -> ()
- | Cumulative_ind acumi ->
- check_subtyping acumi params env_ar mib.mind_packets
- in
- (* check mind_nparams_rec: positivity condition *)
- check_positivity env_ar kn params mib.mind_nparams_rec mib.mind_packets;
- (* check mind_equiv... *)
- (* Now we can add the inductive *)
- add_mind kn mib env
-
diff --git a/checker/indtypes.mli b/checker/indtypes.mli
deleted file mode 100644
index baaa66a6c1..0000000000
--- a/checker/indtypes.mli
+++ /dev/null
@@ -1,39 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(*i*)
-open Names
-open Cic
-open Environ
-(*i*)
-
-val prkn : KerName.t -> Pp.t
-val prcon : Constant.t -> Pp.t
-
-(*s The different kinds of errors that may result of a malformed inductive
- definition. *)
-
-(* Errors related to inductive constructions *)
-type inductive_error =
- | NonPos of env * constr * constr
- | NotEnoughArgs of env * constr * constr
- | NotConstructor of env * constr * constr
- | NonPar of env * constr * int * constr * constr
- | SameNamesTypes of Id.t
- | SameNamesConstructors of Id.t
- | SameNamesOverlap of Id.t list
- | NotAnArity of Id.t
- | BadEntry
-
-exception InductiveError of inductive_error
-
-(*s The following function does checks on inductive declarations. *)
-
-val check_inductive : env -> MutInd.t -> mutual_inductive_body -> env
diff --git a/checker/inductive.ml b/checker/inductive.ml
deleted file mode 100644
index 269a98cb0e..0000000000
--- a/checker/inductive.ml
+++ /dev/null
@@ -1,1205 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-open CErrors
-open Util
-open Names
-open Cic
-open Term
-open Reduction
-open Type_errors
-open Declarations
-open Environ
-open Univ
-
-let inductive_of_constructor = fst
-let index_of_constructor = snd
-let ith_constructor_of_inductive ind i = (ind,i)
-
-type mind_specif = mutual_inductive_body * one_inductive_body
-
-(* raise Not_found if not an inductive type *)
-let lookup_mind_specif env (kn,tyi) =
- let mib = lookup_mind kn env in
- if tyi >= Array.length mib.mind_packets then
- user_err Pp.(str "Inductive.lookup_mind_specif: invalid inductive index");
- (mib, mib.mind_packets.(tyi))
-
-let find_rectype env c =
- let (t, l) = decompose_app (whd_all env c) in
- match t with
- | Ind ind -> (ind, l)
- | _ -> raise Not_found
-
-let find_inductive env c =
- let (t, l) = decompose_app (whd_all env c) in
- match t with
- | Ind (ind,_)
- when (fst (lookup_mind_specif env ind)).mind_finite != CoFinite -> (ind, l)
- | _ -> raise Not_found
-
-let find_coinductive env c =
- let (t, l) = decompose_app (whd_all env c) in
- match t with
- | Ind (ind,_)
- when (fst (lookup_mind_specif env ind)).mind_finite == CoFinite -> (ind, l)
- | _ -> raise Not_found
-
-let inductive_params (mib,_) = mib.mind_nparams
-
-(** Polymorphic inductives *)
-
-let inductive_is_polymorphic mib =
- match mib.mind_universes with
- | Monomorphic_ind _ -> false
- | Polymorphic_ind ctx -> true
- | Cumulative_ind cumi -> true
-
-let inductive_is_cumulative mib =
- match mib.mind_universes with
- | Monomorphic_ind _ -> false
- | Polymorphic_ind ctx -> false
- | Cumulative_ind cumi -> true
-
-(************************************************************************)
-
-(* Build the substitution that replaces Rels by the appropriate *)
-(* inductives *)
-let ind_subst mind mib u =
- let ntypes = mib.mind_ntypes in
- let make_Ik k = Ind ((mind,ntypes-k-1),u) in
- List.init ntypes make_Ik
-
-(* Instantiate inductives in constructor type *)
-let constructor_instantiate mind u mib c =
- let s = ind_subst mind mib u in
- substl s (subst_instance_constr u c)
-
-let instantiate_params full t u args sign =
- let fail () =
- anomaly ~label:"instantiate_params" (Pp.str "type, ctxt and args mismatch.") in
- let (rem_args, subs, ty) =
- fold_rel_context
- (fun decl (largs,subs,ty) ->
- match (decl, largs, ty) with
- | (LocalAssum _, a::args, Prod(_,_,t)) -> (args, a::subs, t)
- | (LocalDef (_,b,_),_,LetIn(_,_,_,t)) ->
- (largs, (substl subs (subst_instance_constr u b))::subs, t)
- | (_,[],_) -> if full then fail() else ([], subs, ty)
- | _ -> fail ())
- sign
- ~init:(args,[],t)
- in
- if rem_args <> [] then fail();
- substl subs ty
-
-let full_inductive_instantiate mib u params sign =
- let dummy = Prop in
- let t = mkArity (Term.subst_instance_context u sign,dummy) in
- fst (destArity (instantiate_params true t u params mib.mind_params_ctxt))
-
-let full_constructor_instantiate ((mind,_),u,(mib,_),params) t =
- let inst_ind = constructor_instantiate mind u mib t in
- instantiate_params true inst_ind u params mib.mind_params_ctxt
-
-(************************************************************************)
-(************************************************************************)
-
-(* Functions to build standard types related to inductive *)
-
-(*
-Computing the actual sort of an applied or partially applied inductive type:
-
-I_i: forall uniformparams:utyps, forall otherparams:otyps, Type(a)
-uniformargs : utyps
-otherargs : otyps
-I_1:forall ...,s_1;...I_n:forall ...,s_n |- sort(C_kj(uniformargs)) = s_kj
-s'_k = max(..s_kj..)
-merge(..s'_k..) = ..s''_k..
---------------------------------------------------------------------
-Gamma |- I_i uniformargs otherargs : phi(s''_i)
-
-where
-
-- if p=0, phi() = Prop
-- if p=1, phi(s) = s
-- if p<>1, phi(s) = sup(Set,s)
-
-Remark: Set (predicative) is encoded as Type(0)
-*)
-
-let sort_as_univ = function
-| Type u -> u
-| Prop -> Univ.type0m_univ
-| Set -> Univ.type0_univ
-
-(* cons_subst add the mapping [u |-> su] in subst if [u] is not *)
-(* in the domain or add [u |-> sup x su] if [u] is already mapped *)
-(* to [x]. *)
-let cons_subst u su subst =
- try
- Univ.LMap.add u (Univ.sup (Univ.LMap.find u subst) su) subst
- with Not_found -> Univ.LMap.add u su subst
-
-(* remember_subst updates the mapping [u |-> x] by [u |-> sup x u] *)
-(* if it is presents and returns the substitution unchanged if not.*)
-let remember_subst u subst =
- try
- let su = Universe.make u in
- Univ.LMap.add u (Univ.sup (Univ.LMap.find u subst) su) subst
- with Not_found -> subst
-
-(* Bind expected levels of parameters to actual levels *)
-(* Propagate the new levels in the signature *)
-let make_subst env =
- let rec make subst = function
- | LocalDef _ :: sign, exp, args ->
- make subst (sign, exp, args)
- | d::sign, None::exp, args ->
- let args = match args with _::args -> args | [] -> [] in
- make subst (sign, exp, args)
- | d::sign, Some u::exp, a::args ->
- (* We recover the level of the argument, but we don't change the *)
- (* level in the corresponding type in the arity; this level in the *)
- (* arity is a global level which, at typing time, will be enforce *)
- (* to be greater than the level of the argument; this is probably *)
- (* a useless extra constraint *)
- let s = sort_as_univ (snd (dest_arity env a)) in
- make (cons_subst u s subst) (sign, exp, args)
- | LocalAssum (na,t) :: sign, Some u::exp, [] ->
- (* No more argument here: we add the remaining universes to the *)
- (* substitution (when [u] is distinct from all other universes in the *)
- (* template, it is identity substitution otherwise (ie. when u is *)
- (* already in the domain of the substitution) [remember_subst] will *)
- (* update its image [x] by [sup x u] in order not to forget the *)
- (* dependency in [u] that remains to be fullfilled. *)
- make (remember_subst u subst) (sign, exp, [])
- | sign, [], _ ->
- (* Uniform parameters are exhausted *)
- subst
- | [], _, _ ->
- assert false
- in
- make Univ.LMap.empty
-
-let instantiate_universes env ctx ar argsorts =
- let args = Array.to_list argsorts in
- let subst = make_subst env (ctx,ar.template_param_levels,args) in
- let level = Univ.subst_univs_universe (Univ.make_subst subst) ar.template_level in
- let ty =
- (* Singleton type not containing types are interpretable in Prop *)
- if Univ.is_type0m_univ level then Prop
- (* Non singleton type not containing types are interpretable in Set *)
- else if Univ.is_type0_univ level then Set
- (* This is a Type with constraints *)
- else Type level
- in
- (ctx, ty)
-
-(* Type of an inductive type *)
-
-let type_of_inductive_gen env ((mib,mip),u) paramtyps =
- match mip.mind_arity with
- | RegularArity a ->
- if not (inductive_is_polymorphic mib) then a.mind_user_arity
- else subst_instance_constr u a.mind_user_arity
- | TemplateArity ar ->
- let ctx = List.rev mip.mind_arity_ctxt in
- let ctx,s = instantiate_universes env ctx ar paramtyps in
- mkArity (List.rev ctx,s)
-
-let type_of_inductive_knowing_parameters env mip args =
- type_of_inductive_gen env mip args
-
-(* Type of a (non applied) inductive type *)
-
-let type_of_inductive env mip =
- type_of_inductive_knowing_parameters env mip [||]
-
-(* The max of an array of universes *)
-
-let cumulate_constructor_univ u = function
- | Prop -> u
- | Set -> Univ.sup Univ.type0_univ u
- | Type u' -> Univ.sup u u'
-
-let max_inductive_sort =
- Array.fold_left cumulate_constructor_univ Univ.type0m_univ
-
-(************************************************************************)
-(* Type of a constructor *)
-
-let type_of_constructor_subst cstr u (mib,mip) =
- let ind = inductive_of_constructor cstr in
- let specif = mip.mind_user_lc in
- let i = index_of_constructor cstr in
- let nconstr = Array.length mip.mind_consnames in
- if i > nconstr then user_err Pp.(str "Not enough constructors in the type.");
- constructor_instantiate (fst ind) u mib specif.(i-1)
-
-let type_of_constructor_gen (cstr,u) (mib,mip as mspec) =
- type_of_constructor_subst cstr u mspec
-
-let type_of_constructor cstru mspec =
- type_of_constructor_gen cstru mspec
-
-let arities_of_specif (kn,u) (mib,mip) =
- let specif = mip.mind_nf_lc in
- Array.map (constructor_instantiate kn u mib) specif
-
-
-
-(************************************************************************)
-
-let error_elim_expln kp ki =
- match kp,ki with
- | (InType | InSet), InProp -> NonInformativeToInformative
- | InType, InSet -> StrongEliminationOnNonSmallType (* if Set impredicative *)
- | _ -> WrongArity
-
-(* Type of case predicates *)
-
-(* Get type of inductive, with parameters instantiated *)
-
-let inductive_sort_family mip =
- match mip.mind_arity with
- | RegularArity s -> family_of_sort s.mind_sort
- | TemplateArity _ -> InType
-
-let mind_arity mip =
- mip.mind_arity_ctxt, inductive_sort_family mip
-
-let get_instantiated_arity (ind,u) (mib,mip) params =
- let sign, s = mind_arity mip in
- full_inductive_instantiate mib u params sign, s
-
-let elim_sorts (_,mip) = mip.mind_kelim
-
-let is_primitive_record (mib,_) =
- match mib.mind_record with
- | PrimRecord _ -> true
- | NotRecord | FakeRecord -> false
-
-let extended_rel_list n hyps =
- let rec reln l p = function
- | LocalAssum _ :: hyps -> reln (Rel (n+p) :: l) (p+1) hyps
- | LocalDef _ :: hyps -> reln l (p+1) hyps
- | [] -> l
- in
- reln [] 1 hyps
-
-let build_dependent_inductive ind (_,mip) params =
- let realargs,_ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in
- applist
- (Ind ind,
- List.map (lift mip.mind_nrealdecls) params
- @ extended_rel_list 0 realargs)
-
-(* This exception is local *)
-exception LocalArity of (sorts_family * sorts_family * arity_error) option
-
-let check_allowed_sort ksort specif =
- if not (List.exists ((=) ksort) (elim_sorts specif)) then
- let s = inductive_sort_family (snd specif) in
- raise (LocalArity (Some(ksort,s,error_elim_expln ksort s)))
-
-let is_correct_arity env c (p,pj) ind specif params =
- let arsign,_ = get_instantiated_arity ind specif params in
- let rec srec env pt ar =
- let pt' = whd_all env pt in
- match pt', ar with
- | Prod (na1,a1,t), LocalAssum (_,a1')::ar' ->
- (try conv env a1 a1'
- with NotConvertible -> raise (LocalArity None));
- srec (push_rel (LocalAssum (na1,a1)) env) t ar'
- | Prod (na1,a1,a2), [] -> (* whnf of t was not needed here! *)
- let env' = push_rel (LocalAssum (na1,a1)) env in
- let ksort = match (whd_all env' a2) with
- | Sort s -> family_of_sort s
- | _ -> raise (LocalArity None) in
- let dep_ind = build_dependent_inductive ind specif params in
- (try conv env a1 dep_ind
- with NotConvertible -> raise (LocalArity None));
- check_allowed_sort ksort specif;
- true
- | Sort s', [] ->
- check_allowed_sort (family_of_sort s') specif;
- false
- | _, (LocalDef _ as d)::ar' ->
- srec (push_rel d env) (lift 1 pt') ar'
- | _ ->
- raise (LocalArity None)
- in
- try srec env pj (List.rev arsign)
- with LocalArity kinds ->
- error_elim_arity env ind (elim_sorts specif) c (p,pj) kinds
-
-
-(************************************************************************)
-(* Type of case branches *)
-
-(* [p] is the predicate, [i] is the constructor number (starting from 0),
- and [cty] is the type of the constructor (params not instantiated) *)
-let build_branches_type (ind,u) (_,mip as specif) params dep p =
- let build_one_branch i cty =
- let typi = full_constructor_instantiate (ind,u,specif,params) cty in
- let (args,ccl) = decompose_prod_assum typi in
- let nargs = rel_context_length args in
- let (_,allargs) = decompose_app ccl in
- let (lparams,vargs) = List.chop (inductive_params specif) allargs in
- let cargs =
- if dep then
- let cstr = ith_constructor_of_inductive ind (i+1) in
- let dep_cstr =
- applist (Construct (cstr,u),lparams@extended_rel_list 0 args) in
- vargs @ [dep_cstr]
- else
- vargs in
- let base = beta_appvect (lift nargs p) (Array.of_list cargs) in
- it_mkProd_or_LetIn base args in
- Array.mapi build_one_branch mip.mind_nf_lc
-
-(* [p] is the predicate, [c] is the match object, [realargs] is the
- list of real args of the inductive type *)
-let build_case_type dep p c realargs =
- let args = if dep then realargs@[c] else realargs in
- beta_appvect p (Array.of_list args)
-
-let type_case_branches env (pind,largs) (p,pj) c =
- let specif = lookup_mind_specif env (fst pind) in
- let nparams = inductive_params specif in
- let (params,realargs) = List.chop nparams largs in
- let dep = is_correct_arity env c (p,pj) pind specif params in
- let lc = build_branches_type pind specif params dep p in
- let ty = build_case_type dep p c realargs in
- (lc, ty)
-
-
-(************************************************************************)
-(* Checking the case annotation is relevant *)
-
-let check_case_info env indsp ci =
- let mib, mip as spec = lookup_mind_specif env indsp in
- if
- not (mind_equiv env indsp ci.ci_ind) ||
- (mib.mind_nparams <> ci.ci_npar) ||
- (mip.mind_consnrealdecls <> ci.ci_cstr_ndecls) ||
- (mip.mind_consnrealargs <> ci.ci_cstr_nargs) ||
- is_primitive_record spec
- then raise (TypeError(env,WrongCaseInfo(indsp,ci)))
-
-(************************************************************************)
-(************************************************************************)
-
-(* Guard conditions for fix and cofix-points *)
-
-(* Check if t is a subterm of Rel n, and gives its specification,
- assuming lst already gives index of
- subterms with corresponding specifications of recursive arguments *)
-
-(* A powerful notion of subterm *)
-
-(* To each inductive definition corresponds an array describing the
- structure of recursive arguments for each constructor, we call it
- the recursive spec of the type (it has type recargs vect). For
- checking the guard, we start from the decreasing argument (Rel n)
- with its recursive spec. During checking the guardness condition,
- we collect patterns variables corresponding to subterms of n, each
- of them with its recursive spec. They are organised in a list lst
- of type (int * recargs) list which is sorted with respect to the
- first argument.
-*)
-
-(*************************************************************)
-(* Environment annotated with marks on recursive arguments *)
-
-(* tells whether it is a strict or loose subterm *)
-type size = Large | Strict
-
-(* merging information *)
-let size_glb s1 s2 =
- match s1,s2 with
- Strict, Strict -> Strict
- | _ -> Large
-
-(* possible specifications for a term:
- - Not_subterm: when the size of a term is not related to the
- recursive argument of the fixpoint
- - Subterm: when the term is a subterm of the recursive argument
- the wf_paths argument specifies which subterms are recursive
- - Dead_code: when the term has been built by elimination over an
- empty type
- *)
-
-type subterm_spec =
- Subterm of (size * wf_paths)
- | Dead_code
- | Not_subterm
-
-let eq_wf_paths = Rtree.equal eq_recarg
-
-let inter_recarg r1 r2 = match r1, r2 with
-| Norec, Norec -> Some r1
-| Mrec i1, Mrec i2
-| Imbr i1, Imbr i2
-| Mrec i1, Imbr i2 -> if Names.eq_ind_chk i1 i2 then Some r1 else None
-| Imbr i1, Mrec i2 -> if Names.eq_ind_chk i1 i2 then Some r2 else None
-| _ -> None
-
-let inter_wf_paths = Rtree.inter eq_recarg inter_recarg Norec
-
-let incl_wf_paths = Rtree.incl eq_recarg inter_recarg Norec
-
-let spec_of_tree t =
- if eq_wf_paths t mk_norec
- then Not_subterm
- else Subterm (Strict, t)
-
-let inter_spec s1 s2 =
- match s1, s2 with
- | _, Dead_code -> s1
- | Dead_code, _ -> s2
- | Not_subterm, _ -> s1
- | _, Not_subterm -> s2
- | Subterm (a1,t1), Subterm (a2,t2) ->
- Subterm (size_glb a1 a2, inter_wf_paths t1 t2)
-
-let subterm_spec_glb =
- Array.fold_left inter_spec Dead_code
-
-type guard_env =
- { env : env;
- (* dB of last fixpoint *)
- rel_min : int;
- (* dB of variables denoting subterms *)
- genv : subterm_spec Lazy.t list;
- }
-
-let make_renv env recarg tree =
- { env = env;
- rel_min = recarg+2; (* recarg = 0 ==> Rel 1 -> recarg; Rel 2 -> fix *)
- genv = [Lazy.from_val(Subterm(Large,tree))] }
-
-let push_var renv (x,ty,spec) =
- { env = push_rel (LocalAssum (x,ty)) renv.env;
- rel_min = renv.rel_min+1;
- genv = spec:: renv.genv }
-
-let assign_var_spec renv (i,spec) =
- { renv with genv = List.assign renv.genv (i-1) spec }
-
-let push_var_renv renv (x,ty) =
- push_var renv (x,ty,Lazy.from_val Not_subterm)
-
-(* Fetch recursive information about a variable p *)
-let subterm_var p renv =
- try Lazy.force (List.nth renv.genv (p-1))
- with Failure _ | Invalid_argument _ -> Not_subterm
-
-let push_ctxt_renv renv ctxt =
- let n = rel_context_length ctxt in
- { env = push_rel_context ctxt renv.env;
- rel_min = renv.rel_min+n;
- genv = iterate (fun ge -> Lazy.from_val Not_subterm::ge) n renv.genv }
-
-let push_fix_renv renv (_,v,_ as recdef) =
- let n = Array.length v in
- { env = push_rec_types recdef renv.env;
- rel_min = renv.rel_min+n;
- genv = iterate (fun ge -> Lazy.from_val Not_subterm::ge) n renv.genv }
-
-
-(* Definition and manipulation of the stack *)
-type stack_element = |SClosure of guard_env*constr |SArg of subterm_spec Lazy.t
-
-let push_stack_closures renv l stack =
- List.fold_right (fun h b -> (SClosure (renv,h))::b) l stack
-
-let push_stack_args l stack =
- List.fold_right (fun h b -> (SArg h)::b) l stack
-
-(******************************)
-(* Computing the recursive subterms of a term (propagation of size
- information through Cases). *)
-
-(*
- c is a branch of an inductive definition corresponding to the spec
- lrec. mind_recvec is the recursive spec of the inductive
- definition of the decreasing argument n.
-
- case_branches_specif renv lrec lc will pass the lambdas
- of c corresponding to pattern variables and collect possibly new
- subterms variables and returns the bodies of the branches with the
- correct envs and decreasing args.
-*)
-
-let lookup_subterms env ind =
- let (_,mip) = lookup_mind_specif env ind in
- mip.mind_recargs
-
-let match_inductive ind ra =
- match ra with
- | (Mrec i | Imbr i) -> eq_ind_chk ind i
- | Norec -> false
-
-(* In {match c as z in ci y_s return P with |C_i x_s => t end}
- [branches_specif renv c_spec ci] returns an array of x_s specs knowing
- c_spec. *)
-let branches_specif renv c_spec ci =
- let car =
- (* We fetch the regular tree associated to the inductive of the match.
- This is just to get the number of constructors (and constructor
- arities) that fit the match branches without forcing c_spec.
- Note that c_spec might be more precise than [v] below, because of
- nested inductive types. *)
- let (_,mip) = lookup_mind_specif renv.env ci.ci_ind in
- let v = dest_subterms mip.mind_recargs in
- Array.map List.length v in
- Array.mapi
- (fun i nca -> (* i+1-th cstructor has arity nca *)
- let lvra = lazy
- (match Lazy.force c_spec with
- Subterm (_,t) when match_inductive ci.ci_ind (dest_recarg t) ->
- let vra = Array.of_list (dest_subterms t).(i) in
- assert (nca = Array.length vra);
- Array.map spec_of_tree vra
- | Dead_code -> Array.make nca Dead_code
- | _ -> Array.make nca Not_subterm) in
- List.init nca (fun j -> lazy (Lazy.force lvra).(j)))
- car
-
-let check_inductive_codomain env p =
- let absctx, ar = dest_lam_assum env p in
- let env = push_rel_context absctx env in
- let arctx, s = dest_prod_assum env ar in
- let env = push_rel_context arctx env in
- let i,l' = decompose_app (whd_all env s) in
- match i with Ind _ -> true | _ -> false
-
-(* The following functions are almost duplicated from indtypes.ml, except
-that they carry here a poorer environment (containing less information). *)
-let ienv_push_var (env, lra) (x,a,ra) =
-(push_rel (LocalAssum (x,a)) env, (Norec,ra)::lra)
-
-let ienv_push_inductive (env, ra_env) ((mind,u),lpar) =
- let mib = Environ.lookup_mind mind env in
- let ntypes = mib.mind_ntypes in
- let push_ind specif env =
- let decl = LocalAssum (Anonymous,
- hnf_prod_applist env (type_of_inductive env ((mib,specif),u)) lpar) in
- push_rel decl env
- in
- let env = Array.fold_right push_ind mib.mind_packets env in
- let rc = Array.mapi (fun j t -> (Imbr (mind,j),t)) (Rtree.mk_rec_calls ntypes) in
- let lra_ind = Array.rev_to_list rc in
- let ra_env = List.map (fun (r,t) -> (r,Rtree.lift ntypes t)) ra_env in
- (env, lra_ind @ ra_env)
-
-let rec ienv_decompose_prod (env,_ as ienv) n c =
- if Int.equal n 0 then (ienv,c) else
- let c' = whd_all env c in
- match c' with
- Prod(na,a,b) ->
- let ienv' = ienv_push_var ienv (na,a,mk_norec) in
- ienv_decompose_prod ienv' (n-1) b
- | _ -> assert false
-
-let lambda_implicit_lift n a =
- let level = Level.make (DirPath.make [Id.of_string "implicit"]) 0 in
- let implicit_sort = Sort (Type (Universe.make level)) in
- let lambda_implicit a = Lambda (Anonymous, implicit_sort, a) in
- iterate lambda_implicit n (lift n a)
-
-let abstract_mind_lc ntyps npars lc =
- if Int.equal npars 0 then
- lc
- else
- let make_abs =
- List.init ntyps
- (function i -> lambda_implicit_lift npars (Rel (i+1)))
- in
- Array.map (substl make_abs) lc
-
-(* [get_recargs_approx env tree ind args] builds an approximation of the recargs
-tree for ind, knowing args. The argument tree is used to know when candidate
-nested types should be traversed, pruning the tree otherwise. This code is very
-close to check_positive in indtypes.ml, but does no positivy check and does not
-compute the number of recursive arguments. *)
-let get_recargs_approx env tree ind args =
- let rec build_recargs (env, ra_env as ienv) tree c =
- let x,largs = decompose_app (whd_all env c) in
- match x with
- | Prod (na,b,d) ->
- assert (List.is_empty largs);
- build_recargs (ienv_push_var ienv (na, b, mk_norec)) tree d
- | Rel k ->
- (* Free variables are allowed and assigned Norec *)
- (try snd (List.nth ra_env (k-1))
- with Failure _ | Invalid_argument _ -> mk_norec)
- | Ind ind_kn ->
- (* When the inferred tree allows it, we consider that we have a potential
- nested inductive type *)
- begin match dest_recarg tree with
- | Imbr kn' | Mrec kn' when eq_ind_chk (fst ind_kn) kn' ->
- build_recargs_nested ienv tree (ind_kn, largs)
- | _ -> mk_norec
- end
- | err ->
- mk_norec
-
- and build_recargs_nested (env,ra_env as ienv) tree (((mind,i),u), largs) =
- (* If the infered tree already disallows recursion, no need to go further *)
- if eq_wf_paths tree mk_norec then tree
- else
- let mib = Environ.lookup_mind mind env in
- let auxnpar = mib.mind_nparams_rec in
- let nonrecpar = mib.mind_nparams - auxnpar in
- let (lpar,_) = List.chop auxnpar largs in
- let auxntyp = mib.mind_ntypes in
- (* Extends the environment with a variable corresponding to
- the inductive def *)
- let (env',_ as ienv') = ienv_push_inductive ienv ((mind,u),lpar) in
- (* Parameters expressed in env' *)
- let lpar' = List.map (lift auxntyp) lpar in
- (* In case of mutual inductive types, we use the recargs tree which was
- computed statically. This is fine because nested inductive types with
- mutually recursive containers are not supported. *)
- let trees =
- if Int.equal auxntyp 1 then [|dest_subterms tree|]
- else Array.map (fun mip -> dest_subterms mip.mind_recargs) mib.mind_packets
- in
- let mk_irecargs j specif =
- (* The nested inductive type with parameters removed *)
- let auxlcvect = abstract_mind_lc auxntyp auxnpar specif.mind_nf_lc in
- let paths = Array.mapi
- (fun k c ->
- let c' = hnf_prod_applist env' c lpar' in
- (* skip non-recursive parameters *)
- let (ienv',c') = ienv_decompose_prod ienv' nonrecpar c' in
- build_recargs_constructors ienv' trees.(j).(k) c')
- auxlcvect
- in
- mk_paths (Imbr (mind,j)) paths
- in
- let irecargs = Array.mapi mk_irecargs mib.mind_packets in
- (Rtree.mk_rec irecargs).(i)
-
- and build_recargs_constructors ienv trees c =
- let rec recargs_constr_rec (env,ra_env as ienv) trees lrec c =
- let x,largs = decompose_app (whd_all env c) in
- match x with
-
- | Prod (na,b,d) ->
- let () = assert (List.is_empty largs) in
- let recarg = build_recargs ienv (List.hd trees) b in
- let ienv' = ienv_push_var ienv (na,b,mk_norec) in
- recargs_constr_rec ienv' (List.tl trees) (recarg::lrec) d
- | hd ->
- List.rev lrec
- in
- recargs_constr_rec ienv trees [] c
- in
- (* starting with ra_env = [] seems safe because any unbounded Rel will be
- assigned Norec *)
- build_recargs_nested (env,[]) tree (ind, args)
-
-(* [restrict_spec env spec p] restricts the size information in spec to what is
- allowed to flow through a match with predicate p in environment env. *)
-let restrict_spec env spec p =
- if spec = Not_subterm then spec
- else 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. *)
- if noccur_with_meta 1 (rel_context_length absctx) ar then spec
- else
- let env = push_rel_context absctx env in
- let arctx, s = dest_prod_assum env ar in
- let env = push_rel_context arctx env in
- let i,args = decompose_app (whd_all env s) in
- match i with
- | Ind i ->
- begin match spec with
- | Dead_code -> spec
- | Subterm(st,tree) ->
- let recargs = get_recargs_approx env tree i args in
- let recargs = inter_wf_paths tree recargs in
- Subterm(st,recargs)
- | _ -> assert false
- end
- | _ -> Not_subterm
-
-(* [subterm_specif renv t] computes the recursive structure of [t] and
- compare its size with the size of the initial recursive argument of
- the fixpoint we are checking. [renv] collects such information
- about variables.
-*)
-
-
-let rec subterm_specif renv stack t =
- (* maybe reduction is not always necessary! *)
- let f,l = decompose_app (whd_all renv.env t) in
- match f with
- | Rel k -> subterm_var k renv
-
- | Case (ci,p,c,lbr) ->
- let stack' = push_stack_closures renv l stack in
- let cases_spec =
- branches_specif renv (lazy_subterm_specif renv [] c) ci
- in
- let stl =
- Array.mapi (fun i br' ->
- let stack_br = push_stack_args (cases_spec.(i)) stack' in
- subterm_specif renv stack_br br')
- lbr in
- let spec = subterm_spec_glb stl in
- restrict_spec renv.env spec p
-
- | Fix ((recindxs,i),(_,typarray,bodies as recdef)) ->
- (* when proving that the fixpoint f(x)=e is less than n, it is enough
- to prove that e is less than n assuming f is less than n
- furthermore when f is applied to a term which is strictly less than
- n, one may assume that x itself is strictly less than n
- *)
- if not (check_inductive_codomain renv.env typarray.(i)) then Not_subterm
- else
- let (ctxt,clfix) = dest_prod renv.env typarray.(i) in
- let oind =
- let env' = push_rel_context ctxt renv.env in
- try Some(fst(find_inductive env' clfix))
- with Not_found -> None in
- (match oind with
- None -> Not_subterm (* happens if fix is polymorphic *)
- | Some ind ->
- let nbfix = Array.length typarray in
- let recargs = lookup_subterms renv.env ind in
- (* pushing the fixpoints *)
- let renv' = push_fix_renv renv recdef in
- let renv' =
- (* Why Strict here ? To be general, it could also be
- Large... *)
- assign_var_spec renv'
- (nbfix-i, lazy (Subterm(Strict,recargs))) in
- let decrArg = recindxs.(i) in
- let theBody = bodies.(i) in
- let nbOfAbst = decrArg+1 in
- let sign,strippedBody = decompose_lam_n_assum nbOfAbst theBody in
- (* pushing the fix parameters *)
- let stack' = push_stack_closures renv l stack in
- let renv'' = push_ctxt_renv renv' sign in
- let renv'' =
- if List.length stack' < nbOfAbst then renv''
- else
- let decrArg = List.nth stack' decrArg in
- let arg_spec = stack_element_specif decrArg in
- assign_var_spec renv'' (1, arg_spec) in
- subterm_specif renv'' [] strippedBody)
-
- | Lambda (x,a,b) ->
- assert (l=[]);
- let spec,stack' = extract_stack stack in
- subterm_specif (push_var renv (x,a,spec)) stack' b
-
- (* Metas and evars are considered OK *)
- | (Meta _|Evar _) -> Dead_code
-
- | Proj (p, c) ->
- let subt = subterm_specif renv stack c in
- (match subt with
- | Subterm (_s, wf) ->
- (* We take the subterm specs of the constructor of the record *)
- let wf_args = (dest_subterms wf).(0) in
- (* We extract the tree of the projected argument *)
- let n = Projection.arg p in
- spec_of_tree (List.nth wf_args n)
- | Dead_code -> Dead_code
- | Not_subterm -> Not_subterm)
-
- (* Other terms are not subterms *)
- | Var _ | Sort _ | Cast _ | Prod _ | LetIn _ | App _ | Const _ | Ind _
- | Construct _ | CoFix _ -> Not_subterm
-
-and lazy_subterm_specif renv stack t =
- lazy (subterm_specif renv stack t)
-
-and stack_element_specif = function
- |SClosure (h_renv,h) -> lazy_subterm_specif h_renv [] h
- |SArg x -> x
-
-and extract_stack = function
- | [] -> Lazy.from_val Not_subterm , []
- | h::t -> stack_element_specif h, t
-
-
-(* Check size x is a correct size for recursive calls. *)
-let check_is_subterm x tree =
- match Lazy.force x with
- | Subterm (Strict,tree') -> incl_wf_paths tree tree'
- | Dead_code -> true
- | _ -> false
-
-(************************************************************************)
-
-exception FixGuardError of env * guard_error
-
-let error_illegal_rec_call renv fx (arg_renv,arg) =
- let (_,le_vars,lt_vars) =
- List.fold_left
- (fun (i,le,lt) sbt ->
- match Lazy.force sbt with
- (Subterm(Strict,_) | Dead_code) -> (i+1, le, i::lt)
- | (Subterm(Large,_)) -> (i+1, i::le, lt)
- | _ -> (i+1, le ,lt))
- (1,[],[]) renv.genv in
- raise (FixGuardError (renv.env,
- RecursionOnIllegalTerm(fx,(arg_renv.env, arg),
- le_vars,lt_vars)))
-
-let error_partial_apply renv fx =
- raise (FixGuardError (renv.env,NotEnoughArgumentsForFixCall fx))
-
-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. *)
- if noccur_with_meta 1 (rel_context_length absctx) ar then stack
- else let env = push_rel_context absctx env in
- let rec filter_stack env ar stack =
- let t = whd_all env ar in
- match stack, t with
- | elt :: stack', Prod (n,a,c0) ->
- let d = LocalAssum (n,a) in
- let ctx, a = dest_prod_assum env a in
- let env = push_rel_context ctx env in
- let ty, args = decompose_app (whd_all env a) in
- let elt = match ty with
- | Ind ind ->
- let spec' = stack_element_specif elt in
- (match (Lazy.force spec') with
- | Not_subterm | Dead_code -> elt
- | Subterm(s,path) ->
- let recargs = get_recargs_approx env path ind args in
- let path = inter_wf_paths path recargs in
- SArg (lazy (Subterm(s,path))))
- | _ -> (SArg (lazy Not_subterm))
- in
- elt :: filter_stack (push_rel d env) c0 stack'
- | _,_ -> List.fold_right (fun _ l -> SArg (lazy Not_subterm) :: l) stack []
- in
- filter_stack env ar stack
-
-(* Check if [def] is a guarded fixpoint body with decreasing arg.
- given [recpos], the decreasing arguments of each mutually defined
- fixpoint. *)
-let check_one_fix renv recpos trees def =
- let nfi = Array.length recpos in
-
- (* Checks if [t] only make valid recursive calls *)
- let rec check_rec_call renv stack t =
- (* if [t] does not make recursive calls, it is guarded: *)
- if noccur_with_meta renv.rel_min nfi t then ()
- else
- let (f,l) = decompose_app (whd_betaiotazeta t) in
- match f with
- | Rel p ->
- (* Test if [p] is a fixpoint (recursive call) *)
- if renv.rel_min <= p && p < renv.rel_min+nfi then
- begin
- List.iter (check_rec_call renv []) l;
- (* the position of the invoked fixpoint: *)
- let glob = renv.rel_min+nfi-1-p in
- (* the decreasing arg of the rec call: *)
- let np = recpos.(glob) in
- let stack' = push_stack_closures renv l stack in
- if List.length stack' <= np then error_partial_apply renv glob
- else
- (* Retrieve the expected tree for the argument *)
- (* Check the decreasing arg is smaller *)
- let z = List.nth stack' np in
- if not (check_is_subterm (stack_element_specif z) trees.(glob)) then
- begin match z with
- |SClosure (z,z') -> error_illegal_rec_call renv glob (z,z')
- |SArg _ -> error_partial_apply renv glob
- end
- end
- else
- begin
- match lookup_rel p renv.env with
- | LocalAssum _ ->
- List.iter (check_rec_call renv []) l
- | LocalDef (_,c,_) ->
- try List.iter (check_rec_call renv []) l
- with FixGuardError _ ->
- check_rec_call renv stack (applist(lift p c,l))
- end
-
- | Case (ci,p,c_0,lrest) ->
- List.iter (check_rec_call renv []) (c_0::p::l);
- (* compute the recarg information for the arguments of
- each branch *)
- 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 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
-
- (* Enables to traverse Fixpoint definitions in a more intelligent
- way, ie, the rule :
- if - g = fix g (y1:T1)...(yp:Tp) {struct yp} := e &
- - f is guarded with respect to the set of pattern variables S
- in a1 ... am &
- - f is guarded with respect to the set of pattern variables S
- in T1 ... Tp &
- - ap is a sub-term of the formal argument of f &
- - f is guarded with respect to the set of pattern variables
- S+{yp} in e
- then f is guarded with respect to S in (g a1 ... am).
- Eduardo 7/9/98 *)
- | Fix ((recindxs,i),(_,typarray,bodies as recdef)) ->
- List.iter (check_rec_call renv []) l;
- Array.iter (check_rec_call renv []) typarray;
- let decrArg = recindxs.(i) in
- let renv' = push_fix_renv renv recdef in
- let stack' = push_stack_closures renv l stack in
- Array.iteri
- (fun j body ->
- if i=j && (List.length stack' > decrArg) then
- let recArg = List.nth stack' decrArg in
- let arg_sp = stack_element_specif recArg in
- check_nested_fix_body renv' (decrArg+1) arg_sp body
- else check_rec_call renv' [] body)
- bodies
-
- | Const (kn,u) ->
- if evaluable_constant kn renv.env then
- try List.iter (check_rec_call renv []) l
- with (FixGuardError _ ) ->
- let value = (applist(constant_value renv.env (kn,u), l)) in
- check_rec_call renv stack value
- else List.iter (check_rec_call renv []) l
-
- | Lambda (x,a,b) ->
- assert (l = []);
- check_rec_call renv [] a ;
- let spec, stack' = extract_stack stack in
- check_rec_call (push_var renv (x,a,spec)) stack' b
-
- | Prod (x,a,b) ->
- assert (l = [] && stack = []);
- check_rec_call renv [] a;
- check_rec_call (push_var_renv renv (x,a)) [] b
-
- | CoFix (i,(_,typarray,bodies as recdef)) ->
- List.iter (check_rec_call renv []) l;
- Array.iter (check_rec_call renv []) typarray;
- let renv' = push_fix_renv renv recdef in
- Array.iter (check_rec_call renv' []) bodies
-
- | (Ind _ | Construct _) ->
- List.iter (check_rec_call renv []) l
-
- | Proj (p, c) ->
- List.iter (check_rec_call renv []) l;
- check_rec_call renv [] c
-
- | Var _ -> anomaly (Pp.str "Section variable in Coqchk!")
-
- | Sort _ -> assert (l = [])
-
- (* l is not checked because it is considered as the meta's context *)
- | (Evar _ | Meta _) -> ()
-
- | (App _ | LetIn _ | Cast _) -> assert false (* beta zeta reduction *)
-
- and check_nested_fix_body renv decr recArgsDecrArg body =
- if decr = 0 then
- check_rec_call (assign_var_spec renv (1,recArgsDecrArg)) [] body
- else
- match body with
- | Lambda (x,a,b) ->
- check_rec_call renv [] a;
- let renv' = push_var_renv renv (x,a) in
- check_nested_fix_body renv' (decr-1) recArgsDecrArg b
- | _ -> anomaly (Pp.str "Not enough abstractions in fix body.")
-
- in
- check_rec_call renv [] def
-
-
-let inductive_of_mutfix env ((nvect,bodynum),(names,types,bodies as recdef)) =
- let nbfix = Array.length bodies in
- if nbfix = 0
- || Array.length nvect <> nbfix
- || Array.length types <> nbfix
- || Array.length names <> nbfix
- || bodynum < 0
- || bodynum >= nbfix
- then anomaly (Pp.str "Ill-formed fix term.");
- let fixenv = push_rec_types recdef env in
- let raise_err env i err =
- error_ill_formed_rec_body env err names i in
- (* Check the i-th definition with recarg k *)
- let find_ind i k def =
- (* check fi does not appear in the k+1 first abstractions,
- gives the type of the k+1-eme abstraction (must be an inductive) *)
- let rec check_occur env n def =
- match (whd_all env def) with
- | Lambda (x,a,b) ->
- if noccur_with_meta n nbfix a then
- let env' = push_rel (LocalAssum (x,a)) env in
- if n = k+1 then
- (* get the inductive type of the fixpoint *)
- let (mind, _) =
- try find_inductive env a
- with Not_found ->
- raise_err env i (RecursionNotOnInductiveType a) in
- (mind, (env', b))
- else check_occur env' (n+1) b
- else anomaly ~label:"check_one_fix" (Pp.str "Bad occurrence of recursive call.")
- | _ -> raise_err env i NotEnoughAbstractionInFixBody in
- check_occur fixenv 1 def in
- (* Do it on every fixpoint *)
- let rv = Array.map2_i find_ind nvect bodies in
- (Array.map fst rv, Array.map snd rv)
-
-
-let check_fix env ((nvect,_),(names,_,bodies as _recdef) as fix) =
- let (minds, rdef) = inductive_of_mutfix env fix in
- let get_tree (kn,i) =
- let mib = Environ.lookup_mind kn env in
- mib.mind_packets.(i).mind_recargs
- in
- let trees = Array.map get_tree minds in
- for i = 0 to Array.length bodies - 1 do
- let (fenv,body) = rdef.(i) in
- let renv = make_renv fenv nvect.(i) trees.(i) in
- try check_one_fix renv nvect trees body
- with FixGuardError (fixenv,err) ->
- error_ill_formed_rec_body fixenv err names i
- done
-
-(*
-let cfkey = CProfile.declare_profile "check_fix";;
-let check_fix env fix = CProfile.profile3 cfkey check_fix env fix;;
-*)
-
-(************************************************************************)
-(* Co-fixpoints. *)
-
-exception CoFixGuardError of env * guard_error
-
-let anomaly_ill_typed () =
- anomaly ~label:"check_one_cofix" (Pp.str "too many arguments applied to constructor.")
-
-let rec codomain_is_coind env c =
- let b = whd_all env c in
- match b with
- | Prod (x,a,b) ->
- codomain_is_coind (push_rel (LocalAssum (x,a)) env) b
- | _ ->
- (try find_coinductive env b
- with Not_found ->
- raise (CoFixGuardError (env, CodomainNotInductiveType b)))
-
-let check_one_cofix env nbfix def deftype =
- let rec check_rec_call env alreadygrd n tree vlra t =
- if not (noccur_with_meta n nbfix t) then
- let c,args = decompose_app (whd_all env t) in
- match c with
- | Rel p when n <= p && p < n+nbfix ->
- (* recursive call: must be guarded and no nested recursive
- call allowed *)
- if not alreadygrd then
- raise (CoFixGuardError (env,UnguardedRecursiveCall t))
- else if not(List.for_all (noccur_with_meta n nbfix) args) then
- raise (CoFixGuardError (env,NestedRecursiveOccurrences))
- | Construct ((_,i as cstr_kn),u) ->
- let lra = vlra.(i-1) in
- let mI = inductive_of_constructor cstr_kn in
- let (mib,mip) = lookup_mind_specif env mI in
- let realargs = List.skipn mib.mind_nparams args in
- let rec process_args_of_constr = function
- | (t::lr), (rar::lrar) ->
- if rar = mk_norec then
- if noccur_with_meta n nbfix t
- then process_args_of_constr (lr, lrar)
- else raise (CoFixGuardError
- (env,RecCallInNonRecArgOfConstructor t))
- else begin
- check_rec_call env true n rar (dest_subterms rar) t;
- process_args_of_constr (lr, lrar)
- end
- | [],_ -> ()
- | _ -> anomaly_ill_typed ()
- in process_args_of_constr (realargs, lra)
-
- | Lambda (x,a,b) ->
- assert (args = []);
- if noccur_with_meta n nbfix a then
- let env' = push_rel (LocalAssum (x,a)) env in
- check_rec_call env' alreadygrd (n+1) tree vlra b
- else
- raise (CoFixGuardError (env,RecCallInTypeOfAbstraction a))
-
- | CoFix (j,(_,varit,vdefs as recdef)) ->
- if List.for_all (noccur_with_meta n nbfix) args
- then
- if Array.for_all (noccur_with_meta n nbfix) varit then
- let nbfix = Array.length vdefs in
- let env' = push_rec_types recdef env in
- (Array.iter (check_rec_call env' alreadygrd (n+nbfix) tree vlra) vdefs;
- List.iter (check_rec_call env alreadygrd n tree vlra) args)
- else
- raise (CoFixGuardError (env,RecCallInTypeOfDef c))
- else
- raise (CoFixGuardError (env,UnguardedRecursiveCall c))
-
- | Case (_,p,tm,vrest) ->
- begin
- let tree = match restrict_spec env (Subterm (Strict, tree)) p with
- | Dead_code -> assert false
- | Subterm (_, tree') -> tree'
- | _ -> raise (CoFixGuardError (env, ReturnPredicateNotCoInductive c))
- in
- if (noccur_with_meta n nbfix p) then
- if (noccur_with_meta n nbfix tm) then
- if (List.for_all (noccur_with_meta n nbfix) args) then
- let vlra = dest_subterms tree in
- Array.iter (check_rec_call env alreadygrd n tree vlra) vrest
- else
- raise (CoFixGuardError (env,RecCallInCaseFun c))
- else
- raise (CoFixGuardError (env,RecCallInCaseArg c))
- else
- raise (CoFixGuardError (env,RecCallInCasePred c))
- end
-
- | Meta _ -> ()
- | Evar _ ->
- List.iter (check_rec_call env alreadygrd n tree vlra) args
-
- | _ -> raise (CoFixGuardError (env,NotGuardedForm t)) in
-
- let (mind, _) = codomain_is_coind env deftype in
- let vlra = lookup_subterms env mind in
- check_rec_call env false 1 vlra (dest_subterms vlra) def
-
-(* The function which checks that the whole block of definitions
- satisfies the guarded condition *)
-
-let check_cofix env (bodynum,(names,types,bodies as recdef)) =
- let nbfix = Array.length bodies in
- for i = 0 to nbfix-1 do
- let fixenv = push_rec_types recdef env in
- try check_one_cofix fixenv nbfix bodies.(i) types.(i)
- with CoFixGuardError (errenv,err) ->
- error_ill_formed_rec_body errenv err names i
- done
diff --git a/checker/inductive.mli b/checker/inductive.mli
deleted file mode 100644
index 0ca0d14a25..0000000000
--- a/checker/inductive.mli
+++ /dev/null
@@ -1,85 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(*i*)
-open Names
-open Cic
-open Environ
-(*i*)
-
-(*s Extracting an inductive type from a construction *)
-
-val find_rectype : env -> constr -> pinductive * constr list
-
-type mind_specif = mutual_inductive_body * one_inductive_body
-
-(*s Fetching information in the environment about an inductive type.
- Raises [Not_found] if the inductive type is not found. *)
-val lookup_mind_specif : env -> inductive -> mind_specif
-
-val inductive_is_polymorphic : mutual_inductive_body -> bool
-
-val inductive_is_cumulative : mutual_inductive_body -> bool
-
-val type_of_inductive : env -> mind_specif puniverses -> constr
-
-(* Return type as quoted by the user *)
-val type_of_constructor : pconstructor -> mind_specif -> constr
-
-val arities_of_specif : MutInd.t puniverses -> mind_specif -> constr array
-
-(* [type_case_branches env (I,args) (p:A) c] computes useful types
- about the following Cases expression:
- <p>Cases (c :: (I args)) of b1..bn end
- It computes the type of every branch (pattern variables are
- introduced by products) and the type for the whole expression.
- *)
-val type_case_branches :
- env -> pinductive * constr list -> constr * constr -> constr
- -> constr array * constr
-
-(* Check a [case_info] actually correspond to a Case expression on the
- given inductive type. *)
-val check_case_info : env -> inductive -> case_info -> unit
-
-(*s Guard conditions for fix and cofix-points. *)
-val check_fix : env -> fixpoint -> unit
-val check_cofix : env -> cofixpoint -> unit
-
-(*s Support for sort-polymorphic inductive types *)
-
-val type_of_inductive_knowing_parameters :
- env -> mind_specif puniverses -> constr array -> constr
-
-val max_inductive_sort : sorts array -> Univ.universe
-
-val instantiate_universes : env -> rel_context ->
- template_arity -> constr array -> rel_context * sorts
-
-(***************************************************************)
-(* Debug *)
-
-type size = Large | Strict
-type subterm_spec =
- Subterm of (size * wf_paths)
- | Dead_code
- | Not_subterm
-type guard_env =
- { env : env;
- (* dB of last fixpoint *)
- rel_min : int;
- (* dB of variables denoting subterms *)
- genv : subterm_spec Lazy.t list;
- }
-
-type stack_element = |SClosure of guard_env*constr |SArg of subterm_spec Lazy.t
-val subterm_specif : guard_env -> stack_element list -> constr -> subterm_spec
-val branches_specif : guard_env -> subterm_spec Lazy.t -> case_info ->
- subterm_spec Lazy.t list array
diff --git a/checker/mod_checking.ml b/checker/mod_checking.ml
index 6b2af71f33..ed617d73c2 100644
--- a/checker/mod_checking.ml
+++ b/checker/mod_checking.ml
@@ -1,11 +1,8 @@
open Pp
open Util
open Names
-open Cic
open Reduction
open Typeops
-open Indtypes
-open Modops
open Subtyping
open Declarations
open Environ
@@ -13,29 +10,29 @@ open Environ
(** {6 Checking constants } *)
let check_constant_declaration env kn cb =
- Flags.if_verbose Feedback.msg_notice (str " checking cst:" ++ prcon kn);
+ Flags.if_verbose Feedback.msg_notice (str " checking cst:" ++ Constant.print kn);
(** Locally set the oracle for further typechecking *)
- let oracle = env.env_conv_oracle in
+ let oracle = env.env_typing_flags.conv_oracle in
let env = Environ.set_oracle env cb.const_typing_flags.conv_oracle in
(** [env'] contains De Bruijn universe variables *)
- let env' =
+ let poly, env' =
match cb.const_universes with
- | Monomorphic_const ctx -> push_context_set ~strict:true ctx env
+ | Monomorphic_const ctx -> false, push_context_set ~strict:true ctx env
| Polymorphic_const auctx ->
let ctx = Univ.AUContext.repr auctx in
- push_context ~strict:false ctx env
+ true, push_context ~strict:false ctx env
in
let ty = cb.const_type in
let _ = infer_type env' ty in
let () =
- match body_of_constant cb with
+ match Environ.body_of_constant_body env cb with
| Some bd ->
- let j = infer env' bd in
- conv_leq env' j ty
+ let j = infer env' (fst bd) in
+ conv_leq env' j.uj_type ty
| None -> ()
in
let env =
- if constant_is_polymorphic cb then add_constant kn cb env
+ if poly then add_constant kn cb env
else add_constant kn cb env'
in
(** Reset the value of the oracle *)
@@ -63,6 +60,7 @@ let mk_mtb mp sign delta =
mod_retroknowledge = ModTypeRK; }
let rec check_module env mp mb =
+ Flags.if_verbose Feedback.msg_notice (str " checking module: " ++ str (ModPath.to_string mp));
let (_:module_signature) =
check_signature env mb.mod_type mb.mod_mp mb.mod_delta
in
@@ -76,10 +74,13 @@ let rec check_module env mp mb =
|Some sign ->
let mtb1 = mk_mtb mp sign mb.mod_delta
and mtb2 = mk_mtb mp mb.mod_type mb.mod_delta in
- let env = add_module_type mp mtb1 env in
- Subtyping.check_subtypes env mtb1 mtb2
+ let env = Modops.add_module_type mp mtb1 env in
+ let cu = check_subtypes env mtb1 mtb2 in
+ if not (Environ.check_constraints cu env) then
+ CErrors.user_err Pp.(str "Incorrect universe constraints for module subtyping");
and check_module_type env mty =
+ Flags.if_verbose Feedback.msg_notice (str " checking module type: " ++ str (ModPath.to_string mty.mod_mp));
let (_:module_signature) =
check_signature env mty.mod_type mty.mod_mp mty.mod_delta in
()
@@ -89,32 +90,35 @@ and check_structure_field env mp lab res = function
let c = Constant.make2 mp lab in
check_constant_declaration env c cb
| SFBmind mib ->
- let kn = MutInd.make2 mp lab in
- let kn = mind_of_delta res kn in
- Indtypes.check_inductive env kn mib
+ let kn = KerName.make mp lab in
+ let kn = Mod_subst.mind_of_delta_kn res kn in
+ CheckInductive.check_inductive env kn mib
| SFBmodule msb ->
let () = check_module env (MPdot(mp,lab)) msb in
Modops.add_module msb env
| SFBmodtype mty ->
check_module_type env mty;
- add_modtype (MPdot(mp,lab)) mty env
+ add_modtype mty env
and check_mexpr env mse mp_mse res = match mse with
| MEident mp ->
let mb = lookup_module mp env in
- (subst_and_strengthen mb mp_mse).mod_type
+ (Modops.strengthen_and_subst_mb mb mp_mse false).mod_type
| MEapply (f,mp) ->
let sign = check_mexpr env f mp_mse res in
- let farg_id, farg_b, fbody_b = destr_functor sign in
- let mtb = module_type_of_module (Some mp) (lookup_module mp env) in
- check_subtypes env mtb farg_b;
- subst_signature (map_mbid farg_id mp) fbody_b
- | MEwith _ -> error_with_module ()
+ let farg_id, farg_b, fbody_b = Modops.destr_functor sign in
+ let mtb = Modops.module_type_of_module (lookup_module mp env) in
+ let cu = check_subtypes env mtb farg_b in
+ if not (Environ.check_constraints cu env) then
+ CErrors.user_err Pp.(str "Incorrect universe constraints for module subtyping");
+ Modops.subst_signature (Mod_subst.map_mbid farg_id mp Mod_subst.empty_delta_resolver) fbody_b
+ | MEwith _ -> CErrors.user_err Pp.(str "Unsupported 'with' constraint in module implementation")
+
and check_mexpression env sign mp_mse res = match sign with
| MoreFunctor (arg_id, mtb, body) ->
check_module_type env mtb;
- let env' = add_module_type (MPbound arg_id) mtb env in
+ let env' = Modops.add_module_type (MPbound arg_id) mtb env in
let body = check_mexpression env' body mp_mse res in
MoreFunctor(arg_id,mtb,body)
| NoFunctor me -> check_mexpr env me mp_mse res
@@ -122,7 +126,7 @@ and check_mexpression env sign mp_mse res = match sign with
and check_signature env sign mp_mse res = match sign with
| MoreFunctor (arg_id, mtb, body) ->
check_module_type env mtb;
- let env' = add_module_type (MPbound arg_id) mtb env in
+ let env' = Modops.add_module_type (MPbound arg_id) mtb env in
let body = check_signature env' body mp_mse res in
MoreFunctor(arg_id,mtb,body)
| NoFunctor struc ->
diff --git a/checker/mod_checking.mli b/checker/mod_checking.mli
index c9e7f9a1aa..6cff3e6b8c 100644
--- a/checker/mod_checking.mli
+++ b/checker/mod_checking.mli
@@ -8,4 +8,4 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-val check_module : Environ.env -> Names.ModPath.t -> Cic.module_body -> unit
+val check_module : Environ.env -> Names.ModPath.t -> Declarations.module_body -> unit
diff --git a/checker/modops.ml b/checker/modops.ml
deleted file mode 100644
index 541d009ff9..0000000000
--- a/checker/modops.ml
+++ /dev/null
@@ -1,151 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(*i*)
-open CErrors
-open Util
-open Pp
-open Names
-open Cic
-open Declarations
-(*i*)
-
-let error_not_a_constant l =
- user_err Pp.(str ("\""^(Label.to_string l)^"\" is not a constant"))
-
-let error_not_a_functor () = user_err Pp.(str "Application of not a functor")
-
-let error_incompatible_modtypes _ _ = user_err Pp.(str "Incompatible module types")
-
-let error_not_match l _ =
- user_err Pp.(str ("Signature components for label "^Label.to_string l^" do not match"))
-
-let error_no_such_label l = user_err Pp.(str ("No such label "^Label.to_string l))
-
-let error_no_such_label_sub l l1 =
- let l1 = ModPath.to_string l1 in
- user_err Pp.(str ("The field "^
- Label.to_string l^" is missing in "^l1^"."))
-
-let error_not_a_module_loc ?loc s =
- user_err ?loc (str ("\""^Label.to_string s^"\" is not a module"))
-
-let error_not_a_module s = error_not_a_module_loc s
-
-let error_with_module () =
- user_err Pp.(str "Unsupported 'with' constraint in module implementation")
-
-let is_functor = function
- | MoreFunctor _ -> true
- | NoFunctor _ -> false
-
-let destr_functor = function
- | MoreFunctor (arg_id,arg_t,body_t) -> (arg_id,arg_t,body_t)
- | NoFunctor _ -> error_not_a_functor ()
-
-let module_body_of_type mp mtb =
- { mtb with mod_mp = mp; mod_expr = Abstract; mod_retroknowledge = ModBodyRK [] }
-
-let rec add_structure mp sign resolver env =
- let add_one env (l,elem) =
- let kn = KerName.make mp l in
- let con = Constant.make1 kn in
- let mind = mind_of_delta resolver (MutInd.make1 kn) in
- match elem with
- | SFBconst cb ->
- (* let con = constant_of_delta resolver con in*)
- Environ.add_constant con cb env
- | SFBmind mib ->
- (* let mind = mind_of_delta resolver mind in*)
- Environ.add_mind mind mib env
- | SFBmodule mb -> add_module mb env
- (* adds components as well *)
- | SFBmodtype mtb -> Environ.add_modtype mtb.mod_mp mtb env
- in
- List.fold_left add_one env sign
-
-and add_module mb env =
- let mp = mb.mod_mp in
- let env = Environ.shallow_add_module mp mb env in
- match mb.mod_type with
- | NoFunctor struc -> add_structure mp struc mb.mod_delta env
- | MoreFunctor _ -> env
-
-let add_module_type mp mtb env = add_module (module_body_of_type mp mtb) env
-
-let strengthen_const mp_from l cb =
- match cb.const_body with
- | Def _ -> cb
- | _ ->
- let con = Constant.make2 mp_from l in
- let u =
- match cb.const_universes with
- | Monomorphic_const _ -> Univ.Instance.empty
- | Polymorphic_const auctx -> Univ.make_abstract_instance auctx
- in
- { cb with
- const_body = Def (Declarations.from_val (Const (con,u))) }
-
-let rec strengthen_mod mp_from mp_to mb =
- if Declarations.mp_in_delta mb.mod_mp mb.mod_delta then mb
- else
- let mk_expr mp_to = Algebraic (NoFunctor (MEident mp_to)) in
- strengthen_body mk_expr mp_from mp_to mb
-
-and strengthen_body : 'a. (_ -> 'a) -> _ -> _ -> 'a generic_module_body -> 'a generic_module_body =
- fun mk_expr mp_from mp_to mb ->
- match mb.mod_type with
- | MoreFunctor _ -> mb
- | NoFunctor sign ->
- let resolve_out,sign_out = strengthen_sig mp_from sign mp_to
- in
- { mb with
- mod_expr = mk_expr mp_to;
- mod_type = NoFunctor sign_out;
- mod_delta = resolve_out }
-
-and strengthen_sig mp_from sign mp_to =
- match sign with
- | [] -> empty_delta_resolver,[]
- | (l,SFBconst cb) :: rest ->
- let item' = l,SFBconst (strengthen_const mp_from l cb) in
- let resolve_out,rest' = strengthen_sig mp_from rest mp_to in
- resolve_out,item'::rest'
- | (_,SFBmind _ as item):: rest ->
- let resolve_out,rest' = strengthen_sig mp_from rest mp_to in
- resolve_out,item::rest'
- | (l,SFBmodule mb) :: rest ->
- let mp_from' = MPdot (mp_from,l) in
- let mp_to' = MPdot(mp_to,l) in
- let mb_out = strengthen_mod mp_from' mp_to' mb in
- let item' = l,SFBmodule (mb_out) in
- let resolve_out,rest' = strengthen_sig mp_from rest mp_to in
- resolve_out (*add_delta_resolver resolve_out mb.mod_delta*),
- item':: rest'
- | (_,SFBmodtype _ as item) :: rest ->
- let resolve_out,rest' = strengthen_sig mp_from rest mp_to in
- resolve_out,item::rest'
-
-let strengthen mtb mp =
- strengthen_body ignore mtb.mod_mp mp mtb
-
-let subst_and_strengthen mb mp =
- strengthen_mod mb.mod_mp mp (subst_module (map_mp mb.mod_mp mp) mb)
-
-let module_type_of_module mp mb =
- let mtb =
- { mb with
- mod_expr = ();
- mod_type_alg = None;
- mod_retroknowledge = ModTypeRK }
- in
- match mp with
- | Some mp -> strengthen {mtb with mod_mp = mp} mp
- | None -> mtb
diff --git a/checker/modops.mli b/checker/modops.mli
deleted file mode 100644
index 9f6f781121..0000000000
--- a/checker/modops.mli
+++ /dev/null
@@ -1,49 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(*i*)
-open Names
-open Cic
-open Environ
-(*i*)
-
-(* Various operations on modules and module types *)
-
-val module_type_of_module :
- ModPath.t option -> module_body -> module_type_body
-
-val is_functor : ('ty,'a) functorize -> bool
-
-val destr_functor : ('ty,'a) functorize -> MBId.t * 'ty * ('ty,'a) functorize
-
-(* adds a module and its components, but not the constraints *)
-val add_module : module_body -> env -> env
-
-val add_module_type : ModPath.t -> module_type_body -> env -> env
-
-val strengthen : module_type_body -> ModPath.t -> module_type_body
-
-val subst_and_strengthen : module_body -> ModPath.t -> module_body
-
-val error_incompatible_modtypes :
- module_type_body -> module_type_body -> 'a
-
-val error_not_match : Label.t -> structure_field_body -> 'a
-
-val error_with_module : unit -> 'a
-
-val error_no_such_label : Label.t -> 'a
-
-val error_no_such_label_sub :
- Label.t -> ModPath.t -> 'a
-
-val error_not_a_constant : Label.t -> 'a
-
-val error_not_a_module : Label.t -> 'a
diff --git a/checker/print.ml b/checker/print.ml
deleted file mode 100644
index 247c811f80..0000000000
--- a/checker/print.ml
+++ /dev/null
@@ -1,96 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-open Format
-open Cic
-open Names
-
-let chk_pp fmt = Pp.pp_with fmt
-let pp_arrayi pp fmt a = Array.iteri (fun i x -> pp fmt (i,x)) a
-let pp_instance fmt i = chk_pp fmt (Univ.Instance.pr i)
-let pp_id fmt id = fprintf fmt "%s" (Id.to_string id)
-
-let print_pure_constr fmt csr =
- let rec pp_term fmt c = match c with
- | Rel n -> fprintf fmt "#%d" n
- | Meta n -> fprintf fmt "Meta(%d)" n
- | Var id -> pp_id fmt id
- | Sort s -> pp_sort fmt s
- | Cast (c,_, t) ->
- fprintf fmt "@[<hov 1>(%a@;::%a)@]" pp_term c pp_term t
- | Prod (Name(id),t,c) ->
- fprintf fmt "@[<hov 1>(%a:%a)@;@[%a@]@]" pp_id id pp_term t pp_term c
- | Prod (Anonymous,t,c) ->
- fprintf fmt "(%a@,->@[%a@])" pp_term t pp_term c
- | Lambda (na,t,c) ->
- fprintf fmt "[%a:@[%a@]]@,@[%a@]" pp_name na pp_term t pp_term c
- | LetIn (na,b,t,c) ->
- fprintf fmt "[%a=@[%a@]@,:@[%a@]]@,@[%a@]" pp_name na pp_term b pp_term t pp_term c
- | App (c,l) ->
- fprintf fmt "(@[%a@]@, @[<hov 1>%a@])" pp_term c (pp_arrayi (fun _ (_,s) -> fprintf fmt "@[%a@]@," pp_term s)) l;
- | Evar _ -> pp_print_string fmt "Evar#"
- | Const (c,u) ->
- fprintf fmt "Cons(@[%a,%a@])" sp_con_display c pp_instance u
- | Ind ((sp,i),u) ->
- fprintf fmt "Ind(@[%a,%d,%a@])" sp_display sp i pp_instance u
- | Construct (((sp,i),j),u) ->
- fprintf fmt "Constr(%a,%d,%d,%a)" sp_display sp i j pp_instance u
- | Case (ci,p,c,bl) ->
- let pp_match fmt (_,mc) = fprintf fmt " @[%a@]" pp_term mc in
- fprintf fmt "@[<v><@[%a@]>@,Case@ @[%a@]@ of@[<v>%a@]@,end@]" pp_term p pp_term c (pp_arrayi pp_match) bl
- | Fix ((t,i),(lna,tl,bl)) ->
- let pp_fixc fmt (k,_) =
- fprintf fmt "@[<v 0> %a/%d@,:@[%a@]@,:=@[%a@]@]@," pp_name lna.(k) t.(k) pp_term tl.(k) pp_term bl.(k) in
- fprintf fmt "Fix(%d)@,@[<v>{%a}@]" i (pp_arrayi pp_fixc) tl
- | CoFix(i,(lna,tl,bl)) ->
- let pp_fixc fmt (k,_) =
- fprintf fmt "@[<v 0> %a@,:@[%a@]@,:=@[%a@]@]@," pp_name lna.(k) pp_term tl.(k) pp_term bl.(k) in
- fprintf fmt "CoFix(%d)@,@[<v>{%a}@]" i (pp_arrayi pp_fixc) tl
- | Proj (p, c) ->
- fprintf fmt "Proj(%a,@,@[%a@])" sp_con_display (Projection.constant p) pp_term c
-
- and pp_sort fmt = function
- | Set -> pp_print_string fmt "Set"
- | Prop -> pp_print_string fmt "Prop"
- | Type u -> fprintf fmt "Type(%a)" chk_pp (Univ.pr_uni u)
-
- and pp_name fmt = function
- | Name id -> pp_id fmt id
- | Anonymous -> pp_print_string fmt "_"
-
-(* Remove the top names for library and Scratch to avoid long names *)
- and sp_display fmt sp =
-(* let dir,l = decode_kn sp in
- let ls =
- match List.rev_map Id.to_string (DirPath.repr dir) with
- ("Top"::l)-> l
- | ("Coq"::_::l) -> l
- | l -> l
- in List.iter (fun x -> print_string x; print_string ".") ls;*)
- pp_print_string fmt (MutInd.debug_to_string sp)
-
- and sp_con_display fmt sp =
- (*
- let dir,l = decode_kn sp in
- let ls =
- match List.rev_map Id.to_string (DirPath.repr dir) with
- ("Top"::l)-> l
- | ("Coq"::_::l) -> l
- | l -> l
- in List.iter (fun x -> print_string x; print_string ".") ls;*)
- pp_print_string fmt (Constant.debug_to_string sp)
-
- in
- try
- fprintf fmt "@[%a@]%!" pp_term csr
- with e ->
- pp_print_string fmt (Printexc.to_string e);
- print_flush ();
- raise e
diff --git a/checker/reduction.ml b/checker/reduction.ml
deleted file mode 100644
index 1158152f63..0000000000
--- a/checker/reduction.ml
+++ /dev/null
@@ -1,627 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-open Names
-open CErrors
-open Util
-open Cic
-open Term
-open Closure
-open Esubst
-open Environ
-
-let rec is_empty_stack = function
- [] -> true
- | Zupdate _::s -> is_empty_stack s
- | Zshift _::s -> is_empty_stack s
- | _ -> false
-
-(* Compute the lift to be performed on a term placed in a given stack *)
-let el_stack el stk =
- let n =
- List.fold_left
- (fun i z ->
- match z with
- Zshift n -> i+n
- | _ -> i)
- 0
- stk in
- el_shft n el
-
-let compare_stack_shape stk1 stk2 =
- let rec compare_rec bal stk1 stk2 =
- match (stk1,stk2) with
- ([],[]) -> bal=0
- | ((Zupdate _|Zshift _)::s1, _) -> compare_rec bal s1 stk2
- | (_, (Zupdate _|Zshift _)::s2) -> compare_rec bal stk1 s2
- | (Zapp l1::s1, _) -> compare_rec (bal+Array.length l1) s1 stk2
- | (_, Zapp l2::s2) -> compare_rec (bal-Array.length l2) stk1 s2
- | (Zproj p1::s1, Zproj p2::s2) ->
- Int.equal bal 0 && compare_rec 0 s1 s2
- | ((ZcaseT(c1,_,_,_))::s1,
- (ZcaseT(c2,_,_,_))::s2) ->
- bal=0 (* && c1.ci_ind = c2.ci_ind *) && compare_rec 0 s1 s2
- | (Zfix(_,a1)::s1, Zfix(_,a2)::s2) ->
- bal=0 && compare_rec 0 a1 a2 && compare_rec 0 s1 s2
- | (_,_) -> false in
- compare_rec 0 stk1 stk2
-
-type lft_constr_stack_elt =
- Zlapp of (lift * fconstr) array
- | Zlproj of Names.Projection.Repr.t * lift
- | Zlfix of (lift * fconstr) * lft_constr_stack
- | Zlcase of case_info * lift * fconstr * fconstr array
-and lft_constr_stack = lft_constr_stack_elt list
-
-let rec zlapp v = function
- Zlapp v2 :: s -> zlapp (Array.append v v2) s
- | s -> Zlapp v :: s
-
-let pure_stack lfts stk =
- let rec pure_rec lfts stk =
- match stk with
- [] -> (lfts,[])
- | zi::s ->
- (match (zi,pure_rec lfts s) with
- (Zupdate _,lpstk) -> lpstk
- | (Zshift n,(l,pstk)) -> (el_shft n l, pstk)
- | (Zapp a, (l,pstk)) ->
- (l,zlapp (Array.map (fun t -> (l,t)) a) pstk)
- | (Zproj p, (l,pstk)) ->
- (l, Zlproj (p,l)::pstk)
- | (Zfix(fx,a),(l,pstk)) ->
- let (lfx,pa) = pure_rec l a in
- (l, Zlfix((lfx,fx),pa)::pstk)
- | (ZcaseT(ci,p,br,env),(l,pstk)) ->
- (l,Zlcase(ci,l,mk_clos env p,mk_clos_vect env br)::pstk)
- ) in
- snd (pure_rec lfts stk)
-
-(****************************************************************************)
-(* Reduction Functions *)
-(****************************************************************************)
-
-let whd_betaiotazeta x =
- match x with
- | (Sort _|Var _|Meta _|Evar _|Const _|Ind _|Construct _|
- Prod _|Lambda _|Fix _|CoFix _) -> x
- | _ -> whd_val (create_clos_infos betaiotazeta empty_env) (inject x)
-
-let whd_all env t =
- match t with
- | (Sort _|Meta _|Evar _|Ind _|Construct _|
- Prod _|Lambda _|Fix _|CoFix _) -> t
- | _ -> whd_val (create_clos_infos betadeltaiota env) (inject t)
-
-let whd_allnolet env t =
- match t with
- | (Sort _|Meta _|Evar _|Ind _|Construct _|
- Prod _|Lambda _|Fix _|CoFix _|LetIn _) -> t
- | _ -> whd_val (create_clos_infos betadeltaiotanolet env) (inject t)
-
-(* Beta *)
-
-let beta_appvect c v =
- let rec stacklam env t stack =
- match t, stack with
- Lambda(_,_,c), arg::stacktl -> stacklam (arg::env) c stacktl
- | _ -> applist (substl env t, stack) in
- stacklam [] c (Array.to_list v)
-
-(********************************************************************)
-(* Conversion *)
-(********************************************************************)
-
-type conv_pb =
- | CONV
- | CUMUL
-
-(* Conversion utility functions *)
-type 'a conversion_function = env -> 'a -> 'a -> unit
-
-exception NotConvertible
-exception NotConvertibleVect of int
-
-let convert_universes univ u u' =
- if Univ.Instance.check_eq univ u u' then ()
- else raise NotConvertible
-
-let compare_stacks f fmind fproj lft1 stk1 lft2 stk2 =
- let rec cmp_rec pstk1 pstk2 =
- match (pstk1,pstk2) with
- | (z1::s1, z2::s2) ->
- cmp_rec s1 s2;
- (match (z1,z2) with
- | (Zlapp a1,Zlapp a2) -> Array.iter2 f a1 a2
- | (Zlfix(fx1,a1),Zlfix(fx2,a2)) ->
- f fx1 fx2; cmp_rec a1 a2
- | (Zlproj (c1,l1),Zlproj (c2,l2)) ->
- if not (fproj c1 c2) then
- raise NotConvertible
- | (Zlcase(ci1,l1,p1,br1),Zlcase(ci2,l2,p2,br2)) ->
- if not (fmind ci1.ci_ind ci2.ci_ind) then
- raise NotConvertible;
- f (l1,p1) (l2,p2);
- Array.iter2 (fun c1 c2 -> f (l1,c1) (l2,c2)) br1 br2
- | _ -> assert false)
- | _ -> () in
- if compare_stack_shape stk1 stk2 then
- cmp_rec (pure_stack lft1 stk1) (pure_stack lft2 stk2)
- else raise NotConvertible
-
-let convert_inductive_instances cv_pb cumi u u' univs =
- let len_instance =
- Univ.AUContext.size (Univ.ACumulativityInfo.univ_context cumi) in
- if not ((len_instance = Univ.Instance.length u) &&
- (len_instance = Univ.Instance.length u')) then
- anomaly (Pp.str "Invalid inductive subtyping encountered!")
- else
- let variance = Univ.ACumulativityInfo.variance cumi in
- let comp_cst =
- match cv_pb with
- | CONV ->
- Univ.Variance.eq_constraints variance u u' Univ.Constraint.empty
- | CUMUL ->
- Univ.Variance.leq_constraints variance u u' Univ.Constraint.empty
- in
- if (Univ.check_constraints comp_cst univs) then () else raise NotConvertible
-
-let convert_inductives
- cv_pb (mind, ind) u1 sv1 u2 sv2 univs =
- match mind.mind_universes with
- | Monomorphic_ind _ | Polymorphic_ind _ -> convert_universes univs u1 u2
- | Cumulative_ind cumi ->
- let num_param_arity =
- mind.mind_nparams + mind.mind_packets.(ind).mind_nrealargs
- in
- if not (num_param_arity = sv1 && num_param_arity = sv2) then
- convert_universes univs u1 u2
- else
- convert_inductive_instances cv_pb cumi u1 u2 univs
-
-let convert_constructors
- (mind, ind, cns) u1 sv1 u2 sv2 univs =
- match mind.mind_universes with
- | Monomorphic_ind _ | Polymorphic_ind _ -> convert_universes univs u1 u2
- | Cumulative_ind cumi ->
- let num_cnstr_args =
- mind.mind_nparams + mind.mind_packets.(ind).mind_consnrealargs.(cns - 1)
- in
- if not (num_cnstr_args = sv1 && num_cnstr_args = sv2) then
- convert_universes univs u1 u2
- else
- (** By invariant, both constructors have a common supertype,
- so they are convertible _at that type_. *)
- ()
-
-(* Convertibility of sorts *)
-
-let sort_cmp env univ pb s0 s1 =
- match (s0,s1) with
- | Prop, Prop | Set, Set -> ()
- | Prop, (Set | Type _) | Set, Type _ ->
- if not (pb = CUMUL) then raise NotConvertible
- | Type u1, Type u2 ->
- (** FIXME: handle type-in-type option here *)
- if (* snd (engagement env) == StratifiedType && *)
- not
- (match pb with
- | CONV -> Univ.check_eq univ u1 u2
- | CUMUL -> Univ.check_leq univ u1 u2)
- then begin
- if !Flags.debug then begin
- let op = match pb with CONV -> "=" | CUMUL -> "<=" in
- Format.eprintf "sort_cmp: @[%a@]\n%!" Pp.pp_with Pp.(
- str"Error: " ++ Univ.pr_uni u1 ++ str op ++ Univ.pr_uni u2 ++ str ":" ++ cut()
- ++ Univ.pr_universes univ)
- end;
- raise NotConvertible
- end
- | Set, Prop | Type _, (Prop | Set) -> raise NotConvertible
-
-let rec no_arg_available = function
- | [] -> true
- | Zupdate _ :: stk -> no_arg_available stk
- | Zshift _ :: stk -> no_arg_available stk
- | Zapp v :: stk -> Array.length v = 0 && no_arg_available stk
- | Zproj _ :: _ -> true
- | ZcaseT _ :: _ -> true
- | Zfix _ :: _ -> true
-
-let rec no_nth_arg_available n = function
- | [] -> true
- | Zupdate _ :: stk -> no_nth_arg_available n stk
- | Zshift _ :: stk -> no_nth_arg_available n stk
- | Zapp v :: stk ->
- let k = Array.length v in
- if n >= k then no_nth_arg_available (n-k) stk
- else false
- | Zproj _ :: _ -> true
- | ZcaseT _ :: _ -> true
- | Zfix _ :: _ -> true
-
-let rec no_case_available = function
- | [] -> true
- | Zupdate _ :: stk -> no_case_available stk
- | Zshift _ :: stk -> no_case_available stk
- | Zapp _ :: stk -> no_case_available stk
- | Zproj _ :: _ -> false
- | ZcaseT _ :: _ -> false
- | Zfix _ :: _ -> true
-
-let in_whnf (t,stk) =
- match fterm_of t with
- | (FLetIn _ | FCaseT _ | FApp _ | FCLOS _ | FLIFT _ | FCast _) -> false
- | FLambda _ -> no_arg_available stk
- | FConstruct _ -> no_case_available stk
- | FCoFix _ -> no_case_available stk
- | FFix(((ri,n),(_,_,_)),_) -> no_nth_arg_available ri.(n) stk
- | (FFlex _ | FProd _ | FEvar _ | FInd _ | FAtom _ | FRel _ | FProj _) -> true
- | FLOCKED -> assert false
-
-let default_level = Level 0
-
-let get_strategy { var_opacity; cst_opacity } = function
- | VarKey id ->
- (try Names.Id.Map.find id var_opacity
- with Not_found -> default_level)
- | ConstKey (c, _) ->
- (try Names.Cmap.find c cst_opacity
- with Not_found -> default_level)
- | RelKey _ -> Expand
-
-let dep_order l2r k1 k2 = match k1, k2 with
-| RelKey _, RelKey _ -> l2r
-| RelKey _, (VarKey _ | ConstKey _) -> true
-| VarKey _, RelKey _ -> false
-| VarKey _, VarKey _ -> l2r
-| VarKey _, ConstKey _ -> true
-| ConstKey _, (RelKey _ | VarKey _) -> false
-| ConstKey _, ConstKey _ -> l2r
-
-let oracle_order infos l2r k1 k2 =
- let o = Closure.oracle_of_infos infos in
- match get_strategy o k1, get_strategy o k2 with
- | Expand, Expand -> dep_order l2r k1 k2
- | Expand, (Opaque | Level _) -> true
- | (Opaque | Level _), Expand -> false
- | Opaque, Opaque -> dep_order l2r k1 k2
- | Level _, Opaque -> true
- | Opaque, Level _ -> false
- | Level n1, Level n2 ->
- if Int.equal n1 n2 then dep_order l2r k1 k2
- else n1 < n2
-
-let eq_table_key univ =
- Names.eq_table_key (fun (c1,u1) (c2,u2) ->
- Constant.UserOrd.equal c1 c2 &&
- Univ.Instance.check_eq univ u1 u2)
-
-let proj_equiv_infos infos p1 p2 =
- Int.equal (Projection.Repr.arg p1) (Projection.Repr.arg p2) &&
- mind_equiv (infos_env infos) (Projection.Repr.inductive p1) (Projection.Repr.inductive p2)
-
-(* Conversion between [lft1]term1 and [lft2]term2 *)
-let rec ccnv univ cv_pb infos lft1 lft2 term1 term2 =
- eqappr univ cv_pb infos (lft1, (term1,[])) (lft2, (term2,[]))
-
-(* Conversion between [lft1](hd1 v1) and [lft2](hd2 v2) *)
-and eqappr univ cv_pb infos (lft1,st1) (lft2,st2) =
- Control.check_for_interrupt ();
- (* First head reduce both terms *)
- let rec whd_both (t1,stk1) (t2,stk2) =
- let st1' = whd_stack infos t1 stk1 in
- let st2' = whd_stack infos t2 stk2 in
- (* Now, whd_stack on term2 might have modified st1 (due to sharing),
- and st1 might not be in whnf anymore. If so, we iterate ccnv. *)
- if in_whnf st1' then (st1',st2') else whd_both st1' st2' in
- let ((hd1,v1),(hd2,v2)) = whd_both st1 st2 in
- let appr1 = (lft1,(hd1,v1)) and appr2 = (lft2,(hd2,v2)) in
- (* compute the lifts that apply to the head of the term (hd1 and hd2) *)
- let el1 = el_stack lft1 v1 in
- let el2 = el_stack lft2 v2 in
- match (fterm_of hd1, fterm_of hd2) with
- (* case of leaves *)
- | (FAtom a1, FAtom a2) ->
- (match a1, a2 with
- | (Sort s1, Sort s2) ->
- assert (is_empty_stack v1 && is_empty_stack v2);
- sort_cmp (infos_env infos) univ cv_pb s1 s2
- | (Meta n, Meta m) ->
- if n=m
- then convert_stacks univ infos lft1 lft2 v1 v2
- else raise NotConvertible
- | _ -> raise NotConvertible)
- | (FEvar (ev1,args1), FEvar (ev2,args2)) ->
- if ev1=ev2 then
- (convert_stacks univ infos lft1 lft2 v1 v2;
- convert_vect univ infos el1 el2 args1 args2)
- else raise NotConvertible
-
- (* 2 index known to be bound to no constant *)
- | (FRel n, FRel m) ->
- if reloc_rel n el1 = reloc_rel m el2
- then convert_stacks univ infos lft1 lft2 v1 v2
- else raise NotConvertible
-
- (* 2 constants, 2 local defined vars or 2 defined rels *)
- | (FFlex fl1, FFlex fl2) ->
- (try (* try first intensional equality *)
- if eq_table_key univ fl1 fl2
- then convert_stacks univ infos lft1 lft2 v1 v2
- else raise NotConvertible
- with NotConvertible ->
- (* else the oracle tells which constant is to be expanded *)
- let (app1,app2) =
- if oracle_order infos false fl1 fl2 then
- match unfold_reference infos fl1 with
- | Some def1 -> ((lft1, whd_stack infos def1 v1), appr2)
- | None ->
- (match unfold_reference infos fl2 with
- | Some def2 -> (appr1, (lft2, whd_stack infos def2 v2))
- | None -> raise NotConvertible)
- else
- match unfold_reference infos fl2 with
- | Some def2 -> (appr1, (lft2, whd_stack infos def2 v2))
- | None ->
- (match unfold_reference infos fl1 with
- | Some def1 -> ((lft1, whd_stack infos def1 v1), appr2)
- | None -> raise NotConvertible) in
- eqappr univ cv_pb infos app1 app2)
-
- | (FProj (p1,c1), _) ->
- let s1 = unfold_projection (infos_env infos) p1 in
- eqappr univ cv_pb infos (lft1, whd_stack infos c1 (s1 :: v1)) appr2
-
- | (_, FProj (p2,c2)) ->
- let s2 = unfold_projection (infos_env infos) p2 in
- eqappr univ cv_pb infos appr1 (lft2, whd_stack infos c2 (s2 :: v2))
-
- (* other constructors *)
- | (FLambda _, FLambda _) ->
- (* Inconsistency: we tolerate that v1, v2 contain shift and update but
- we throw them away *)
- assert (is_empty_stack v1 && is_empty_stack v2);
- let (_,ty1,bd1) = destFLambda mk_clos hd1 in
- let (_,ty2,bd2) = destFLambda mk_clos hd2 in
- ccnv univ CONV infos el1 el2 ty1 ty2;
- ccnv univ CONV infos (el_lift el1) (el_lift el2) bd1 bd2
-
- | (FProd (_,c1,c2), FProd (_,c'1,c'2)) ->
- assert (is_empty_stack v1 && is_empty_stack v2);
- (* Luo's system *)
- ccnv univ CONV infos el1 el2 c1 c'1;
- ccnv univ cv_pb infos (el_lift el1) (el_lift el2) c2 c'2
-
- (* Eta-expansion on the fly *)
- | (FLambda _, _) ->
- if v1 <> [] then
- anomaly (Pp.str "conversion was given unreduced term (FLambda).");
- let (_,_ty1,bd1) = destFLambda mk_clos hd1 in
- eqappr univ CONV infos
- (el_lift lft1,(bd1,[])) (el_lift lft2,(hd2,eta_expand_stack v2))
- | (_, FLambda _) ->
- if v2 <> [] then
- anomaly (Pp.str "conversion was given unreduced term (FLambda).");
- let (_,_ty2,bd2) = destFLambda mk_clos hd2 in
- eqappr univ CONV infos
- (el_lift lft1,(hd1,eta_expand_stack v1)) (el_lift lft2,(bd2,[]))
-
- (* only one constant, defined var or defined rel *)
- | (FFlex fl1, c2) ->
- (match unfold_reference infos fl1 with
- | Some def1 ->
- eqappr univ cv_pb infos (lft1, whd_stack infos def1 v1) appr2
- | None ->
- match c2 with
- | FConstruct ((ind2,j2),u2) ->
- (try
- let v2, v1 =
- eta_expand_ind_stack (infos_env infos) ind2 hd2 v2 (snd appr1)
- in convert_stacks univ infos lft1 lft2 v1 v2
- with Not_found -> raise NotConvertible)
- | _ -> raise NotConvertible)
-
- | (c1, FFlex fl2) ->
- (match unfold_reference infos fl2 with
- | Some def2 ->
- eqappr univ cv_pb infos appr1 (lft2, whd_stack infos def2 v2)
- | None ->
- match c1 with
- | FConstruct ((ind1,j1),u1) ->
- (try let v1, v2 =
- eta_expand_ind_stack (infos_env infos) ind1 hd1 v1 (snd appr2)
- in convert_stacks univ infos lft1 lft2 v1 v2
- with Not_found -> raise NotConvertible)
- | _ -> raise NotConvertible)
-
- (* Inductive types: MutInd MutConstruct Fix Cofix *)
-
- | (FInd (ind1,u1), FInd (ind2,u2)) ->
- if mind_equiv_infos infos ind1 ind2 then
- if Univ.Instance.length u1 = 0 || Univ.Instance.length u2 = 0 then
- begin
- convert_universes univ u1 u2;
- convert_stacks univ infos lft1 lft2 v1 v2
- end
- else
- let mind = Environ.lookup_mind (fst ind1) (infos_env infos) in
- let () =
- convert_inductives cv_pb (mind, snd ind1) u1 (stack_args_size v1)
- u2 (stack_args_size v2) univ
- in
- convert_stacks univ infos lft1 lft2 v1 v2
- else raise NotConvertible
-
- | (FConstruct ((ind1,j1),u1), FConstruct ((ind2,j2),u2)) ->
- if Int.equal j1 j2 && mind_equiv_infos infos ind1 ind2 then
- if Univ.Instance.length u1 = 0 || Univ.Instance.length u2 = 0 then
- begin
- convert_universes univ u1 u2;
- convert_stacks univ infos lft1 lft2 v1 v2
- end
- else
- let mind = Environ.lookup_mind (fst ind1) (infos_env infos) in
- let () =
- convert_constructors
- (mind, snd ind1, j1) u1 (stack_args_size v1)
- u2 (stack_args_size v2) univ
- in
- convert_stacks univ infos lft1 lft2 v1 v2
- else raise NotConvertible
-
- (* Eta expansion of records *)
- | (FConstruct ((ind1,j1),u1), _) ->
- (try
- let v1, v2 =
- eta_expand_ind_stack (infos_env infos) ind1 hd1 v1 (snd appr2)
- in convert_stacks univ infos lft1 lft2 v1 v2
- with Not_found -> raise NotConvertible)
-
- | (_, FConstruct ((ind2,j2),u2)) ->
- (try
- let v2, v1 =
- eta_expand_ind_stack (infos_env infos) ind2 hd2 v2 (snd appr1)
- in convert_stacks univ infos lft1 lft2 v1 v2
- with Not_found -> raise NotConvertible)
-
- | (FFix ((op1,(_,tys1,cl1)),e1), FFix((op2,(_,tys2,cl2)),e2)) ->
- if op1 = op2
- then
- let n = Array.length cl1 in
- let fty1 = Array.map (mk_clos e1) tys1 in
- let fty2 = Array.map (mk_clos e2) tys2 in
- let fcl1 = Array.map (mk_clos (subs_liftn n e1)) cl1 in
- let fcl2 = Array.map (mk_clos (subs_liftn n e2)) cl2 in
- convert_vect univ infos el1 el2 fty1 fty2;
- convert_vect univ infos
- (el_liftn n el1) (el_liftn n el2) fcl1 fcl2;
- convert_stacks univ infos lft1 lft2 v1 v2
- else raise NotConvertible
-
- | (FCoFix ((op1,(_,tys1,cl1)),e1), FCoFix((op2,(_,tys2,cl2)),e2)) ->
- if op1 = op2
- then
- let n = Array.length cl1 in
- let fty1 = Array.map (mk_clos e1) tys1 in
- let fty2 = Array.map (mk_clos e2) tys2 in
- let fcl1 = Array.map (mk_clos (subs_liftn n e1)) cl1 in
- let fcl2 = Array.map (mk_clos (subs_liftn n e2)) cl2 in
- convert_vect univ infos el1 el2 fty1 fty2;
- convert_vect univ infos
- (el_liftn n el1) (el_liftn n el2) fcl1 fcl2;
- convert_stacks univ infos lft1 lft2 v1 v2
- else raise NotConvertible
-
- (* Should not happen because both (hd1,v1) and (hd2,v2) are in whnf *)
- | ( (FLetIn _, _) | (FCaseT _,_) | (FApp _,_) | (FCLOS _,_) | (FLIFT _,_)
- | (_, FLetIn _) | (_,FCaseT _) | (_,FApp _) | (_,FCLOS _) | (_,FLIFT _)
- | (FLOCKED,_) | (_,FLOCKED) ) -> assert false
-
- (* In all other cases, terms are not convertible *)
- | _ -> raise NotConvertible
-
-and convert_stacks univ infos lft1 lft2 stk1 stk2 =
- compare_stacks
- (fun (l1,t1) (l2,t2) -> ccnv univ CONV infos l1 l2 t1 t2)
- (mind_equiv_infos infos) (proj_equiv_infos infos)
- lft1 stk1 lft2 stk2
-
-and convert_vect univ infos lft1 lft2 v1 v2 =
- Array.iter2 (fun t1 t2 -> ccnv univ CONV infos lft1 lft2 t1 t2) v1 v2
-
-let clos_fconv cv_pb eager_delta env t1 t2 =
- let infos =
- create_clos_infos
- (if eager_delta then betadeltaiota else betaiotazeta) env in
- let univ = universes env in
- ccnv univ cv_pb infos el_id el_id (inject t1) (inject t2)
-
-let fconv cv_pb eager_delta env t1 t2 =
- if eq_constr t1 t2 then ()
- else clos_fconv cv_pb eager_delta env t1 t2
-
-let conv = fconv CONV false
-let conv_leq = fconv CUMUL false
-
-(* option for conversion : no compilation for the checker *)
-
-let vm_conv cv_pb = fconv cv_pb true
-
-(********************************************************************)
-(* Special-Purpose Reduction *)
-(********************************************************************)
-
-(* pseudo-reduction rule:
- * [hnf_prod_app env s (Prod(_,B)) N --> B[N]
- * with an HNF on the first argument to produce a product.
- * if this does not work, then we use the string S as part of our
- * error message. *)
-
-let hnf_prod_app env t n =
- match whd_all env t with
- | Prod (_,_,b) -> subst1 n b
- | _ -> anomaly ~label:"hnf_prod_app" (Pp.str "Need a product.")
-
-let hnf_prod_applist env t nl =
- List.fold_left (hnf_prod_app env) t nl
-
-(* Dealing with arities *)
-
-let dest_prod env =
- let rec decrec env m c =
- let t = whd_all env c in
- match t with
- | Prod (n,a,c0) ->
- let d = LocalAssum (n,a) in
- decrec (push_rel d env) (d::m) c0
- | _ -> m,t
- in
- decrec env empty_rel_context
-
-(* The same but preserving lets in the context, not internal ones. *)
-let dest_prod_assum env =
- let rec prodec_rec env l ty =
- let rty = whd_allnolet env ty in
- match rty with
- | Prod (x,t,c) ->
- let d = LocalAssum (x,t) in
- prodec_rec (push_rel d env) (d::l) c
- | LetIn (x,b,t,c) ->
- let d = LocalDef (x,b,t) in
- prodec_rec (push_rel d env) (d::l) c
- | _ ->
- let rty' = whd_all env rty in
- if Term.eq_constr rty' rty then l, rty
- else prodec_rec env l rty'
- in
- prodec_rec env empty_rel_context
-
-let dest_lam_assum env =
- let rec lamec_rec env l ty =
- let rty = whd_allnolet env ty in
- match rty with
- | Lambda (x,t,c) ->
- let d = LocalAssum (x,t) in
- lamec_rec (push_rel d env) (d::l) c
- | LetIn (x,b,t,c) ->
- let d = LocalDef (x,b,t) in
- lamec_rec (push_rel d env) (d::l) c
- | _ -> l,rty
- in
- lamec_rec env empty_rel_context
-
-
-let dest_arity env c =
- let l, c = dest_prod_assum env c in
- match c with
- | Sort s -> l,s
- | _ -> user_err Pp.(str "not an arity")
-
diff --git a/checker/reduction.mli b/checker/reduction.mli
deleted file mode 100644
index 3bbf46544d..0000000000
--- a/checker/reduction.mli
+++ /dev/null
@@ -1,55 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(*i*)
-open Cic
-open Term
-open Environ
-(*i*)
-
-(************************************************************************)
-(*s Reduction functions *)
-
-val whd_betaiotazeta : constr -> constr
-val whd_all : env -> constr -> constr
-val whd_allnolet : env -> constr -> constr
-
-(************************************************************************)
-(*s conversion functions *)
-
-exception NotConvertible
-exception NotConvertibleVect of int
-type 'a conversion_function = env -> 'a -> 'a -> unit
-
-type conv_pb = CONV | CUMUL
-
-val conv : constr conversion_function
-val conv_leq : constr conversion_function
-
-val vm_conv : conv_pb -> constr conversion_function
-
-(************************************************************************)
-
-(* Builds an application node, reducing beta redexes it may produce. *)
-val beta_appvect : constr -> constr array -> constr
-
-(* Pseudo-reduction rule Prod(x,A,B) a --> B[x\a] *)
-val hnf_prod_applist : env -> constr -> constr list -> constr
-
-
-(************************************************************************)
-(*s Recognizing products and arities modulo reduction *)
-
-val dest_prod : env -> constr -> rel_context * constr
-val dest_prod_assum : env -> constr -> rel_context * constr
-val dest_lam_assum : env -> constr -> rel_context * constr
-
-
-val dest_arity : env -> constr -> arity
diff --git a/checker/safe_typing.mli b/checker/safe_checking.ml
index 51e5ca3203..6dc2953060 100644
--- a/checker/safe_typing.mli
+++ b/checker/safe_checking.ml
@@ -8,15 +8,16 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-(*i*)
-open Cic
+open Declarations
open Environ
-(*i*)
-val get_env : unit -> env
+let import senv clib univs digest =
+ let mb = Safe_typing.module_of_library clib in
+ let env = Safe_typing.env_of_safe_env senv in
+ let env = push_context_set ~strict:true mb.mod_constraints env in
+ let env = push_context_set ~strict:true univs env in
+ Mod_checking.check_module env mb.mod_mp mb;
+ let (_,senv) = Safe_typing.import clib univs digest senv in senv
-val set_engagement : engagement -> unit
-val import :
- CUnix.physical_path -> compiled_library -> Univ.ContextSet.t -> Cic.vodigest -> unit
-val unsafe_import :
- CUnix.physical_path -> compiled_library -> Univ.ContextSet.t -> Cic.vodigest -> unit
+let unsafe_import senv clib univs digest =
+ let (_,senv) = Safe_typing.import clib univs digest senv in senv
diff --git a/checker/print.mli b/checker/safe_checking.mli
index da1362ca5c..44cd2b3a2e 100644
--- a/checker/print.mli
+++ b/checker/safe_checking.mli
@@ -8,6 +8,9 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-open Cic
+(*i*)
+open Safe_typing
+(*i*)
-val print_pure_constr : Format.formatter -> constr -> unit
+val import : safe_environment -> compiled_library -> Univ.ContextSet.t -> vodigest -> safe_environment
+val unsafe_import : safe_environment -> compiled_library -> Univ.ContextSet.t -> vodigest -> safe_environment
diff --git a/checker/safe_typing.ml b/checker/safe_typing.ml
deleted file mode 100644
index e3640c3795..0000000000
--- a/checker/safe_typing.ml
+++ /dev/null
@@ -1,96 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-open Pp
-open CErrors
-open Util
-open Cic
-open Names
-open Environ
-
-let pr_dirpath dp = str (DirPath.to_string dp)
-
-(************************************************************************)
-(*
- * Global environment
- *)
-
-let genv = ref empty_env
-let get_env () = !genv
-
-let set_engagement c =
- genv := set_engagement c !genv
-
-(* full_add_module adds module with universes and constraints *)
-let full_add_module dp mb univs digest =
- let env = !genv in
- let env = push_context_set ~strict:true mb.mod_constraints env in
- let env = push_context_set ~strict:true univs env in
- let env = Modops.add_module mb env in
- genv := add_digest env dp digest
-
-(* Check that the engagement expected by a library extends the initial one *)
-let check_engagement env expected_impredicative_set =
- let impredicative_set = Environ.engagement env in
- begin
- match impredicative_set, expected_impredicative_set with
- | PredicativeSet, ImpredicativeSet ->
- CErrors.user_err Pp.(str "Needs option -impredicative-set.")
- | _ -> ()
- end;
- ()
-
-(* Libraries = Compiled modules *)
-
-let report_clash f caller dir =
- let msg =
- str "compiled library " ++ pr_dirpath caller ++
- spc() ++ str "makes inconsistent assumptions over library" ++ spc() ++
- pr_dirpath dir ++ fnl() in
- f msg
-
-
-let check_imports f caller env needed =
- let check (dp,stamp) =
- try
- let actual_stamp = lookup_digest env dp in
- if stamp <> actual_stamp then report_clash f caller dp
- with Not_found ->
- user_err Pp.(str ("Reference to unknown module " ^ (DirPath.to_string dp)))
- in
- Array.iter check needed
-
-(* This function should append a certificate to the .vo file.
- The digest must be part of the certicate to rule out attackers
- that could change the .vo file between the time it was read and
- the time the stamp is written.
- For the moment, .vo are not signed. *)
-let stamp_library file digest = ()
-
-(* When the module is checked, digests do not need to match, but a
- warning is issued in case of mismatch *)
-let import file clib univs digest =
- let env = !genv in
- check_imports Feedback.msg_warning clib.comp_name env clib.comp_deps;
- check_engagement env clib.comp_enga;
- let mb = clib.comp_mod in
- Mod_checking.check_module
- (push_context_set ~strict:true univs
- (push_context_set ~strict:true mb.mod_constraints env)) mb.mod_mp mb;
- stamp_library file digest;
- full_add_module clib.comp_name mb univs digest
-
-(* When the module is admitted, digests *must* match *)
-let unsafe_import file clib univs digest =
- let env = !genv in
- if !Flags.debug then check_imports Feedback.msg_warning clib.comp_name env clib.comp_deps
- else check_imports (user_err ~hdr:"unsafe_import") clib.comp_name env clib.comp_deps;
- check_engagement env clib.comp_enga;
- full_add_module clib.comp_name clib.comp_mod univs digest
diff --git a/checker/subtyping.ml b/checker/subtyping.ml
deleted file mode 100644
index e2c605dde8..0000000000
--- a/checker/subtyping.ml
+++ /dev/null
@@ -1,344 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(*i*)
-open Util
-open Names
-open Cic
-open Declarations
-open Environ
-open Reduction
-open Inductive
-open Modops
-(*i*)
-
-(* This local type is used to subtype a constant with a constructor or
- an inductive type. It can also be useful to allow reorderings in
- inductive types *)
-type namedobject =
- | Constant of constant_body
- | IndType of inductive * mutual_inductive_body
- | IndConstr of constructor * mutual_inductive_body
-
-type namedmodule =
- | Module of module_body
- | Modtype of module_type_body
-
-(* adds above information about one mutual inductive: all types and
- constructors *)
-
-let add_mib_nameobjects mp l mib map =
- let ind = MutInd.make2 mp l in
- let add_mip_nameobjects j oib map =
- let ip = (ind,j) in
- let map =
- Array.fold_right_i
- (fun i id map ->
- Label.Map.add (Label.of_id id) (IndConstr((ip,i+1), mib)) map)
- oib.mind_consnames
- map
- in
- Label.Map.add (Label.of_id oib.mind_typename) (IndType (ip, mib)) map
- in
- Array.fold_right_i add_mip_nameobjects mib.mind_packets map
-
-
-(* creates (namedobject/namedmodule) map for the whole signature *)
-
-type labmap = { objs : namedobject Label.Map.t; mods : namedmodule Label.Map.t }
-
-let empty_labmap = { objs = Label.Map.empty; mods = Label.Map.empty }
-
-let get_obj mp map l =
- try Label.Map.find l map.objs
- with Not_found -> error_no_such_label_sub l mp
-
-let get_mod mp map l =
- try Label.Map.find l map.mods
- with Not_found -> error_no_such_label_sub l mp
-
-let make_labmap mp list =
- let add_one (l,e) map =
- match e with
- | SFBconst cb -> { map with objs = Label.Map.add l (Constant cb) map.objs }
- | SFBmind mib -> { map with objs = add_mib_nameobjects mp l mib map.objs }
- | SFBmodule mb -> { map with mods = Label.Map.add l (Module mb) map.mods }
- | SFBmodtype mtb -> { map with mods = Label.Map.add l (Modtype mtb) map.mods }
- in
- List.fold_right add_one list empty_labmap
-
-
-let check_conv_error error f env a1 a2 =
- try
- f env a1 a2
- with
- NotConvertible -> error ()
-
-let check_polymorphic_instance error env auctx1 auctx2 =
- if not (Univ.AUContext.size auctx1 == Univ.AUContext.size auctx2) then
- error ()
- else if not (Univ.check_subtype (Environ.universes env) auctx2 auctx1) then
- error ()
- else
- Environ.push_context ~strict:false (Univ.AUContext.repr auctx2) env
-
-(* for now we do not allow reorderings *)
-let check_inductive env mp1 l info1 mib2 spec2 subst1 subst2=
- let kn = MutInd.make2 mp1 l in
- let error () = error_not_match l spec2 in
- let check_conv f = check_conv_error error f in
- let mib1 =
- match info1 with
- | IndType ((_,0), mib) -> subst_mind subst1 mib
- | _ -> error ()
- in
- let mib2 = subst_mind subst2 mib2 in
- let check eq f = if not (eq (f mib1) (f mib2)) then error () in
- let env, u =
- match mib1.mind_universes, mib2.mind_universes with
- | Monomorphic_ind _, Monomorphic_ind _ -> env, Univ.Instance.empty
- | Polymorphic_ind auctx, Polymorphic_ind auctx' ->
- let env = check_polymorphic_instance error env auctx auctx' in
- env, Univ.make_abstract_instance auctx'
- | Cumulative_ind cumi, Cumulative_ind cumi' ->
- (** Currently there is no way to control variance of inductive types, but
- just in case we require that they are in a subtyping relation. *)
- let () =
- let v = Univ.ACumulativityInfo.variance cumi in
- let v' = Univ.ACumulativityInfo.variance cumi' in
- if not (Array.for_all2 Univ.Variance.check_subtype v' v) then
- CErrors.anomaly Pp.(str "Variance mismatch for " ++ MutInd.print kn)
- in
- let auctx = Univ.ACumulativityInfo.univ_context cumi in
- let auctx' = Univ.ACumulativityInfo.univ_context cumi' in
- let env = check_polymorphic_instance error env auctx auctx' in
- env, Univ.make_abstract_instance auctx'
- | _ -> error ()
- in
- let check_inductive_type t1 t2 = check_conv conv_leq env t1 t2 in
-
- let check_packet p1 p2 =
- let check eq f = if not (eq (f p1) (f p2)) then error () in
- check
- (fun a1 a2 -> Array.equal Id.equal a1 a2)
- (fun p -> p.mind_consnames);
- check Id.equal (fun p -> p.mind_typename);
- (* nf_lc later *)
- (* nf_arity later *)
- (* user_lc ignored *)
- (* user_arity ignored *)
- check Int.equal (fun p -> p.mind_nrealargs);
- (* kelim ignored *)
- (* listrec ignored *)
- (* finite done *)
- (* nparams done *)
- (* params_ctxt done because part of the inductive types *)
- (* Don't check the sort of the type if polymorphic *)
- check_inductive_type
- (type_of_inductive env ((mib1,p1), u)) (type_of_inductive env ((mib2,p2),u))
- in
- let check_cons_types i p1 p2 =
- Array.iter2 (check_conv conv env)
- (arities_of_specif (kn,u) (mib1,p1))
- (arities_of_specif (kn,u) (mib2,p2))
- in
- check (==) (fun mib -> mib.mind_finite);
- check Int.equal (fun mib -> mib.mind_ntypes);
- assert (Array.length mib1.mind_packets >= 1
- && Array.length mib2.mind_packets >= 1);
-
- (* Check that the expected numbers of uniform parameters are the same *)
- (* No need to check the contexts of parameters: it is checked *)
- (* at the time of checking the inductive arities in check_packet. *)
- (* Notice that we don't expect the local definitions to match: only *)
- (* the inductive types and constructors types have to be convertible *)
- check Int.equal (fun mib -> mib.mind_nparams);
-
- (*begin
- match mib2.mind_equiv with
- | None -> ()
- | Some kn2' ->
- let kn2 = scrape_mind env kn2' in
- let kn1 = match mib1.mind_equiv with
- None -> kn
- | Some kn1' -> scrape_mind env kn1'
- in
- if kn1 <> kn2 then error ()
- end;*)
- (* we check that records and their field names are preserved. *)
- let record_equal x y =
- match x, y with
- | NotRecord, NotRecord -> true
- | FakeRecord, FakeRecord -> true
- | PrimRecord info1, PrimRecord info2 ->
- let check (id1, p1, pb1) (id2, p2, pb2) =
- (* we don't care about the id, the types are inferred from the inductive
- (ie checked before now) *)
- Array.for_all2 Label.equal p1 p2
- in
- Array.equal check info1 info2
- | _, _ -> false
- in
- check record_equal (fun mib -> mib.mind_record);
- if mib1.mind_record != NotRecord then begin
- let rec names_prod_letin t = match t with
- | Prod(n,_,t) -> n::(names_prod_letin t)
- | LetIn(n,_,_,t) -> n::(names_prod_letin t)
- | Cast(t,_,_) -> names_prod_letin t
- | _ -> []
- in
- assert (Array.length mib1.mind_packets = 1);
- assert (Array.length mib2.mind_packets = 1);
- assert (Array.length mib1.mind_packets.(0).mind_user_lc = 1);
- assert (Array.length mib2.mind_packets.(0).mind_user_lc = 1);
- check (List.equal Name.equal)
- (fun mib ->
- let nparamdecls = List.length mib.mind_params_ctxt in
- let names = names_prod_letin (mib.mind_packets.(0).mind_user_lc.(0)) in
- snd (List.chop nparamdecls names))
- end;
- (* we first check simple things *)
- Array.iter2 check_packet mib1.mind_packets mib2.mind_packets;
- (* and constructor types in the end *)
- let _ = Array.map2_i check_cons_types mib1.mind_packets mib2.mind_packets
- in ()
-
-let check_constant env l info1 cb2 spec2 subst1 subst2 =
- let error () = error_not_match l spec2 in
- let check_conv f = check_conv_error error f in
- let check_type env t1 t2 = check_conv conv_leq env t1 t2 in
- match info1 with
- | Constant cb1 ->
- let cb1 = subst_const_body subst1 cb1 in
- let cb2 = subst_const_body subst2 cb2 in
- (*Start by checking universes *)
- let env =
- match cb1.const_universes, cb2.const_universes with
- | Monomorphic_const _, Monomorphic_const _ -> env
- | Polymorphic_const auctx1, Polymorphic_const auctx2 ->
- check_polymorphic_instance error env auctx1 auctx2
- | Monomorphic_const _, Polymorphic_const _ ->
- error ()
- | Polymorphic_const _, Monomorphic_const _ ->
- error ()
- in
- (* Now check types *)
- let typ1 = cb1.const_type in
- let typ2 = cb2.const_type in
- check_type env typ1 typ2;
- (* Now we check the bodies:
- - A transparent constant can only be implemented by a compatible
- transparent constant.
- - In the signature, an opaque is handled just as a parameter:
- anything of the right type can implement it, even if bodies differ.
- *)
- (match cb2.const_body with
- | Undef _ | OpaqueDef _ -> ()
- | Def lc2 ->
- (match cb1.const_body with
- | Undef _ | OpaqueDef _ -> error ()
- | Def lc1 ->
- (* NB: cb1 might have been strengthened and appear as transparent.
- Anyway [check_conv] will handle that afterwards. *)
- let c1 = force_constr lc1 in
- let c2 = force_constr lc2 in
- check_conv conv env c1 c2))
- | IndType ((kn,i),mind1) ->
- CErrors.user_err (Pp.str (
- "The kernel does not recognize yet that a parameter can be " ^
- "instantiated by an inductive type. Hint: you can rename the " ^
- "inductive type and give a definition to map the old name to the new " ^
- "name."))
- | IndConstr (((kn,i),j),mind1) ->
- CErrors.user_err (Pp.str (
- "The kernel does not recognize yet that a parameter can be " ^
- "instantiated by a constructor. Hint: you can rename the " ^
- "constructor and give a definition to map the old name to the new " ^
- "name."))
-
-let rec check_modules env msb1 msb2 subst1 subst2 =
- let mty1 = module_type_of_module None msb1 in
- let mty2 = module_type_of_module None msb2 in
- check_modtypes env mty1 mty2 subst1 subst2 false;
-
-
-and check_signatures env mp1 sig1 sig2 subst1 subst2 =
- let map1 = make_labmap mp1 sig1 in
- let check_one_body (l,spec2) =
- match spec2 with
- | SFBconst cb2 ->
- check_constant env l (get_obj mp1 map1 l)
- cb2 spec2 subst1 subst2
- | SFBmind mib2 ->
- check_inductive env mp1 l (get_obj mp1 map1 l)
- mib2 spec2 subst1 subst2
- | SFBmodule msb2 ->
- begin
- match get_mod mp1 map1 l with
- | Module msb -> check_modules env msb msb2
- subst1 subst2
- | _ -> error_not_match l spec2
- end
- | SFBmodtype mtb2 ->
- let mtb1 =
- match get_mod mp1 map1 l with
- | Modtype mtb -> mtb
- | _ -> error_not_match l spec2
- in
- let env =
- add_module_type mtb2.mod_mp mtb2
- (add_module_type mtb1.mod_mp mtb1 env)
- in
- check_modtypes env mtb1 mtb2 subst1 subst2 true
- in
- List.iter check_one_body sig2
-
-and check_modtypes env mtb1 mtb2 subst1 subst2 equiv =
- if mtb1==mtb2 || mtb1.mod_type == mtb2.mod_type then ()
- else
- let rec check_structure env str1 str2 equiv subst1 subst2 =
- match str1,str2 with
- | NoFunctor (list1),
- NoFunctor (list2) ->
- check_signatures env mtb1.mod_mp list1 list2 subst1 subst2;
- if equiv then
- check_signatures env mtb2.mod_mp list2 list1 subst1 subst2
- else
- ()
- | MoreFunctor (arg_id1,arg_t1,body_t1),
- MoreFunctor (arg_id2,arg_t2,body_t2) ->
- check_modtypes env
- arg_t2 arg_t1
- (map_mp arg_t1.mod_mp arg_t2.mod_mp) subst2
- equiv;
- (* contravariant *)
- let env = add_module_type (MPbound arg_id2) arg_t2 env in
- let env =
- if is_functor body_t1 then env
- else
- let env = shallow_remove_module mtb1.mod_mp env in
- add_module {mod_mp = mtb1.mod_mp;
- mod_expr = Abstract;
- mod_type = body_t1;
- mod_type_alg = None;
- mod_constraints = mtb1.mod_constraints;
- mod_retroknowledge = ModBodyRK [];
- mod_delta = mtb1.mod_delta} env
- in
- check_structure env body_t1 body_t2 equiv
- (join (map_mbid arg_id1 (MPbound arg_id2)) subst1)
- subst2
- | _ , _ -> error_incompatible_modtypes mtb1 mtb2
- in
- check_structure env mtb1.mod_type mtb2.mod_type equiv subst1 subst2
-
-let check_subtypes env sup super =
- check_modtypes env (strengthen sup sup.mod_mp) super empty_subst
- (map_mp super.mod_mp sup.mod_mp) false
diff --git a/checker/term.ml b/checker/term.ml
deleted file mode 100644
index d84491b38f..0000000000
--- a/checker/term.ml
+++ /dev/null
@@ -1,447 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(* This module instantiates the structure of generic de Bruijn terms to Coq *)
-
-open CErrors
-open Util
-open Names
-open Esubst
-
-open Cic
-
-(* Sorts. *)
-
-let family_of_sort = function
- | Prop -> InProp
- | Set -> InSet
- | Type _ -> InType
-
-let family_equal = (==)
-
-let sort_of_univ u =
- if Univ.is_type0m_univ u then Prop
- else if Univ.is_type0_univ u then Set
- else Type u
-
-(********************************************************************)
-(* Constructions as implemented *)
-(********************************************************************)
-
-let rec strip_outer_cast c = match c with
- | Cast (c,_,_) -> strip_outer_cast c
- | _ -> c
-
-let collapse_appl c = match c with
- | App (f,cl) ->
- let rec collapse_rec f cl2 =
- match (strip_outer_cast f) with
- | App (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
- | _ -> App (f,cl2)
- in
- collapse_rec f cl
- | _ -> c
-
-let decompose_app c =
- match collapse_appl c with
- | App (f,cl) -> (f, Array.to_list cl)
- | _ -> (c,[])
-
-
-let applist (f,l) = App (f, Array.of_list l)
-
-
-(****************************************************************************)
-(* Functions for dealing with constr terms *)
-(****************************************************************************)
-
-(*********************)
-(* Occurring *)
-(*********************)
-
-let iter_constr_with_binders g f n c = match c with
- | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
- | Construct _) -> ()
- | Cast (c,_,t) -> f n c; f n t
- | Prod (_,t,c) -> f n t; f (g n) c
- | Lambda (_,t,c) -> f n t; f (g n) c
- | LetIn (_,b,t,c) -> f n b; f n t; f (g n) c
- | App (c,l) -> f n c; Array.iter (f n) l
- | Evar (_,l) -> Array.iter (f n) l
- | Case (_,p,c,bl) -> f n p; f n c; Array.iter (f n) bl
- | Fix (_,(_,tl,bl)) ->
- Array.iter (f n) tl;
- Array.iter (f (iterate g (Array.length tl) n)) bl
- | CoFix (_,(_,tl,bl)) ->
- Array.iter (f n) tl;
- Array.iter (f (iterate g (Array.length tl) n)) bl
- | Proj (p, c) -> f n c
-
-exception LocalOccur
-
-(* (closedn n M) raises FreeVar if a variable of height greater than n
- occurs in M, returns () otherwise *)
-
-let closedn n c =
- let rec closed_rec n c = match c with
- | Rel m -> if m>n then raise LocalOccur
- | _ -> iter_constr_with_binders succ closed_rec n c
- in
- try closed_rec n c; true with LocalOccur -> false
-
-(* [closed0 M] is true iff [M] is a (de Bruijn) closed term *)
-
-let closed0 = closedn 0
-
-(* (noccurn n M) returns true iff (Rel n) does NOT occur in term M *)
-
-let noccurn n term =
- let rec occur_rec n c = match c with
- | Rel m -> if Int.equal m n then raise LocalOccur
- | _ -> iter_constr_with_binders succ occur_rec n c
- in
- try occur_rec n term; true with LocalOccur -> false
-
-(* (noccur_between n m M) returns true iff (Rel p) does NOT occur in term M
- for n <= p < n+m *)
-
-let noccur_between n m term =
- let rec occur_rec n c = match c with
- | Rel(p) -> if n<=p && p<n+m then raise LocalOccur
- | _ -> iter_constr_with_binders succ occur_rec n c
- in
- try occur_rec n term; true with LocalOccur -> false
-
-(* Checking function for terms containing existential variables.
- The function [noccur_with_meta] considers the fact that
- each existential variable (as well as each isevar)
- in the term appears applied to its local context,
- which may contain the CoFix variables. These occurrences of CoFix variables
- are not considered *)
-
-let noccur_with_meta n m term =
- let rec occur_rec n c = match c with
- | Rel p -> if n<=p && p<n+m then raise LocalOccur
- | App(f,cl) ->
- (match f with
- | (Cast (Meta _,_,_)| Meta _) -> ()
- | _ -> iter_constr_with_binders succ occur_rec n c)
- | Evar (_, _) -> ()
- | _ -> iter_constr_with_binders succ occur_rec n c
- in
- try (occur_rec n term; true) with LocalOccur -> false
-
-
-(*********************)
-(* Lifting *)
-(*********************)
-
-let map_constr_with_binders g f l c = match c with
- | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
- | Construct _) -> c
- | Cast (c,k,t) -> Cast (f l c, k, f l t)
- | Prod (na,t,c) -> Prod (na, f l t, f (g l) c)
- | Lambda (na,t,c) -> Lambda (na, f l t, f (g l) c)
- | LetIn (na,b,t,c) -> LetIn (na, f l b, f l t, f (g l) c)
- | App (c,al) -> App (f l c, Array.map (f l) al)
- | Evar (e,al) -> Evar (e, Array.map (f l) al)
- | Case (ci,p,c,bl) -> Case (ci, f l p, f l c, Array.map (f l) bl)
- | Fix (ln,(lna,tl,bl)) ->
- let l' = iterate g (Array.length tl) l in
- Fix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl))
- | CoFix(ln,(lna,tl,bl)) ->
- let l' = iterate g (Array.length tl) l in
- CoFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl))
- | Proj (p, c) -> Proj (p, f l c)
-
-(* The generic lifting function *)
-let rec exliftn el c = match c with
- | Rel i -> Rel(reloc_rel i el)
- | _ -> map_constr_with_binders el_lift exliftn el c
-
-(* Lifting the binding depth across k bindings *)
-
-let liftn k n =
- match el_liftn (pred n) (el_shft k el_id) with
- | ELID -> (fun c -> c)
- | el -> exliftn el
-
-let lift k = liftn k 1
-
-(*********************)
-(* Substituting *)
-(*********************)
-
-(* (subst1 M c) substitutes M for Rel(1) in c
- we generalise it to (substl [M1,...,Mn] c) which substitutes in parallel
- M1,...,Mn for respectively Rel(1),...,Rel(n) in c *)
-
-(* 1st : general case *)
-type info = Closed | Open | Unknown
-type 'a substituend = { mutable sinfo: info; sit: 'a }
-
-let rec lift_substituend depth s =
- match s.sinfo with
- | Closed -> s.sit
- | Open -> lift depth s.sit
- | Unknown ->
- s.sinfo <- if closed0 s.sit then Closed else Open;
- lift_substituend depth s
-
-let make_substituend c = { sinfo=Unknown; sit=c }
-
-let substn_many lamv n c =
- let lv = Array.length lamv in
- if Int.equal lv 0 then c
- else
- let rec substrec depth c = match c with
- | Rel k ->
- if k<=depth then c
- else if k-depth <= lv then lift_substituend depth lamv.(k-depth-1)
- else Rel (k-lv)
- | _ -> map_constr_with_binders succ substrec depth c in
- substrec n c
-
-let substnl laml n =
- substn_many (Array.map make_substituend (Array.of_list laml)) n
-let substl laml = substnl laml 0
-let subst1 lam = substl [lam]
-
-
-(***************************************************************************)
-(* Type of assumptions and contexts *)
-(***************************************************************************)
-
-let empty_rel_context = []
-let rel_context_length = List.length
-let rel_context_nhyps hyps =
- let rec nhyps acc = function
- | [] -> acc
- | LocalAssum _ :: hyps -> nhyps (1+acc) hyps
- | LocalDef _ :: hyps -> nhyps acc hyps in
- nhyps 0 hyps
-let fold_rel_context f l ~init = List.fold_right f l init
-
-let fold_rel_context_outside f l ~init = List.fold_right f l init
-
-let map_rel_decl f = function
- | LocalAssum (n, typ) as decl ->
- let typ' = f typ in
- if typ' == typ then decl else
- LocalAssum (n, typ')
- | LocalDef (n, body, typ) as decl ->
- let body' = f body in
- let typ' = f typ in
- if body' == body && typ' == typ then decl else
- LocalDef (n, body', typ')
-
-let map_rel_context f =
- List.Smart.map (map_rel_decl f)
-
-let extended_rel_list n hyps =
- let rec reln l p = function
- | LocalAssum _ :: hyps -> reln (Rel (n+p) :: l) (p+1) hyps
- | LocalDef _ :: hyps -> reln l (p+1) hyps
- | [] -> l
- in
- reln [] 1 hyps
-
-(* Iterate lambda abstractions *)
-
-(* compose_lam [xn:Tn;..;x1:T1] b = [x1:T1]..[xn:Tn]b *)
-let compose_lam l b =
- let rec lamrec = function
- | ([], b) -> b
- | ((v,t)::l, b) -> lamrec (l, Lambda (v,t,b))
- in
- lamrec (l,b)
-
-(* Transforms a lambda term [x1:T1]..[xn:Tn]T into the pair
- ([(xn,Tn);...;(x1,T1)],T), where T is not a lambda *)
-let decompose_lam =
- let rec lamdec_rec l c = match c with
- | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) c
- | Cast (c,_,_) -> lamdec_rec l c
- | _ -> l,c
- in
- lamdec_rec []
-
-(* Decompose lambda abstractions and lets, until finding n
- abstractions *)
-let decompose_lam_n_assum n =
- if n < 0 then
- user_err Pp.(str "decompose_lam_n_assum: integer parameter must be positive");
- let rec lamdec_rec l n c =
- if Int.equal n 0 then l,c
- else match c with
- | Lambda (x,t,c) -> lamdec_rec (LocalAssum (x,t) :: l) (n-1) c
- | LetIn (x,b,t,c) -> lamdec_rec (LocalDef (x,b,t) :: l) n c
- | Cast (c,_,_) -> lamdec_rec l n c
- | c -> user_err Pp.(str "decompose_lam_n_assum: not enough abstractions")
- in
- lamdec_rec empty_rel_context n
-
-(* Iterate products, with or without lets *)
-
-(* Constructs either [(x:t)c] or [[x=b:t]c] *)
-let mkProd_or_LetIn decl c =
- match decl with
- | LocalAssum (na,t) -> Prod (na, t, c)
- | LocalDef (na,b,t) -> LetIn (na, b, t, c)
-
-let it_mkProd_or_LetIn = List.fold_left (fun c d -> mkProd_or_LetIn d c)
-
-let decompose_prod_assum =
- let rec prodec_rec l c =
- match c with
- | Prod (x,t,c) -> prodec_rec (LocalAssum (x,t) :: l) c
- | LetIn (x,b,t,c) -> prodec_rec (LocalDef (x,b,t) :: l) c
- | Cast (c,_,_) -> prodec_rec l c
- | _ -> l,c
- in
- prodec_rec empty_rel_context
-
-let decompose_prod_n_assum n =
- if n < 0 then
- user_err Pp.(str "decompose_prod_n_assum: integer parameter must be positive");
- let rec prodec_rec l n c =
- if Int.equal n 0 then l,c
- else match c with
- | Prod (x,t,c) -> prodec_rec (LocalAssum (x,t) :: l) (n-1) c
- | LetIn (x,b,t,c) -> prodec_rec (LocalDef (x,b,t) :: l) (n-1) c
- | Cast (c,_,_) -> prodec_rec l n c
- | c -> user_err Pp.(str "decompose_prod_n_assum: not enough assumptions")
- in
- prodec_rec empty_rel_context n
-
-
-(***************************)
-(* Other term constructors *)
-(***************************)
-
-type arity = rel_context * sorts
-
-let mkArity (sign,s) = it_mkProd_or_LetIn (Sort s) sign
-
-let destArity =
- let rec prodec_rec l c =
- match c with
- | Prod (x,t,c) -> prodec_rec (LocalAssum (x,t)::l) c
- | LetIn (x,b,t,c) -> prodec_rec (LocalDef (x,b,t)::l) c
- | Cast (c,_,_) -> prodec_rec l c
- | Sort s -> l,s
- | _ -> anomaly ~label:"destArity" (Pp.str "not an arity.")
- in
- prodec_rec []
-
-let rec isArity c =
- match c with
- | Prod (_,_,c) -> isArity c
- | LetIn (_,b,_,c) -> isArity (subst1 b c)
- | Cast (c,_,_) -> isArity c
- | Sort _ -> true
- | _ -> false
-
-(*******************************)
-(* alpha conversion functions *)
-(*******************************)
-
-(* alpha conversion : ignore print names and casts *)
-
-let compare_sorts s1 s2 = match s1, s2 with
-| Prop, Prop | Set, Set -> true
-| Type u1, Type u2 -> Univ.Universe.equal u1 u2
-| Prop, Set | Set, Prop -> false
-| (Prop | Set), Type _ -> false
-| Type _, (Prop | Set) -> false
-
-let eq_puniverses f (c1,u1) (c2,u2) =
- Univ.Instance.equal u1 u2 && f c1 c2
-
-let compare_constr f t1 t2 =
- match t1, t2 with
- | Rel n1, Rel n2 -> Int.equal n1 n2
- | Meta m1, Meta m2 -> Int.equal m1 m2
- | Var id1, Var id2 -> Id.equal id1 id2
- | Sort s1, Sort s2 -> compare_sorts s1 s2
- | Cast (c1,_,_), _ -> f c1 t2
- | _, Cast (c2,_,_) -> f t1 c2
- | Prod (_,t1,c1), Prod (_,t2,c2) -> f t1 t2 && f c1 c2
- | Lambda (_,t1,c1), Lambda (_,t2,c2) -> f t1 t2 && f c1 c2
- | LetIn (_,b1,t1,c1), LetIn (_,b2,t2,c2) -> f b1 b2 && f t1 t2 && f c1 c2
- | App (c1,l1), App (c2,l2) ->
- if Int.equal (Array.length l1) (Array.length l2) then
- f c1 c2 && Array.for_all2 f l1 l2
- else
- let (h1,l1) = decompose_app t1 in
- let (h2,l2) = decompose_app t2 in
- if Int.equal (List.length l1) (List.length l2) then
- f h1 h2 && List.for_all2 f l1 l2
- else false
- | Evar (e1,l1), Evar (e2,l2) -> Int.equal e1 e2 && Array.equal f l1 l2
- | Const c1, Const c2 -> eq_puniverses Constant.UserOrd.equal c1 c2
- | Ind c1, Ind c2 -> eq_puniverses eq_ind_chk c1 c2
- | Construct ((c1,i1),u1), Construct ((c2,i2),u2) -> Int.equal i1 i2 && eq_ind_chk c1 c2
- && Univ.Instance.equal u1 u2
- | Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) ->
- f p1 p2 && f c1 c2 && Array.equal f bl1 bl2
- | Fix ((ln1, i1),(_,tl1,bl1)), Fix ((ln2, i2),(_,tl2,bl2)) ->
- Int.equal i1 i2 && Array.equal Int.equal ln1 ln2 &&
- Array.equal f tl1 tl2 && Array.equal f bl1 bl2
- | CoFix(ln1,(_,tl1,bl1)), CoFix(ln2,(_,tl2,bl2)) ->
- Int.equal ln1 ln2 && Array.equal f tl1 tl2 && Array.equal f bl1 bl2
- | Proj (p1,c1), Proj(p2,c2) -> Projection.equal p1 p2 && f c1 c2
- | _ -> false
-
-let rec eq_constr m n =
- (m == n) ||
- compare_constr eq_constr m n
-
-let eq_constr m n = eq_constr m n (* to avoid tracing a recursive fun *)
-
-(* Universe substitutions *)
-
-let map_constr f c = map_constr_with_binders (fun x -> x) (fun _ c -> f c) 0 c
-
-let subst_instance_constr subst c =
- if Univ.Instance.is_empty subst then c
- else
- let f u = Univ.subst_instance_instance subst u in
- let rec aux t =
- match t with
- | Const (c, u) ->
- if Univ.Instance.is_empty u then t
- else
- let u' = f u in
- if u' == u then t
- else (Const (c, u'))
- | Ind (i, u) ->
- if Univ.Instance.is_empty u then t
- else
- let u' = f u in
- if u' == u then t
- else (Ind (i, u'))
- | Construct (c, u) ->
- if Univ.Instance.is_empty u then t
- else
- let u' = f u in
- if u' == u then t
- else (Construct (c, u'))
- | Sort (Type u) ->
- let u' = Univ.subst_instance_universe subst u in
- if u' == u then t else
- (Sort (sort_of_univ u'))
- | _ -> map_constr aux t
- in
- aux c
-
-let subst_instance_context s ctx =
- if Univ.Instance.is_empty s then ctx
- else map_rel_context (fun x -> subst_instance_constr s x) ctx
diff --git a/checker/term.mli b/checker/term.mli
deleted file mode 100644
index 2524dff189..0000000000
--- a/checker/term.mli
+++ /dev/null
@@ -1,59 +0,0 @@
-open Names
-open Cic
-
-val family_of_sort : sorts -> sorts_family
-val family_equal : sorts_family -> sorts_family -> bool
-
-val decompose_app : constr -> constr * constr list
-val applist : constr * constr list -> constr
-val iter_constr_with_binders :
- ('a -> 'a) -> ('a -> constr -> unit) -> 'a -> constr -> unit
-exception LocalOccur
-val closedn : int -> constr -> bool
-val closed0 : constr -> bool
-val noccurn : int -> constr -> bool
-val noccur_between : int -> int -> constr -> bool
-val noccur_with_meta : int -> int -> constr -> bool
-val map_constr_with_binders :
- ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr
-val exliftn : Esubst.lift -> constr -> constr
-val liftn : int -> int -> constr -> constr
-val lift : int -> constr -> constr
-type info = Closed | Open | Unknown
-type 'a substituend = { mutable sinfo : info; sit : 'a; }
-val lift_substituend : int -> constr substituend -> constr
-val make_substituend : 'a -> 'a substituend
-val substn_many : constr substituend array -> int -> constr -> constr
-val substnl : constr list -> int -> constr -> constr
-val substl : constr list -> constr -> constr
-val subst1 : constr -> constr -> constr
-
-val empty_rel_context : rel_context
-val rel_context_length : rel_context -> int
-val rel_context_nhyps : rel_context -> int
-val fold_rel_context :
- (rel_declaration -> 'a -> 'a) -> rel_context -> init:'a -> 'a
-val fold_rel_context_outside :
- (rel_declaration -> 'a -> 'a) -> rel_context -> init:'a -> 'a
-val map_rel_decl : (constr -> constr) -> rel_declaration -> rel_declaration
-val map_rel_context : (constr -> constr) -> rel_context -> rel_context
-val extended_rel_list : int -> rel_context -> constr list
-val compose_lam : (Name.t * constr) list -> constr -> constr
-val decompose_lam : constr -> (Name.t * constr) list * constr
-val decompose_lam_n_assum : int -> constr -> rel_context * constr
-val mkProd_or_LetIn : rel_declaration -> constr -> constr
-val it_mkProd_or_LetIn : constr -> rel_context -> constr
-val decompose_prod_assum : constr -> rel_context * constr
-val decompose_prod_n_assum : int -> constr -> rel_context * constr
-
-type arity = rel_context * sorts
-
-val mkArity : arity -> constr
-val destArity : constr -> arity
-val isArity : constr -> bool
-val compare_constr : (constr -> constr -> bool) -> constr -> constr -> bool
-val eq_constr : constr -> constr -> bool
-
-(** Instance substitution for polymorphism. *)
-val subst_instance_constr : Univ.universe_instance -> constr -> constr
-val subst_instance_context : Univ.universe_instance -> rel_context -> rel_context
diff --git a/checker/type_errors.ml b/checker/type_errors.ml
deleted file mode 100644
index 5079643783..0000000000
--- a/checker/type_errors.ml
+++ /dev/null
@@ -1,115 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-open Names
-open Cic
-open Environ
-
-type unsafe_judgment = constr * constr
-
-let nf_betaiota c = c
-
-(* Type errors. *)
-
-type guard_error =
- (* Fixpoints *)
- | NotEnoughAbstractionInFixBody
- | RecursionNotOnInductiveType of constr
- | RecursionOnIllegalTerm of int * (env * constr) * int list * int list
- | NotEnoughArgumentsForFixCall of int
- (* CoFixpoints *)
- | CodomainNotInductiveType of constr
- | NestedRecursiveOccurrences
- | UnguardedRecursiveCall of constr
- | RecCallInTypeOfAbstraction of constr
- | RecCallInNonRecArgOfConstructor of constr
- | RecCallInTypeOfDef of constr
- | RecCallInCaseFun of constr
- | RecCallInCaseArg of constr
- | RecCallInCasePred of constr
- | NotGuardedForm of constr
- | ReturnPredicateNotCoInductive of constr
-
-type arity_error =
- | NonInformativeToInformative
- | StrongEliminationOnNonSmallType
- | WrongArity
-
-type type_error =
- | UnboundRel of int
- | UnboundVar of variable
- | NotAType of unsafe_judgment
- | BadAssumption of unsafe_judgment
- | ReferenceVariables of constr
- | ElimArity of pinductive * sorts_family list * constr * unsafe_judgment
- * (sorts_family * sorts_family * arity_error) option
- | CaseNotInductive of unsafe_judgment
- | WrongCaseInfo of inductive * case_info
- | NumberBranches of unsafe_judgment * int
- | IllFormedBranch of constr * int * constr * constr
- | Generalization of (Name.t * constr) * unsafe_judgment
- | ActualType of unsafe_judgment * constr
- | CantApplyBadType of
- (int * constr * constr) * unsafe_judgment * unsafe_judgment array
- | CantApplyNonFunctional of unsafe_judgment * unsafe_judgment array
- | IllFormedRecBody of guard_error * Name.t array * int
- | IllTypedRecBody of
- int * Name.t array * unsafe_judgment array * constr array
- | UnsatisfiedConstraints of Univ.constraints
-
-exception TypeError of env * type_error
-
-let nfj (c,ct) = (c, nf_betaiota ct)
-
-let error_unbound_rel env n =
- raise (TypeError (env, UnboundRel n))
-
-let error_unbound_var env v =
- raise (TypeError (env, UnboundVar v))
-
-let error_not_type env j =
- raise (TypeError (env, NotAType j))
-
-let error_assumption env j =
- raise (TypeError (env, BadAssumption j))
-
-let error_reference_variables env id =
- raise (TypeError (env, ReferenceVariables id))
-
-let error_elim_arity env ind aritylst c pj okinds =
- raise (TypeError (env, ElimArity (ind,aritylst,c,pj,okinds)))
-
-let error_case_not_inductive env j =
- raise (TypeError (env, CaseNotInductive j))
-
-let error_number_branches env cj expn =
- raise (TypeError (env, NumberBranches (nfj cj,expn)))
-
-let error_ill_formed_branch env c i actty expty =
- raise (TypeError (env,
- IllFormedBranch (c,i,nf_betaiota actty, nf_betaiota expty)))
-
-let error_actual_type env j expty =
- raise (TypeError (env, ActualType (j,expty)))
-
-let error_cant_apply_not_functional env rator randl =
- raise (TypeError (env, CantApplyNonFunctional (rator,randl)))
-
-let error_cant_apply_bad_type env t rator randl =
- raise (TypeError (env, CantApplyBadType (t,rator,randl)))
-
-let error_ill_formed_rec_body env why lna i =
- raise (TypeError (env, IllFormedRecBody (why,lna,i)))
-
-let error_ill_typed_rec_body env i lna vdefj vargs =
- raise (TypeError (env, IllTypedRecBody (i,lna,vdefj,vargs)))
-
-let error_unsatisfied_constraints env c =
- raise (TypeError (env, UnsatisfiedConstraints c))
diff --git a/checker/type_errors.mli b/checker/type_errors.mli
deleted file mode 100644
index 09703458ac..0000000000
--- a/checker/type_errors.mli
+++ /dev/null
@@ -1,106 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(*i*)
-open Names
-open Cic
-open Environ
-(*i*)
-
-type unsafe_judgment = constr * constr
-
-(* Type errors. \label{typeerrors} *)
-
-(*i Rem: NotEnoughAbstractionInFixBody should only occur with "/i" Fix
- notation i*)
-type guard_error =
- (* Fixpoints *)
- | NotEnoughAbstractionInFixBody
- | RecursionNotOnInductiveType of constr
- | RecursionOnIllegalTerm of int * (env * constr) * int list * int list
- | NotEnoughArgumentsForFixCall of int
- (* CoFixpoints *)
- | CodomainNotInductiveType of constr
- | NestedRecursiveOccurrences
- | UnguardedRecursiveCall of constr
- | RecCallInTypeOfAbstraction of constr
- | RecCallInNonRecArgOfConstructor of constr
- | RecCallInTypeOfDef of constr
- | RecCallInCaseFun of constr
- | RecCallInCaseArg of constr
- | RecCallInCasePred of constr
- | NotGuardedForm of constr
- | ReturnPredicateNotCoInductive of constr
-
-type arity_error =
- | NonInformativeToInformative
- | StrongEliminationOnNonSmallType
- | WrongArity
-
-type type_error =
- | UnboundRel of int
- | UnboundVar of variable
- | NotAType of unsafe_judgment
- | BadAssumption of unsafe_judgment
- | ReferenceVariables of constr
- | ElimArity of pinductive * sorts_family list * constr * unsafe_judgment
- * (sorts_family * sorts_family * arity_error) option
- | CaseNotInductive of unsafe_judgment
- | WrongCaseInfo of inductive * case_info
- | NumberBranches of unsafe_judgment * int
- | IllFormedBranch of constr * int * constr * constr
- | Generalization of (Name.t * constr) * unsafe_judgment
- | ActualType of unsafe_judgment * constr
- | CantApplyBadType of
- (int * constr * constr) * unsafe_judgment * unsafe_judgment array
- | CantApplyNonFunctional of unsafe_judgment * unsafe_judgment array
- | IllFormedRecBody of guard_error * Name.t array * int
- | IllTypedRecBody of
- int * Name.t array * unsafe_judgment array * constr array
- | UnsatisfiedConstraints of Univ.constraints
-
-exception TypeError of env * type_error
-
-val error_unbound_rel : env -> int -> 'a
-
-val error_unbound_var : env -> variable -> 'a
-
-val error_not_type : env -> unsafe_judgment -> 'a
-
-val error_assumption : env -> unsafe_judgment -> 'a
-
-val error_reference_variables : env -> constr -> 'a
-
-val error_elim_arity :
- env -> pinductive -> sorts_family list -> constr -> unsafe_judgment ->
- (sorts_family * sorts_family * arity_error) option -> 'a
-
-val error_case_not_inductive : env -> unsafe_judgment -> 'a
-
-val error_number_branches : env -> unsafe_judgment -> int -> 'a
-
-val error_ill_formed_branch : env -> constr -> int -> constr -> constr -> 'a
-
-val error_actual_type : env -> unsafe_judgment -> constr -> 'a
-
-val error_cant_apply_not_functional :
- env -> unsafe_judgment -> unsafe_judgment array -> 'a
-
-val error_cant_apply_bad_type :
- env -> int * constr * constr ->
- unsafe_judgment -> unsafe_judgment array -> 'a
-
-val error_ill_formed_rec_body :
- env -> guard_error -> Name.t array -> int -> 'a
-
-val error_ill_typed_rec_body :
- env -> int -> Name.t array -> unsafe_judgment array -> constr array -> 'a
-
-val error_unsatisfied_constraints : env -> Univ.constraints -> 'a
diff --git a/checker/typeops.ml b/checker/typeops.ml
deleted file mode 100644
index e4c3f4ae4b..0000000000
--- a/checker/typeops.ml
+++ /dev/null
@@ -1,382 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-open CErrors
-open Util
-open Names
-open Cic
-open Term
-open Reduction
-open Type_errors
-open Inductive
-open Environ
-
-let inductive_of_constructor = fst
-
-let conv_leq_vecti env v1 v2 =
- Array.fold_left2_i
- (fun i _ t1 t2 ->
- (try conv_leq env t1 t2
- with NotConvertible -> raise (NotConvertibleVect i)); ())
- ()
- v1
- v2
-
-let check_constraints cst env =
- if Environ.check_constraints cst env then ()
- else error_unsatisfied_constraints env cst
-
-(* This should be a type (a priori without intension to be an assumption) *)
-let type_judgment env (c,ty as j) =
- match whd_all env ty with
- | Sort s -> (c,s)
- | _ -> error_not_type env j
-
-(* This should be a type intended to be assumed. The error message is *)
-(* not as useful as for [type_judgment]. *)
-let assumption_of_judgment env j =
- try fst(type_judgment env j)
- with TypeError _ ->
- error_assumption env j
-
-(************************************************)
-(* Incremental typing rules: builds a typing judgement given the *)
-(* judgements for the subterms. *)
-
-(*s Type of sorts *)
-
-(* Prop and Set *)
-
-let judge_of_prop = Sort (Type Univ.type1_univ)
-
-(* Type of Type(i). *)
-
-let judge_of_type u = Sort (Type (Univ.super u))
-
-(*s Type of a de Bruijn index. *)
-
-let judge_of_relative env n =
- try
- let LocalAssum (_,typ) | LocalDef (_,_,typ) = lookup_rel n env in
- lift n typ
- with Not_found ->
- error_unbound_rel env n
-
-(* Type of constants *)
-
-let judge_of_constant env (kn,u as cst) =
- let _cb =
- try lookup_constant kn env
- with Not_found ->
- failwith ("Cannot find constant: "^Constant.to_string kn)
- in
- let ty, cu = constant_type env cst in
- let () = check_constraints cu env in
- ty
-
-(* Type of an application. *)
-
-let judge_of_apply env (f,funj) argjv =
- let rec apply_rec n typ = function
- | [] -> typ
- | (h,hj)::restjl ->
- (match whd_all env typ with
- | Prod (_,c1,c2) ->
- (try conv_leq env hj c1
- with NotConvertible ->
- error_cant_apply_bad_type env (n,c1, hj) (f,funj) argjv);
- apply_rec (n+1) (subst1 h c2) restjl
- | _ ->
- error_cant_apply_not_functional env (f,funj) argjv)
- in
- apply_rec 1 funj (Array.to_list argjv)
-
-(* Type of product *)
-
-let sort_of_product env domsort rangsort =
- match (domsort, rangsort) with
- (* Product rule (s,Prop,Prop) *)
- | _, Prop -> rangsort
- (* Product rule (Prop/Set,Set,Set) *)
- | (Prop | Set), Set -> rangsort
- (* Product rule (Type,Set,?) *)
- | Type u1, Set ->
- if engagement env = ImpredicativeSet then
- (* Rule is (Type,Set,Set) in the Set-impredicative calculus *)
- rangsort
- else
- (* Rule is (Type_i,Set,Type_i) in the Set-predicative calculus *)
- Type (Univ.sup u1 Univ.type0_univ)
- (* Product rule (Prop,Type_i,Type_i) *)
- | Set, Type u2 -> Type (Univ.sup Univ.type0_univ u2)
- (* Product rule (Prop,Type_i,Type_i) *)
- | Prop, Type _ -> rangsort
- (* Product rule (Type_i,Type_i,Type_i) *)
- | Type u1, Type u2 -> Type (Univ.sup u1 u2)
-
-(* Type of a type cast *)
-
-(* [judge_of_cast env (c,typ1) (typ2,s)] implements the rule
-
- env |- c:typ1 env |- typ2:s env |- typ1 <= typ2
- ---------------------------------------------------------------------
- env |- c:typ2
-*)
-
-let judge_of_cast env (c,cj) k tj =
- let conversion =
- match k with
- | VMcast | NATIVEcast -> vm_conv CUMUL
- | DEFAULTcast -> conv_leq in
- try
- conversion env cj tj
- with NotConvertible ->
- error_actual_type env (c,cj) tj
-
-(* Inductive types. *)
-
-(* The type is parametric over the uniform parameters whose conclusion
- is in Type; to enforce the internal constraints between the
- parameters and the instances of Type occurring in the type of the
- constructors, we use the level variables _statically_ assigned to
- the conclusions of the parameters as mediators: e.g. if a parameter
- has conclusion Type(alpha), static constraints of the form alpha<=v
- exist between alpha and the Type's occurring in the constructor
- types; when the parameters is finally instantiated by a term of
- conclusion Type(u), then the constraints u<=alpha is computed in
- the App case of execute; from this constraints, the expected
- dynamic constraints of the form u<=v are enforced *)
-
-let judge_of_inductive_knowing_parameters env (ind,u) (paramstyp:constr array) =
- let specif =
- try lookup_mind_specif env ind
- with Not_found ->
- failwith ("Cannot find mutual inductive block: "^MutInd.to_string (fst ind))
- in
- type_of_inductive_knowing_parameters env (specif,u) paramstyp
-
-let judge_of_inductive env ind =
- judge_of_inductive_knowing_parameters env ind [||]
-
-(* Constructors. *)
-
-let judge_of_constructor env (c,u) =
- let ind = inductive_of_constructor c in
- let specif =
- try lookup_mind_specif env ind
- with Not_found ->
- failwith ("Cannot find mutual inductive block: "^MutInd.to_string (fst ind))
- in
- type_of_constructor (c,u) specif
-
-(* Case. *)
-
-let check_branch_types env (c,cj) (lfj,explft) =
- try conv_leq_vecti env lfj explft
- with
- NotConvertibleVect i ->
- error_ill_formed_branch env c i lfj.(i) explft.(i)
- | Invalid_argument _ ->
- error_number_branches env (c,cj) (Array.length explft)
-
-let judge_of_case env ci pj (c,cj) lfj =
- let indspec =
- try find_rectype env cj
- with Not_found -> error_case_not_inductive env (c,cj) in
- let _ = check_case_info env (fst (fst indspec)) ci in
- let (bty,rslty) = type_case_branches env indspec pj c in
- check_branch_types env (c,cj) (lfj,bty);
- rslty
-
-(* Projection. *)
-
-let judge_of_projection env p c ct =
- let pty = lookup_projection p env in
- let (ind,u), args =
- try find_rectype env ct
- with Not_found -> error_case_not_inductive env (c, ct)
- in
- let ty = subst_instance_constr u pty in
- substl (c :: List.rev args) ty
-
-(* Fixpoints. *)
-
-(* Checks the type of a general (co)fixpoint, i.e. without checking *)
-(* the specific guard condition. *)
-
-let type_fixpoint env lna lar lbody vdefj =
- let lt = Array.length vdefj in
- assert (Array.length lar = lt && Array.length lbody = lt);
- try
- conv_leq_vecti env vdefj (Array.map (fun ty -> lift lt ty) lar)
- with NotConvertibleVect i ->
- let vdefj = Array.map2 (fun b ty -> b,ty) lbody vdefj in
- error_ill_typed_rec_body env i lna vdefj lar
-
-(************************************************************************)
-(************************************************************************)
-
-
-(* let refresh_arity env ar = *)
-(* let ctxt, hd = decompose_prod_assum ar in *)
-(* match hd with *)
-(* Sort (Type u) when not (is_univ_variable u) -> *)
-(* let u' = fresh_local_univ() in *)
-(* let env' = add_constraints (enforce_leq u u' empty_constraint) env in *)
-(* env', mkArity (ctxt,Type u') *)
-(* | _ -> env, ar *)
-
-
-(* The typing machine. *)
-let rec execute env cstr =
- match cstr with
- (* Atomic terms *)
- | Sort (Prop | Set) -> judge_of_prop
-
- | Sort (Type u) -> judge_of_type u
-
- | Rel n -> judge_of_relative env n
-
- | Var _ -> anomaly (Pp.str "Section variable in Coqchk!")
-
- | Const c -> judge_of_constant env c
-
- (* Lambda calculus operators *)
- | App (App (f,args),args') ->
- execute env (App(f,Array.append args args'))
-
- | App (f,args) ->
- let jl = execute_array env args in
- let j =
- match f with
- | Ind ind ->
- judge_of_inductive_knowing_parameters env ind jl
- | _ ->
- (* No template polymorphism *)
- execute env f
- in
- let jl = Array.map2 (fun c ty -> c,ty) args jl in
- judge_of_apply env (f,j) jl
-
- | Proj (p, c) ->
- let ct = execute env c in
- judge_of_projection env p c ct
-
- | Lambda (name,c1,c2) ->
- let _ = execute_type env c1 in
- let env1 = push_rel (LocalAssum (name,c1)) env in
- let j' = execute env1 c2 in
- Prod(name,c1,j')
-
- | Prod (name,c1,c2) ->
- let varj = execute_type env c1 in
- let env1 = push_rel (LocalAssum (name,c1)) env in
- let varj' = execute_type env1 c2 in
- Sort (sort_of_product env varj varj')
-
- | LetIn (name,c1,c2,c3) ->
- let j1 = execute env c1 in
- (* /!\ c2 can be an inferred type => refresh
- (but the pushed type is still c2) *)
- let _ =
- let env',c2' = (* refresh_arity env *) env, c2 in
- let _ = execute_type env' c2' in
- judge_of_cast env' (c1,j1) DEFAULTcast c2' in
- let env1 = push_rel (LocalDef (name,c1,c2)) env in
- let j' = execute env1 c3 in
- subst1 c1 j'
-
- | Cast (c,k,t) ->
- let cj = execute env c in
- let _ = execute_type env t in
- judge_of_cast env (c,cj) k t;
- t
-
- (* Inductive types *)
- | Ind ind -> judge_of_inductive env ind
-
- | Construct c -> judge_of_constructor env c
-
- | Case (ci,p,c,lf) ->
- let cj = execute env c in
- let pj = execute env p in
- let lfj = execute_array env lf in
- judge_of_case env ci (p,pj) (c,cj) lfj
- | Fix ((_,i as vni),recdef) ->
- let fix_ty = execute_recdef env recdef i in
- let fix = (vni,recdef) in
- check_fix env fix;
- fix_ty
-
- | CoFix (i,recdef) ->
- let fix_ty = execute_recdef env recdef i in
- let cofix = (i,recdef) in
- check_cofix env cofix;
- fix_ty
-
- (* Partial proofs: unsupported by the kernel *)
- | Meta _ ->
- anomaly (Pp.str "the kernel does not support metavariables.")
-
- | Evar _ ->
- anomaly (Pp.str "the kernel does not support existential variables.")
-
-and execute_type env constr =
- let j = execute env constr in
- snd (type_judgment env (constr,j))
-
-and execute_recdef env (names,lar,vdef) i =
- let larj = execute_array env lar in
- let larj = Array.map2 (fun c ty -> c,ty) lar larj in
- let lara = Array.map (assumption_of_judgment env) larj in
- let env1 = push_rec_types (names,lara,vdef) env in
- let vdefj = execute_array env1 vdef in
- type_fixpoint env1 names lara vdef vdefj;
- lara.(i)
-
-and execute_array env = Array.map (execute env)
-
-(* Derived functions *)
-let infer env constr = execute env constr
-let infer_type env constr = execute_type env constr
-
-(* Typing of several terms. *)
-
-let check_ctxt env rels =
- fold_rel_context (fun d env ->
- match d with
- | LocalAssum (_,ty) ->
- let _ = infer_type env ty in
- push_rel d env
- | LocalDef (_,bd,ty) ->
- let j1 = infer env bd in
- let _ = infer env ty in
- conv_leq env j1 ty;
- push_rel d env)
- rels ~init:env
-
-(* Polymorphic arities utils *)
-
-let check_kind env ar u =
- match (snd (dest_prod env ar)) with
- | Sort (Type u') when Univ.Universe.equal u' (Univ.Universe.make u) -> ()
- | _ -> failwith "not the correct sort"
-
-let check_polymorphic_arity env params par =
- let pl = par.template_param_levels in
- let rec check_p env pl params =
- match pl, params with
- Some u::pl, LocalAssum (na,ty)::params ->
- check_kind env ty u;
- check_p (push_rel (LocalAssum (na,ty)) env) pl params
- | None::pl,d::params -> check_p (push_rel d env) pl params
- | [], _ -> ()
- | _ -> failwith "check_poly: not the right number of params" in
- check_p env pl (List.rev params)
diff --git a/checker/univ.ml b/checker/univ.ml
deleted file mode 100644
index a0511ad21b..0000000000
--- a/checker/univ.ml
+++ /dev/null
@@ -1,1127 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(* Created in Caml by Gérard Huet for CoC 4.8 [Dec 1988] *)
-(* Functional code by Jean-Christophe Filliâtre for Coq V7.0 [1999] *)
-(* Extension with algebraic universes by HH for Coq V7.0 [Sep 2001] *)
-(* Additional support for sort-polymorphic inductive types by HH [Mar 2006] *)
-
-(* Revisions by Bruno Barras, Hugo Herbelin, Pierre Letouzey *)
-
-open Pp
-open CErrors
-open Util
-
-(* Universes are stratified by a partial ordering $\le$.
- Let $\~{}$ be the associated equivalence. We also have a strict ordering
- $<$ between equivalence classes, and we maintain that $<$ is acyclic,
- and contained in $\le$ in the sense that $[U]<[V]$ implies $U\le V$.
-
- At every moment, we have a finite number of universes, and we
- maintain the ordering in the presence of assertions $U<V$ and $U\le V$.
-
- The equivalence $\~{}$ is represented by a tree structure, as in the
- union-find algorithm. The assertions $<$ and $\le$ are represented by
- adjacency lists *)
-
-module RawLevel =
-struct
- open Names
- type t =
- | Prop
- | Set
- | Level of int * DirPath.t
- | Var of int
-
- (* Hash-consing *)
-
- let equal x y =
- x == y ||
- match x, y with
- | Prop, Prop -> true
- | Set, Set -> true
- | Level (n,d), Level (n',d') ->
- Int.equal n n' && DirPath.equal d d'
- | Var n, Var n' -> Int.equal n n'
- | _ -> false
-
- let compare u v =
- match u, v with
- | Prop,Prop -> 0
- | Prop, _ -> -1
- | _, Prop -> 1
- | Set, Set -> 0
- | Set, _ -> -1
- | _, Set -> 1
- | Level (i1, dp1), Level (i2, dp2) ->
- if i1 < i2 then -1
- else if i1 > i2 then 1
- else DirPath.compare dp1 dp2
- | Level _, _ -> -1
- | _, Level _ -> 1
- | Var n, Var m -> Int.compare n m
-
- open Hashset.Combine
-
- let hash = function
- | Prop -> combinesmall 1 0
- | Set -> combinesmall 1 1
- | Var n -> combinesmall 2 n
- | Level (n, d) -> combinesmall 3 (combine n (Names.DirPath.hash d))
-end
-
-module Level = struct
-
- open Names
-
- type raw_level = RawLevel.t =
- | Prop
- | Set
- | Level of int * DirPath.t
- | Var of int
-
- (** Embed levels with their hash value *)
- type t = {
- hash : int;
- data : RawLevel.t }
-
- let equal x y =
- x == y || Int.equal x.hash y.hash && RawLevel.equal x.data y.data
-
- let hash x = x.hash
-
- let data x = x.data
-
- let make l = { hash = RawLevel.hash l; data = l }
-
- let set = make Set
- let prop = make Prop
- let var i = make (Var i)
-
- let is_small x =
- match data x with
- | Level _ -> false
- | _ -> true
-
- let is_prop x =
- match data x with
- | Prop -> true
- | _ -> false
-
- let is_set x =
- match data x with
- | Set -> true
- | _ -> false
-
- let compare u v =
- if u == v then 0
- else
- let c = Int.compare (hash u) (hash v) in
- if c == 0 then RawLevel.compare (data u) (data v)
- else c
-
- let to_string x =
- match data x with
- | Prop -> "Prop"
- | Set -> "Set"
- | Level (n,d) -> Names.DirPath.to_string d^"."^string_of_int n
- | Var i -> "Var("^string_of_int i^")"
-
- let pr u = str (to_string u)
-
- let make m n = make (Level (n, m))
-
-end
-
-(** Level sets and maps *)
-module LMap = HMap.Make (Level)
-module LSet = struct
- include LMap.Set
-
- let pr s =
- str"{" ++ prlist_with_sep spc Level.pr (elements s) ++ str"}"
-
-end
-
-type 'a universe_map = 'a LMap.t
-
-type universe_level = Level.t
-
-type universe_level_subst_fn = universe_level -> universe_level
-
-(* An algebraic universe [universe] is either a universe variable
- [Level.t] or a formal universe known to be greater than some
- universe variables and strictly greater than some (other) universe
- variables
-
- Universes variables denote universes initially present in the term
- to type-check and non variable algebraic universes denote the
- universes inferred while type-checking: it is either the successor
- of a universe present in the initial term to type-check or the
- maximum of two algebraic universes
-*)
-
-module Universe =
-struct
- (* Invariants: non empty, sorted and without duplicates *)
-
- module Expr =
- struct
- type t = Level.t * int
-
- let make l = (l, 0)
-
- let prop = (Level.prop, 0)
- let set = (Level.set, 0)
- let type1 = (Level.set, 1)
-
- let is_prop = function
- | (l,0) -> Level.is_prop l
- | _ -> false
-
- let equal x y = x == y ||
- (let (u,n) = x and (v,n') = y in
- Int.equal n n' && Level.equal u v)
-
- let successor (u,n) =
- if Level.is_prop u then type1
- else (u, n + 1)
-
- let addn k (u,n as x) =
- if k = 0 then x
- else if Level.is_prop u then
- (Level.set,n+k)
- else (u,n+k)
-
- let super (u,n as x) (v,n' as y) =
- let cmp = Level.compare u v in
- if Int.equal cmp 0 then
- if n < n' then Inl true
- else Inl false
- else if is_prop x then Inl true
- else if is_prop y then Inl false
- else Inr cmp
-
- let to_string (v, n) =
- if Int.equal n 0 then Level.to_string v
- else Level.to_string v ^ "+" ^ string_of_int n
-
- let pr x = str(to_string x)
-
- let level = function
- | (v,0) -> Some v
- | _ -> None
-
- let map f (v, n as x) =
- let v' = f v in
- if v' == v then x
- else if Level.is_prop v' && n != 0 then
- (Level.set, n)
- else (v', n)
-
- end
-
- type t = Expr.t list
-
- let tip u = [u]
- let cons u v = u :: v
-
- let equal x y = x == y || List.equal Expr.equal x y
-
- let make l = tip (Expr.make l)
-
- let pr l = match l with
- | [u] -> Expr.pr u
- | _ ->
- str "max(" ++ hov 0
- (prlist_with_sep pr_comma Expr.pr l) ++
- str ")"
-
- let level l = match l with
- | [l] -> Expr.level l
- | _ -> None
-
- (* The lower predicative level of the hierarchy that contains (impredicative)
- Prop and singleton inductive types *)
- let type0m = tip Expr.prop
-
- (* The level of sets *)
- let type0 = tip Expr.set
-
- (* When typing [Prop] and [Set], there is no constraint on the level,
- hence the definition of [type1_univ], the type of [Prop] *)
- let type1 = tip (Expr.successor Expr.set)
-
- let is_type0m x = equal type0m x
- let is_type0 x = equal type0 x
-
- (* Returns the formal universe that lies juste above the universe variable u.
- Used to type the sort u. *)
- let super l =
- List.map (fun x -> Expr.successor x) l
-
- let addn n l =
- List.map (fun x -> Expr.addn n x) l
-
- let rec merge_univs l1 l2 =
- match l1, l2 with
- | [], _ -> l2
- | _, [] -> l1
- | h1 :: t1, h2 :: t2 ->
- (match Expr.super h1 h2 with
- | Inl true (* h1 < h2 *) -> merge_univs t1 l2
- | Inl false -> merge_univs l1 t2
- | Inr c ->
- if c <= 0 (* h1 < h2 is name order *)
- then cons h1 (merge_univs t1 l2)
- else cons h2 (merge_univs l1 t2))
-
- let sort u =
- let rec aux a l =
- match l with
- | b :: l' ->
- (match Expr.super a b with
- | Inl false -> aux a l'
- | Inl true -> l
- | Inr c ->
- if c <= 0 then cons a l
- else cons b (aux a l'))
- | [] -> cons a l
- in
- List.fold_right (fun a acc -> aux a acc) u []
-
- (* Returns the formal universe that is greater than the universes u and v.
- Used to type the products. *)
- let sup x y = merge_univs x y
-
- let empty = []
-
- let exists = List.exists
-
- let for_all = List.for_all
-
- let smart_map = List.Smart.map
-
-end
-
-type universe = Universe.t
-
-(* The level of predicative Set *)
-let type0m_univ = Universe.type0m
-let type0_univ = Universe.type0
-let type1_univ = Universe.type1
-let is_type0m_univ = Universe.is_type0m
-let is_type0_univ = Universe.is_type0
-let is_univ_variable l = Universe.level l != None
-let pr_uni = Universe.pr
-
-let sup = Universe.sup
-let super = Universe.super
-
-open Universe
-
-(* Comparison on this type is pointer equality *)
-type canonical_arc =
- { univ: Level.t;
- lt: Level.t list;
- le: Level.t list;
- rank : int;
- predicative : bool}
-
-let terminal u = {univ=u; lt=[]; le=[]; rank=0; predicative=false}
-
-module UMap :
-sig
- type key = Level.t
- type +'a t
- val empty : 'a t
- val add : key -> 'a -> 'a t -> 'a t
- val find : key -> 'a t -> 'a
- val fold : (key -> 'a -> 'b -> 'b) -> 'a t -> 'b -> 'b
-end = HMap.Make(Level)
-
-(* A Level.t is either an alias for another one, or a canonical one,
- for which we know the universes that are above *)
-
-type univ_entry =
- Canonical of canonical_arc
- | Equiv of Level.t
-
-type universes = univ_entry UMap.t
-
-let enter_equiv_arc u v g =
- UMap.add u (Equiv v) g
-
-let enter_arc ca g =
- UMap.add ca.univ (Canonical ca) g
-
-(* Every Level.t has a unique canonical arc representative *)
-
-(* repr : universes -> Level.t -> canonical_arc *)
-(* canonical representative : we follow the Equiv links *)
-
-let repr g u =
- let rec repr_rec u =
- let a =
- try UMap.find u g
- with Not_found -> anomaly ~label:"Univ.repr"
- (str"Universe " ++ Level.pr u ++ str" undefined.")
- in
- match a with
- | Equiv v -> repr_rec v
- | Canonical arc -> arc
- in
- repr_rec u
-
-let get_set_arc g = repr g Level.set
-
-exception AlreadyDeclared
-
-let add_universe vlev strict g =
- try
- let _arcv = UMap.find vlev g in
- raise AlreadyDeclared
- with Not_found ->
- let v = terminal vlev in
- let arc =
- let arc = get_set_arc g in
- if strict then
- { arc with lt=vlev::arc.lt}
- else
- { arc with le=vlev::arc.le}
- in
- let g = enter_arc arc g in
- enter_arc v g
-
-(* reprleq : canonical_arc -> canonical_arc list *)
-(* All canonical arcv such that arcu<=arcv with arcv#arcu *)
-let reprleq g arcu =
- let rec searchrec w = function
- | [] -> w
- | v :: vl ->
- let arcv = repr g v in
- if List.memq arcv w || arcu==arcv then
- searchrec w vl
- else
- searchrec (arcv :: w) vl
- in
- searchrec [] arcu.le
-
-
-(* between : Level.t -> canonical_arc -> canonical_arc list *)
-(* between u v = { w | u<=w<=v, w canonical } *)
-(* between is the most costly operation *)
-
-let between g arcu arcv =
- (* good are all w | u <= w <= v *)
- (* bad are all w | u <= w ~<= v *)
- (* find good and bad nodes in {w | u <= w} *)
- (* explore b u = (b or "u is good") *)
- let rec explore ((good, bad, b) as input) arcu =
- if List.memq arcu good then
- (good, bad, true) (* b or true *)
- else if List.memq arcu bad then
- input (* (good, bad, b or false) *)
- else
- let leq = reprleq g arcu in
- (* is some universe >= u good ? *)
- let good, bad, b_leq =
- List.fold_left explore (good, bad, false) leq
- in
- if b_leq then
- arcu::good, bad, true (* b or true *)
- else
- good, arcu::bad, b (* b or false *)
- in
- let good,_,_ = explore ([arcv],[],false) arcu in
- good
-
-(* We assume compare(u,v) = LE with v canonical (see compare below).
- In this case List.hd(between g u v) = repr u
- Otherwise, between g u v = []
- *)
-
-type constraint_type = Lt | Le | Eq
-
-let constraint_type_ord c1 c2 = match c1, c2 with
-| Lt, Lt -> 0
-| Lt, _ -> -1
-| Le, Lt -> 1
-| Le, Le -> 0
-| Le, Eq -> -1
-| Eq, Eq -> 0
-| Eq, _ -> 1
-
-(** [compare_neq] : is [arcv] in the transitive upward closure of [arcu] ?
-
- In [strict] mode, we fully distinguish between LE and LT, while in
- non-strict mode, we simply answer LE for both situations.
-
- If [arcv] is encountered in a LT part, we could directly answer
- without visiting unneeded parts of this transitive closure.
- In [strict] mode, if [arcv] is encountered in a LE part, we could only
- change the default answer (1st arg [c]) from NLE to LE, since a strict
- constraint may appear later. During the recursive traversal,
- [lt_done] and [le_done] are universes we have already visited,
- they do not contain [arcv]. The 4rd arg is [(lt_todo,le_todo)],
- two lists of universes not yet considered, known to be above [arcu],
- strictly or not.
-
- We use depth-first search, but the presence of [arcv] in [new_lt]
- is checked as soon as possible : this seems to be slightly faster
- on a test.
-*)
-
-type fast_order = FastEQ | FastLT | FastLE | FastNLE
-
-let fast_compare_neq strict g arcu arcv =
- (* [c] characterizes whether arcv has already been related
- to arcu among the lt_done,le_done universe *)
- let rec cmp c lt_done le_done lt_todo le_todo = match lt_todo, le_todo with
- | [],[] -> c
- | arc::lt_todo, le_todo ->
- if List.memq arc lt_done then
- cmp c lt_done le_done lt_todo le_todo
- else
- let rec find lt_todo lt le = match le with
- | [] ->
- begin match lt with
- | [] -> cmp c (arc :: lt_done) le_done lt_todo le_todo
- | u :: lt ->
- let arc = repr g u in
- if arc == arcv then
- if strict then FastLT else FastLE
- else find (arc :: lt_todo) lt le
- end
- | u :: le ->
- let arc = repr g u in
- if arc == arcv then
- if strict then FastLT else FastLE
- else find (arc :: lt_todo) lt le
- in
- find lt_todo arc.lt arc.le
- | [], arc::le_todo ->
- if arc == arcv then
- (* No need to continue inspecting universes above arc:
- if arcv is strictly above arc, then we would have a cycle.
- But we cannot answer LE yet, a stronger constraint may
- come later from [le_todo]. *)
- if strict then cmp FastLE lt_done le_done [] le_todo else FastLE
- else
- if (List.memq arc lt_done) || (List.memq arc le_done) then
- cmp c lt_done le_done [] le_todo
- else
- let rec find lt_todo lt = match lt with
- | [] ->
- let fold accu u =
- let node = repr g u in
- node :: accu
- in
- let le_new = List.fold_left fold le_todo arc.le in
- cmp c lt_done (arc :: le_done) lt_todo le_new
- | u :: lt ->
- let arc = repr g u in
- if arc == arcv then
- if strict then FastLT else FastLE
- else find (arc :: lt_todo) lt
- in
- find [] arc.lt
- in
- cmp FastNLE [] [] [] [arcu]
-
-let fast_compare g arcu arcv =
- if arcu == arcv then FastEQ else fast_compare_neq true g arcu arcv
-
-let is_leq g arcu arcv =
- arcu == arcv ||
- (match fast_compare_neq false g arcu arcv with
- | FastNLE -> false
- | (FastEQ|FastLE|FastLT) -> true)
-
-let is_lt g arcu arcv =
- if arcu == arcv then false
- else
- match fast_compare_neq true g arcu arcv with
- | FastLT -> true
- | (FastEQ|FastLE|FastNLE) -> false
-
-(* Invariants : compare(u,v) = EQ <=> compare(v,u) = EQ
- compare(u,v) = LT or LE => compare(v,u) = NLE
- compare(u,v) = NLE => compare(v,u) = NLE or LE or LT
-
- Adding u>=v is consistent iff compare(v,u) # LT
- and then it is redundant iff compare(u,v) # NLE
- Adding u>v is consistent iff compare(v,u) = NLE
- and then it is redundant iff compare(u,v) = LT *)
-
-(** * Universe checks [check_eq] and [check_leq], used in coqchk *)
-
-(** First, checks on universe levels *)
-
-let check_equal g u v =
- let arcu = repr g u in
- let arcv = repr g v in
- arcu == arcv
-
-let check_eq_level g u v = u == v || check_equal g u v
-
-let is_set_arc u = Level.is_set u.univ
-let is_prop_arc u = Level.is_prop u.univ
-
-let check_smaller g strict u v =
- let arcu = repr g u in
- let arcv = repr g v in
- if strict then
- is_lt g arcu arcv
- else
- is_prop_arc arcu
- || (is_set_arc arcu && arcv.predicative)
- || is_leq g arcu arcv
-
-(** Then, checks on universes *)
-
-type 'a check_function = universes -> 'a -> 'a -> bool
-
-let check_equal_expr g x y =
- x == y || (let (u, n) = x and (v, m) = y in
- Int.equal n m && check_equal g u v)
-
-let check_eq_univs g l1 l2 =
- let f x1 x2 = check_equal_expr g x1 x2 in
- let exists x1 l = List.exists (fun x2 -> f x1 x2) l in
- List.for_all (fun x1 -> exists x1 l2) l1
- && List.for_all (fun x2 -> exists x2 l1) l2
-
-let check_eq g u v =
- Universe.equal u v || check_eq_univs g u v
-
-let check_smaller_expr g (u,n) (v,m) =
- let diff = n - m in
- match diff with
- | 0 -> check_smaller g false u v
- | 1 -> check_smaller g true u v
- | x when x < 0 -> check_smaller g false u v
- | _ -> false
-
-let exists_bigger g ul l =
- Universe.exists (fun ul' ->
- check_smaller_expr g ul ul') l
-
-let real_check_leq g u v =
- Universe.for_all (fun ul -> exists_bigger g ul v) u
-
-let check_leq g u v =
- Universe.equal u v ||
- Universe.is_type0m u ||
- check_eq_univs g u v || real_check_leq g u v
-
-(** Enforcing new constraints : [setlt], [setleq], [merge], [merge_disc] *)
-
-(** To speed up tests of Set </<= i *)
-let set_predicative g arcv =
- enter_arc {arcv with predicative = true} g
-
-(* setlt : Level.t -> Level.t -> reason -> unit *)
-(* forces u > v *)
-(* this is normally an update of u in g rather than a creation. *)
-let setlt g arcu arcv =
- let arcu' = {arcu with lt=arcv.univ::arcu.lt} in
- let g =
- if is_set_arc arcu then set_predicative g arcv
- else g
- in
- enter_arc arcu' g, arcu'
-
-(* checks that non-redundant *)
-let setlt_if (g,arcu) v =
- let arcv = repr g v in
- if is_lt g arcu arcv then g, arcu
- else setlt g arcu arcv
-
-(* setleq : Level.t -> Level.t -> unit *)
-(* forces u >= v *)
-(* this is normally an update of u in g rather than a creation. *)
-let setleq g arcu arcv =
- let arcu' = {arcu with le=arcv.univ::arcu.le} in
- let g =
- if is_set_arc arcu' then
- set_predicative g arcv
- else g
- in
- enter_arc arcu' g, arcu'
-
-(* checks that non-redundant *)
-let setleq_if (g,arcu) v =
- let arcv = repr g v in
- if is_leq g arcu arcv then g, arcu
- else setleq g arcu arcv
-
-(* merge : Level.t -> Level.t -> unit *)
-(* we assume compare(u,v) = LE *)
-(* merge u v forces u ~ v with repr u as canonical repr *)
-let merge g arcu arcv =
- (* we find the arc with the biggest rank, and we redirect all others to it *)
- let arcu, g, v =
- let best_ranked (max_rank, old_max_rank, best_arc, rest) arc =
- if Level.is_small arc.univ || arc.rank >= max_rank
- then (arc.rank, max_rank, arc, best_arc::rest)
- else (max_rank, old_max_rank, best_arc, arc::rest)
- in
- match between g arcu arcv with
- | [] -> anomaly (str "Univ.between.")
- | arc::rest ->
- let (max_rank, old_max_rank, best_arc, rest) =
- List.fold_left best_ranked (arc.rank, min_int, arc, []) rest in
- if max_rank > old_max_rank then best_arc, g, rest
- else begin
- (* one redirected node also has max_rank *)
- let arcu = {best_arc with rank = max_rank + 1} in
- arcu, enter_arc arcu g, rest
- end
- in
- let redirect (g,w,w') arcv =
- let g' = enter_equiv_arc arcv.univ arcu.univ g in
- (g',List.unionq arcv.lt w,arcv.le@w')
- in
- let (g',w,w') = List.fold_left redirect (g,[],[]) v in
- let g_arcu = (g',arcu) in
- let g_arcu = List.fold_left setlt_if g_arcu w in
- let g_arcu = List.fold_left setleq_if g_arcu w' in
- fst g_arcu
-
-(* merge_disc : Level.t -> Level.t -> unit *)
-(* we assume compare(u,v) = compare(v,u) = NLE *)
-(* merge_disc u v forces u ~ v with repr u as canonical repr *)
-let merge_disc g arc1 arc2 =
- let arcu, arcv = if arc1.rank < arc2.rank then arc2, arc1 else arc1, arc2 in
- let arcu, g =
- if not (Int.equal arc1.rank arc2.rank) then arcu, g
- else
- let arcu = {arcu with rank = succ arcu.rank} in
- arcu, enter_arc arcu g
- in
- let g' = enter_equiv_arc arcv.univ arcu.univ g in
- let g_arcu = (g',arcu) in
- let g_arcu = List.fold_left setlt_if g_arcu arcv.lt in
- let g_arcu = List.fold_left setleq_if g_arcu arcv.le in
- fst g_arcu
-
-(* Universe inconsistency: error raised when trying to enforce a relation
- that would create a cycle in the graph of universes. *)
-
-type univ_inconsistency = constraint_type * universe * universe
-
-exception UniverseInconsistency of univ_inconsistency
-
-let error_inconsistency o u v =
- raise (UniverseInconsistency (o,make u,make v))
-
-(* enforc_univ_eq : Level.t -> Level.t -> unit *)
-(* enforc_univ_eq u v will force u=v if possible, will fail otherwise *)
-
-let enforce_univ_eq u v g =
- let arcu = repr g u in
- let arcv = repr g v in
- match fast_compare g arcu arcv with
- | FastEQ -> g
- | FastLT -> error_inconsistency Eq v u
- | FastLE -> merge g arcu arcv
- | FastNLE ->
- (match fast_compare g arcv arcu with
- | FastLT -> error_inconsistency Eq u v
- | FastLE -> merge g arcv arcu
- | FastNLE -> merge_disc g arcu arcv
- | FastEQ -> anomaly (Pp.str "Univ.compare."))
-
-(* enforce_univ_leq : Level.t -> Level.t -> unit *)
-(* enforce_univ_leq u v will force u<=v if possible, will fail otherwise *)
-let enforce_univ_leq u v g =
- let arcu = repr g u in
- let arcv = repr g v in
- if is_leq g arcu arcv then g
- else
- match fast_compare g arcv arcu with
- | FastLT -> error_inconsistency Le u v
- | FastLE -> merge g arcv arcu
- | FastNLE -> fst (setleq g arcu arcv)
- | FastEQ -> anomaly (Pp.str "Univ.compare.")
-
-(* enforce_univ_lt u v will force u<v if possible, will fail otherwise *)
-let enforce_univ_lt u v g =
- let arcu = repr g u in
- let arcv = repr g v in
- match fast_compare g arcu arcv with
- | FastLT -> g
- | FastLE -> fst (setlt g arcu arcv)
- | FastEQ -> error_inconsistency Lt u v
- | FastNLE ->
- match fast_compare_neq false g arcv arcu with
- FastNLE -> fst (setlt g arcu arcv)
- | FastEQ -> anomaly (Pp.str "Univ.compare.")
- | FastLE | FastLT -> error_inconsistency Lt u v
-
-(* Prop = Set is forbidden here. *)
-let initial_universes =
- let g = enter_arc (terminal Level.set) UMap.empty in
- let g = enter_arc (terminal Level.prop) g in
- enforce_univ_lt Level.prop Level.set g
-
-(* Constraints and sets of constraints. *)
-
-type univ_constraint = Level.t * constraint_type * Level.t
-
-let enforce_constraint cst g =
- match cst with
- | (u,Lt,v) -> enforce_univ_lt u v g
- | (u,Le,v) -> enforce_univ_leq u v g
- | (u,Eq,v) -> enforce_univ_eq u v g
-
-module UConstraintOrd =
-struct
- type t = univ_constraint
- let compare (u,c,v) (u',c',v') =
- let i = constraint_type_ord c c' in
- if not (Int.equal i 0) then i
- else
- let i' = Level.compare u u' in
- if not (Int.equal i' 0) then i'
- else Level.compare v v'
-end
-
-let pr_constraint_type op =
- let op_str = match op with
- | Lt -> " < "
- | Le -> " <= "
- | Eq -> " = "
- in str op_str
-
-module Constraint =
-struct
- module S = Set.Make(UConstraintOrd)
- include S
-
- let pr prl c =
- fold (fun (u1,op,u2) pp_std ->
- pp_std ++ prl u1 ++ pr_constraint_type op ++
- prl u2 ++ fnl () ) c (str "")
-end
-
-let empty_constraint = Constraint.empty
-let merge_constraints c g =
- Constraint.fold enforce_constraint c g
-
-type constraints = Constraint.t
-
-(** A value with universe constraints. *)
-type 'a constrained = 'a * constraints
-
-(** Constraint functions. *)
-
-type 'a constraint_function = 'a -> 'a -> constraints -> constraints
-
-let check_constraint g (l,d,r) =
- match d with
- | Eq -> check_equal g l r
- | Le -> check_smaller g false l r
- | Lt -> check_smaller g true l r
-
-let check_constraints c g =
- Constraint.for_all (check_constraint g) c
-
-(**********************************************************************)
-(** Universe polymorphism *)
-(**********************************************************************)
-
-(** A universe level substitution, note that no algebraic universes are
- involved *)
-
-type universe_level_subst = universe_level universe_map
-
-(** A full substitution might involve algebraic universes *)
-type universe_subst = universe universe_map
-
-module Instance : sig
- type t = Level.t array
-
- val empty : t
- val is_empty : t -> bool
- val equal : t -> t -> bool
- val subst_fn : universe_level_subst_fn -> t -> t
- val subst : universe_level_subst -> t -> t
- val pr : t -> Pp.t
- val check_eq : t check_function
- val length : t -> int
- val append : t -> t -> t
- val of_array : Level.t array -> t
-end =
-struct
- type t = Level.t array
-
- let empty = [||]
-
- let is_empty x = Int.equal (Array.length x) 0
-
- let subst_fn fn t =
- let t' = CArray.Smart.map fn t in
- if t' == t then t else t'
-
- let subst s t =
- let t' =
- CArray.Smart.map (fun x -> try LMap.find x s with Not_found -> x) t
- in if t' == t then t else t'
-
- let pr =
- prvect_with_sep spc Level.pr
-
- let equal t u =
- t == u ||
- (Array.is_empty t && Array.is_empty u) ||
- (CArray.for_all2 Level.equal t u
- (* Necessary as universe instances might come from different modules and
- unmarshalling doesn't preserve sharing *))
-
- let check_eq g t1 t2 =
- t1 == t2 ||
- (Int.equal (Array.length t1) (Array.length t2) &&
- let rec aux i =
- (Int.equal i (Array.length t1)) || (check_eq_level g t1.(i) t2.(i) && aux (i + 1))
- in aux 0)
-
- let length = Array.length
-
- let append = Array.append
-
- let of_array i = i
-
-end
-
-(** Substitute instance inst for ctx in csts *)
-
-let subst_instance_level s l =
- match l.Level.data with
- | Level.Var n -> s.(n)
- | _ -> l
-
-let subst_instance_instance s i =
- Array.Smart.map (fun l -> subst_instance_level s l) i
-
-let subst_instance_universe s u =
- let f x = Universe.Expr.map (fun u -> subst_instance_level s u) x in
- let u' = Universe.smart_map f u in
- if u == u' then u
- else Universe.sort u'
-
-let subst_instance_constraint s (u,d,v as c) =
- let u' = subst_instance_level s u in
- let v' = subst_instance_level s v in
- if u' == u && v' == v then c
- else (u',d,v')
-
-let subst_instance_constraints s csts =
- Constraint.fold
- (fun c csts -> Constraint.add (subst_instance_constraint s c) csts)
- csts Constraint.empty
-
-type universe_instance = Instance.t
-
-type 'a puniverses = 'a * Instance.t
-(** A context of universe levels with universe constraints,
- representiong local universe variables and constraints *)
-
-module UContext =
-struct
- type t = Instance.t constrained
-
- (** Universe contexts (variables as a list) *)
- let empty = (Instance.empty, Constraint.empty)
- let make x = x
- let instance (univs, cst) = univs
- let constraints (univs, cst) = cst
- let size (univs, _) = Instance.length univs
-
- let is_empty (univs, cst) = Instance.is_empty univs && Constraint.is_empty cst
- let pr prl (univs, cst as ctx) =
- if is_empty ctx then mt() else
- h 0 (Instance.pr univs ++ str " |= ") ++ h 0 (v 0 (Constraint.pr prl cst))
-end
-
-type universe_context = UContext.t
-
-module AUContext =
-struct
- include UContext
-
- let repr (inst, cst) =
- (Array.mapi (fun i l -> Level.var i) inst, cst)
-
- let instantiate inst (u, cst) =
- assert (Array.length u = Array.length inst);
- subst_instance_constraints inst cst
-
-end
-
-type abstract_universe_context = AUContext.t
-
-module Variance =
-struct
- (** A universe position in the instance given to a cumulative
- inductive can be the following. Note there is no Contravariant
- case because [forall x : A, B <= forall x : A', B'] requires [A =
- A'] as opposed to [A' <= A]. *)
- type t = Irrelevant | Covariant | Invariant
-
- let check_subtype x y = match x, y with
- | (Irrelevant | Covariant | Invariant), Irrelevant -> true
- | Irrelevant, Covariant -> false
- | (Covariant | Invariant), Covariant -> true
- | (Irrelevant | Covariant), Invariant -> false
- | Invariant, Invariant -> true
-
- let leq_constraint csts variance u u' =
- match variance with
- | Irrelevant -> csts
- | Covariant -> Constraint.add (u, Le, u') csts
- | Invariant -> Constraint.add (u, Eq, u') csts
-
- let eq_constraint csts variance u u' =
- match variance with
- | Irrelevant -> csts
- | Covariant | Invariant -> Constraint.add (u, Eq, u') csts
-
- let leq_constraints variance u u' csts =
- let len = Array.length u in
- assert (len = Array.length u' && len = Array.length variance);
- Array.fold_left3 leq_constraint csts variance u u'
-
- let eq_constraints variance u u' csts =
- let len = Array.length u in
- assert (len = Array.length u' && len = Array.length variance);
- Array.fold_left3 eq_constraint csts variance u u'
-end
-
-module CumulativityInfo =
-struct
- type t = universe_context * Variance.t array
-
- let univ_context (univcst, subtypcst) = univcst
- let variance (univs, variance) = variance
-
-end
-
-module ACumulativityInfo = CumulativityInfo
-type abstract_cumulativity_info = ACumulativityInfo.t
-
-module ContextSet =
-struct
- type t = LSet.t constrained
- let empty = LSet.empty, Constraint.empty
- let constraints (_, cst) = cst
- let levels (ctx, _) = ctx
- let make ctx cst = (ctx, cst)
-end
-type universe_context_set = ContextSet.t
-
-(** Instance subtyping *)
-
-let check_subtype univs ctxT ctx =
- if AUContext.size ctx == AUContext.size ctx then
- let (inst, cst) = AUContext.repr ctx in
- let cstT = UContext.constraints (AUContext.repr ctxT) in
- let push accu v = add_universe v false accu in
- let univs = Array.fold_left push univs inst in
- let univs = merge_constraints cstT univs in
- check_constraints cst univs
- else false
-
-(** Substitutions. *)
-
-let is_empty_subst = LMap.is_empty
-let empty_level_subst = LMap.empty
-let is_empty_level_subst = LMap.is_empty
-
-(** Substitution functions *)
-
-(** With level to level substitutions. *)
-let subst_univs_level_level subst l =
- try LMap.find l subst
- with Not_found -> l
-
-let subst_univs_level_universe subst u =
- let f x = Universe.Expr.map (fun u -> subst_univs_level_level subst u) x in
- let u' = Universe.smart_map f u in
- if u == u' then u
- else Universe.sort u'
-
-let make_abstract_instance (ctx, _) =
- Array.mapi (fun i l -> Level.var i) ctx
-
-(** With level to universe substitutions. *)
-type universe_subst_fn = universe_level -> universe
-
-let make_subst subst = fun l -> LMap.find l subst
-
-let subst_univs_expr_opt fn (l,n) =
- Universe.addn n (fn l)
-
-let subst_univs_universe fn ul =
- let subst, nosubst =
- List.fold_right (fun u (subst,nosubst) ->
- try let a' = subst_univs_expr_opt fn u in
- (a' :: subst, nosubst)
- with Not_found -> (subst, u :: nosubst))
- ul ([], [])
- in
- if CList.is_empty subst then ul
- else
- let substs =
- List.fold_left Universe.merge_univs Universe.empty subst
- in
- List.fold_left (fun acc u -> Universe.merge_univs acc (Universe.tip u))
- substs nosubst
-
-let merge_context strict ctx g =
- let g = Array.fold_left
- (fun g v -> add_universe v strict g)
- g (UContext.instance ctx)
- in merge_constraints (UContext.constraints ctx) g
-
-let merge_context_set strict ctx g =
- let g = LSet.fold
- (* Include and side effects still have double-declared universes *)
- (fun v g -> try add_universe v strict g with AlreadyDeclared -> g)
- (ContextSet.levels ctx) g
- in merge_constraints (ContextSet.constraints ctx) g
-
-(** Pretty-printing *)
-
-let pr_constraints prl = Constraint.pr prl
-
-let pr_universe_context = UContext.pr
-
-let pr_arc = function
- | _, Canonical {univ=u; lt=[]; le=[]} ->
- mt ()
- | _, Canonical {univ=u; lt=lt; le=le} ->
- let opt_sep = match lt, le with
- | [], _ | _, [] -> mt ()
- | _ -> spc ()
- in
- Level.pr u ++ str " " ++
- v 0
- (pr_sequence (fun v -> str "< " ++ Level.pr v) lt ++
- opt_sep ++
- pr_sequence (fun v -> str "<= " ++ Level.pr v) le) ++
- fnl ()
- | u, Equiv v ->
- Level.pr u ++ str " = " ++ Level.pr v ++ fnl ()
-
-let pr_universes g =
- let graph = UMap.fold (fun u a l -> (u,a)::l) g [] in
- prlist pr_arc graph
diff --git a/checker/univ.mli b/checker/univ.mli
deleted file mode 100644
index 3b29b158f2..0000000000
--- a/checker/univ.mli
+++ /dev/null
@@ -1,301 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-(** Universes. *)
-
-module Level :
-sig
- type t
- (** Type of universe levels. A universe level is essentially a unique name
- that will be associated to constraints later on. *)
-
- val make : Names.DirPath.t -> int -> t
- (** Create a new universe level from a unique identifier and an associated
- module path. *)
-
- val var : int -> t
-
- val pr : t -> Pp.t
- (** Pretty-printing *)
-
- val equal : t -> t -> bool
-end
-
-type universe_level = Level.t
-(** Alias name. *)
-
-module Universe :
-sig
- type t
- (** Type of universes. A universe is defined as a set of level expressions.
- A level expression is built from levels and successors of level expressions, i.e.:
- le ::= l + n, n \in N.
-
- A universe is said atomic if it consists of a single level expression with
- no increment, and algebraic otherwise (think the least upper bound of a set of
- level expressions).
- *)
-
- val equal : t -> t -> bool
- (** Equality function on formal universes *)
-
- val make : Level.t -> t
- (** Create a universe representing the given level. *)
-
- val pr : t -> Pp.t
-end
-
-type universe = Universe.t
-
-(** Alias name. *)
-
-val pr_uni : universe -> Pp.t
-
-(** The universes hierarchy: Type 0- = Prop <= Type 0 = Set <= Type 1 <= ...
- Typing of universes: Type 0-, Type 0 : Type 1; Type i : Type (i+1) if i>0 *)
-val type0m_univ : universe
-val type0_univ : universe
-val type1_univ : universe
-
-val is_type0_univ : universe -> bool
-val is_type0m_univ : universe -> bool
-val is_univ_variable : universe -> bool
-
-val sup : universe -> universe -> universe
-val super : universe -> universe
-
-(** {6 Graphs of universes. } *)
-
-type universes
-
-type 'a check_function = universes -> 'a -> 'a -> bool
-val check_leq : universe check_function
-val check_eq : universe check_function
-
-
-
-(** The initial graph of universes: Prop < Set *)
-val initial_universes : universes
-
-(** Adds a universe to the graph, ensuring it is >= or > Set.
- @raise AlreadyDeclared if the level is already declared in the graph. *)
-
-exception AlreadyDeclared
-
-val add_universe : universe_level -> bool -> universes -> universes
-
-(** {6 Constraints. } *)
-
-type constraint_type = Lt | Le | Eq
-type univ_constraint = universe_level * constraint_type * universe_level
-
-module Constraint : Set.S with type elt = univ_constraint
-
-type constraints = Constraint.t
-
-val empty_constraint : constraints
-
-(** A value with universe constraints. *)
-type 'a constrained = 'a * constraints
-
-(** Enforcing constraints. *)
-
-type 'a constraint_function = 'a -> 'a -> constraints -> constraints
-
-(** {6 ... } *)
-(** Merge of constraints in a universes graph.
- The function [merge_constraints] merges a set of constraints in a given
- universes graph. It raises the exception [UniverseInconsistency] if the
- constraints are not satisfiable. *)
-
-(** Type explanation is used to decorate error messages to provide
- useful explanation why a given constraint is rejected. It is composed
- of a path of universes and relation kinds [(r1,u1);..;(rn,un)] means
- .. <(r1) u1 <(r2) ... <(rn) un (where <(ri) is the relation symbol
- denoted by ri, currently only < and <=). The lowest end of the chain
- is supposed known (see UniverseInconsistency exn). The upper end may
- differ from the second univ of UniverseInconsistency because all
- universes in the path are canonical. Note that each step does not
- necessarily correspond to an actual constraint, but reflect how the
- system stores the graph and may result from combination of several
- constraints...
-*)
-type univ_inconsistency = constraint_type * universe * universe
-
-exception UniverseInconsistency of univ_inconsistency
-
-val merge_constraints : constraints -> universes -> universes
-
-val check_constraints : constraints -> universes -> bool
-
-(** {6 Support for universe polymorphism } *)
-
-(** Polymorphic maps from universe levels to 'a *)
-module LMap : CSig.MapS with type key = universe_level
-module LSet :
-sig
- include CSig.SetS with type elt = Level.t
-
- val pr : t -> Pp.t
- (** Pretty-printing *)
-end
-
-type 'a universe_map = 'a LMap.t
-
-(** {6 Substitution} *)
-
-type universe_subst_fn = universe_level -> universe
-type universe_level_subst_fn = universe_level -> universe_level
-
-(** A full substitution, might involve algebraic universes *)
-type universe_subst = universe universe_map
-type universe_level_subst = universe_level universe_map
-
-(** {6 Universe instances} *)
-
-module Instance :
-sig
- type t
- (** A universe instance represents a vector of argument universes
- to a polymorphic definition (constant, inductive or constructor). *)
-
- val empty : t
- val is_empty : t -> bool
-
- val equal : t -> t -> bool
-
- val subst_fn : universe_level_subst_fn -> t -> t
- (** Substitution by a level-to-level function. *)
-
- val subst : universe_level_subst -> t -> t
- (** Substitution by a level-to-level function. *)
-
- val pr : t -> Pp.t
- (** Pretty-printing, no comments *)
-
- val check_eq : t check_function
- (** Check equality of instances w.r.t. a universe graph *)
-
- val length : t -> int
- (** Compute the length of the instance *)
-
- val of_array : Level.t array -> t
-
- val append : t -> t -> t
- (** Append two universe instances *)
-end
-
-type universe_instance = Instance.t
-
-type 'a puniverses = 'a * universe_instance
-
-(** A vector of universe levels with universe constraints,
- representiong local universe variables and associated constraints *)
-
-module UContext :
-sig
- type t
-
- val empty : t
- val make : universe_instance constrained -> t
- val instance : t -> Instance.t
- val constraints : t -> constraints
- val is_empty : t -> bool
-
-end
-
-type universe_context = UContext.t
-
-module AUContext :
-sig
- type t
-
- val size : t -> int
-
- val instantiate : Instance.t -> t -> Constraint.t
- val repr : t -> UContext.t
-
- val pr : (Level.t -> Pp.t) -> t -> Pp.t
-
-end
-
-type abstract_universe_context = AUContext.t
-
-module Variance :
-sig
- (** A universe position in the instance given to a cumulative
- inductive can be the following. Note there is no Contravariant
- case because [forall x : A, B <= forall x : A', B'] requires [A =
- A'] as opposed to [A' <= A]. *)
- type t = Irrelevant | Covariant | Invariant
- val check_subtype : t -> t -> bool
- val leq_constraints : t array -> Instance.t constraint_function
- val eq_constraints : t array -> Instance.t constraint_function
-end
-
-
-module ACumulativityInfo :
-sig
- type t
-
- val univ_context : t -> abstract_universe_context
- val variance : t -> Variance.t array
-
-end
-
-type abstract_cumulativity_info = ACumulativityInfo.t
-
-module ContextSet :
- sig
- type t
- val make : LSet.t -> constraints -> t
- val empty : t
- val constraints : t -> constraints
- end
-
-type universe_context_set = ContextSet.t
-
-val merge_context : bool -> universe_context -> universes -> universes
-val merge_context_set : bool -> universe_context_set -> universes -> universes
-
-val empty_level_subst : universe_level_subst
-val is_empty_level_subst : universe_level_subst -> bool
-
-(** Substitution of universes. *)
-val subst_univs_level_level : universe_level_subst -> universe_level -> universe_level
-val subst_univs_level_universe : universe_level_subst -> universe -> universe
-
-(** Level to universe substitutions. *)
-
-val is_empty_subst : universe_subst -> bool
-val make_subst : universe_subst -> universe_subst_fn
-
-val subst_univs_universe : universe_subst_fn -> universe -> universe
-
-(** Substitution of instances *)
-val subst_instance_instance : universe_instance -> universe_instance -> universe_instance
-val subst_instance_universe : universe_instance -> universe -> universe
-
-(* val make_instance_subst : universe_instance -> universe_level_subst *)
-(* val make_inverse_instance_subst : universe_instance -> universe_level_subst *)
-
-(** Build the relative instance corresponding to the context *)
-val make_abstract_instance : abstract_universe_context -> universe_instance
-
-(** Check instance subtyping *)
-val check_subtype : universes -> AUContext.t -> AUContext.t -> bool
-
-(** {6 Pretty-printing of universes. } *)
-
-val pr_constraint_type : constraint_type -> Pp.t
-val pr_constraints : (Level.t -> Pp.t) -> constraints -> Pp.t
-val pr_universe_context : (Level.t -> Pp.t) -> universe_context -> Pp.t
-
-val pr_universes : universes -> Pp.t
diff --git a/dev/checker_printers.ml b/dev/checker_printers.ml
index 40ae1a7b05..4f89bbd34e 100644
--- a/dev/checker_printers.ml
+++ b/dev/checker_printers.ml
@@ -59,15 +59,11 @@ let pP s = pp (hov 0 s)
let ppuni u = pp(Universe.pr u)
let ppuni_level u = pp (Level.pr u)
-let ppuniverse_set l = pp (LSet.pr l)
-let ppuniverse_instance l = pp (Instance.pr l)
-let ppauniverse_context l = pp (AUContext.pr Level.pr l)
let ppuniverse_context l = pp (pr_universe_context Level.pr l)
let ppconstraints c = pp (pr_constraints Level.pr c)
let ppuniverse_context_future c =
let ctx = Future.force c in
ppuniverse_context ctx
-let ppuniverses u = pp (Univ.pr_universes u)
let pploc x = let (l,r) = Loc.unloc x in
print_string"(";print_int l;print_string",";print_int r;print_string")"
diff --git a/dev/checker_printers.mli b/dev/checker_printers.mli
index 2f9500c5c3..8be9b87257 100644
--- a/dev/checker_printers.mli
+++ b/dev/checker_printers.mli
@@ -43,12 +43,8 @@ val ppididmap : Names.Id.t Names.Id.Map.t -> unit
(* Universes *)
val ppuni : Univ.Universe.t -> unit
val ppuni_level : Univ.Level.t -> unit (* raw *)
-val ppuniverse_set : Univ.LSet.t -> unit
-val ppuniverse_instance : Univ.Instance.t -> unit
-val ppauniverse_context : Univ.AUContext.t -> unit
val ppuniverse_context : Univ.UContext.t -> unit
val ppconstraints : Univ.Constraint.t -> unit
val ppuniverse_context_future : Univ.UContext.t Future.computation -> unit
-val ppuniverses : Univ.universes -> unit
val pploc : Loc.t -> unit
diff --git a/dev/ocamldebug-coq.run b/dev/ocamldebug-coq.run
index 85bb04efe0..d330f517be 100644
--- a/dev/ocamldebug-coq.run
+++ b/dev/ocamldebug-coq.run
@@ -14,40 +14,24 @@
export CAML_LD_LIBRARY_PATH=$COQTOP/kernel/byterun:$CAML_LD_LIBRARY_PATH
-GUESS_CHECKER=
-for arg in "$@"; do
- if [ "${arg##*/}" = coqchk.byte ]; then
- GUESS_CHECKER=1
- fi
-done
-
-if [ -z "$GUESS_CHECKER" ]; then
- exec $OCAMLDEBUG \
- -I $CAMLP5LIB -I +threads \
- -I $COQTOP \
- -I $COQTOP/config -I $COQTOP/printing -I $COQTOP/grammar -I $COQTOP/clib \
- -I $COQTOP/lib -I $COQTOP/kernel -I $COQTOP/kernel/byterun \
- -I $COQTOP/library -I $COQTOP/engine \
- -I $COQTOP/pretyping -I $COQTOP/parsing -I $COQTOP/vernac \
- -I $COQTOP/interp -I $COQTOP/proofs -I $COQTOP/tactics -I $COQTOP/stm \
- -I $COQTOP/toplevel -I $COQTOP/dev -I $COQTOP/config -I $COQTOP/ltac \
- -I $COQTOP/plugins/cc -I $COQTOP/plugins/dp \
- -I $COQTOP/plugins/extraction -I $COQTOP/plugins/field \
- -I $COQTOP/plugins/firstorder \
- -I $COQTOP/plugins/funind -I $COQTOP/plugins/groebner \
- -I $COQTOP/plugins/interface -I $COQTOP/plugins/micromega \
- -I $COQTOP/plugins/omega -I $COQTOP/plugins/quote \
- -I $COQTOP/plugins/ring \
- -I $COQTOP/plugins/rtauto -I $COQTOP/plugins/setoid_ring \
- -I $COQTOP/plugins/subtac -I $COQTOP/plugins/syntax \
- -I $COQTOP/plugins/xml -I $COQTOP/plugins/ltac \
- -I $COQTOP/ide \
- "$@"
-else
- exec $OCAMLDEBUG \
- -I $CAMLP5LIB -I +threads \
- -I $COQTOP \
- -I $COQTOP/config -I $COQTOP/clib \
- -I $COQTOP/lib -I $COQTOP/checker \
- "$@"
-fi
+exec $OCAMLDEBUG \
+ -I $CAMLP5LIB -I +threads \
+ -I $COQTOP \
+ -I $COQTOP/config -I $COQTOP/printing -I $COQTOP/grammar -I $COQTOP/clib \
+ -I $COQTOP/lib -I $COQTOP/kernel -I $COQTOP/kernel/byterun \
+ -I $COQTOP/library -I $COQTOP/engine \
+ -I $COQTOP/pretyping -I $COQTOP/parsing -I $COQTOP/vernac \
+ -I $COQTOP/interp -I $COQTOP/proofs -I $COQTOP/tactics -I $COQTOP/stm \
+ -I $COQTOP/toplevel -I $COQTOP/dev -I $COQTOP/config -I $COQTOP/ltac \
+ -I $COQTOP/plugins/cc -I $COQTOP/plugins/dp \
+ -I $COQTOP/plugins/extraction -I $COQTOP/plugins/field \
+ -I $COQTOP/plugins/firstorder \
+ -I $COQTOP/plugins/funind -I $COQTOP/plugins/groebner \
+ -I $COQTOP/plugins/interface -I $COQTOP/plugins/micromega \
+ -I $COQTOP/plugins/omega -I $COQTOP/plugins/quote \
+ -I $COQTOP/plugins/ring \
+ -I $COQTOP/plugins/rtauto -I $COQTOP/plugins/setoid_ring \
+ -I $COQTOP/plugins/subtac -I $COQTOP/plugins/syntax \
+ -I $COQTOP/plugins/xml -I $COQTOP/plugins/ltac \
+ -I $COQTOP/ide \
+ "$@"
diff --git a/engine/termops.ml b/engine/termops.ml
index 52880846f8..ada6311067 100644
--- a/engine/termops.ml
+++ b/engine/termops.ml
@@ -24,84 +24,13 @@ module CompactedDecl = Context.Compacted.Declaration
module Internal = struct
-(* Sorts and sort family *)
-
-let print_sort = function
- | Set -> (str "Set")
- | Prop -> (str "Prop")
- | Type u -> (str "Type(" ++ Univ.Universe.pr u ++ str ")")
-
-let pr_sort_family = function
- | InSet -> (str "Set")
- | InProp -> (str "Prop")
- | InType -> (str "Type")
-
-let pr_con sp = str(Constant.to_string sp)
-
-let pr_fix pr_constr ((t,i),(lna,tl,bl)) =
- let fixl = Array.mapi (fun i na -> (na,t.(i),tl.(i),bl.(i))) lna in
- hov 1
- (str"fix " ++ int i ++ spc() ++ str"{" ++
- v 0 (prlist_with_sep spc (fun (na,i,ty,bd) ->
- Name.print na ++ str"/" ++ int i ++ str":" ++ pr_constr ty ++
- cut() ++ str":=" ++ pr_constr bd) (Array.to_list fixl)) ++
- str"}")
-
-let pr_puniverses p u =
- if Univ.Instance.is_empty u then p
- else p ++ str"(*" ++ Univ.Instance.pr UnivNames.pr_with_global_universes u ++ str"*)"
-
-(* Minimalistic constr printer, typically for debugging *)
-
-let rec pr_constr c = match kind c with
- | Rel n -> str "#"++int n
- | Meta n -> str "Meta(" ++ int n ++ str ")"
- | Var id -> Id.print id
- | Sort s -> print_sort s
- | Cast (c,_, t) -> hov 1
- (str"(" ++ pr_constr c ++ cut() ++
- str":" ++ pr_constr t ++ str")")
- | Prod (Name(id),t,c) -> hov 1
- (str"forall " ++ Id.print id ++ str":" ++ pr_constr t ++ str"," ++
- spc() ++ pr_constr c)
- | Prod (Anonymous,t,c) -> hov 0
- (str"(" ++ pr_constr t ++ str " ->" ++ spc() ++
- pr_constr c ++ str")")
- | Lambda (na,t,c) -> hov 1
- (str"fun " ++ Name.print na ++ str":" ++
- pr_constr t ++ str" =>" ++ spc() ++ pr_constr c)
- | LetIn (na,b,t,c) -> hov 0
- (str"let " ++ Name.print na ++ str":=" ++ pr_constr b ++
- str":" ++ brk(1,2) ++ pr_constr t ++ cut() ++
- pr_constr c)
- | App (c,l) -> hov 1
- (str"(" ++ pr_constr c ++ spc() ++
- prlist_with_sep spc pr_constr (Array.to_list l) ++ str")")
- | Evar (e,l) -> hov 1
- (str"Evar#" ++ int (Evar.repr e) ++ str"{" ++
- prlist_with_sep spc pr_constr (Array.to_list l) ++str"}")
- | Const (c,u) -> str"Cst(" ++ pr_puniverses (pr_con c) u ++ str")"
- | Ind ((sp,i),u) -> str"Ind(" ++ pr_puniverses (MutInd.print sp ++ str"," ++ int i) u ++ str")"
- | Construct (((sp,i),j),u) ->
- str"Constr(" ++ pr_puniverses (MutInd.print sp ++ str"," ++ int i ++ str"," ++ int j) u ++ str")"
- | Proj (p,c) -> str"Proj(" ++ pr_con (Projection.constant p) ++ str"," ++ bool (Projection.unfolded p) ++ pr_constr c ++ str")"
- | Case (ci,p,c,bl) -> v 0
- (hv 0 (str"<"++pr_constr p++str">"++ cut() ++ str"Case " ++
- pr_constr c ++ str"of") ++ cut() ++
- prlist_with_sep (fun _ -> brk(1,2)) pr_constr (Array.to_list bl) ++
- cut() ++ str"end")
- | Fix f -> pr_fix pr_constr f
- | CoFix(i,(lna,tl,bl)) ->
- let fixl = Array.mapi (fun i na -> (na,tl.(i),bl.(i))) lna in
- hov 1
- (str"cofix " ++ int i ++ spc() ++ str"{" ++
- v 0 (prlist_with_sep spc (fun (na,ty,bd) ->
- Name.print na ++ str":" ++ pr_constr ty ++
- cut() ++ str":=" ++ pr_constr bd) (Array.to_list fixl)) ++
- str"}")
-
-let debug_print_constr c = pr_constr EConstr.Unsafe.(to_constr c)
-let debug_print_constr_env env sigma c = pr_constr EConstr.(to_constr sigma c)
+let pr_sort_family = Sorts.pr_sort_family
+[@@ocaml.deprecated "Use [Sorts.pr_sort_family]"]
+let pr_fix = Constr.debug_print_fix
+[@@ocaml.deprecated "Use [Constr.debug_print_fix]"]
+
+let debug_print_constr c = Constr.debug_print EConstr.Unsafe.(to_constr c)
+let debug_print_constr_env env sigma c = Constr.debug_print EConstr.(to_constr sigma c)
let term_printer = ref debug_print_constr_env
let print_constr_env env sigma t = !term_printer (env:env) sigma (t:Evd.econstr)
diff --git a/engine/termops.mli b/engine/termops.mli
index 07c9541f25..6c3d4fa612 100644
--- a/engine/termops.mli
+++ b/engine/termops.mli
@@ -17,9 +17,10 @@ open Environ
open EConstr
(** printers *)
-val print_sort : Sorts.t -> Pp.t
val pr_sort_family : Sorts.family -> Pp.t
+[@@ocaml.deprecated "Use [Sorts.pr_sort_family]"]
val pr_fix : ('a -> Pp.t) -> ('a, 'a) pfixpoint -> Pp.t
+[@@ocaml.deprecated "Use [Constr.debug_print_fix]"]
(** about contexts *)
val push_rel_assum : Name.t * types -> env -> env
diff --git a/kernel/constr.ml b/kernel/constr.ml
index d7f35da10d..704e6de6b8 100644
--- a/kernel/constr.ml
+++ b/kernel/constr.ml
@@ -1338,3 +1338,70 @@ type compacted_declaration = (constr, types) Context.Compacted.Declaration.pt
type rel_context = rel_declaration list
type named_context = named_declaration list
type compacted_context = compacted_declaration list
+
+(* Sorts and sort family *)
+
+let debug_print_fix pr_constr ((t,i),(lna,tl,bl)) =
+ let open Pp in
+ let fixl = Array.mapi (fun i na -> (na,t.(i),tl.(i),bl.(i))) lna in
+ hov 1
+ (str"fix " ++ int i ++ spc() ++ str"{" ++
+ v 0 (prlist_with_sep spc (fun (na,i,ty,bd) ->
+ Name.print na ++ str"/" ++ int i ++ str":" ++ pr_constr ty ++
+ cut() ++ str":=" ++ pr_constr bd) (Array.to_list fixl)) ++
+ str"}")
+
+let pr_puniverses p u =
+ if Univ.Instance.is_empty u then p
+ else Pp.(p ++ str"(*" ++ Univ.Instance.pr Univ.Level.pr u ++ str"*)")
+
+(* Minimalistic constr printer, typically for debugging *)
+
+let rec debug_print c =
+ let open Pp in
+ match kind c with
+ | Rel n -> str "#"++int n
+ | Meta n -> str "Meta(" ++ int n ++ str ")"
+ | Var id -> Id.print id
+ | Sort s -> Sorts.debug_print s
+ | Cast (c,_, t) -> hov 1
+ (str"(" ++ debug_print c ++ cut() ++
+ str":" ++ debug_print t ++ str")")
+ | Prod (Name(id),t,c) -> hov 1
+ (str"forall " ++ Id.print id ++ str":" ++ debug_print t ++ str"," ++
+ spc() ++ debug_print c)
+ | Prod (Anonymous,t,c) -> hov 0
+ (str"(" ++ debug_print t ++ str " ->" ++ spc() ++
+ debug_print c ++ str")")
+ | Lambda (na,t,c) -> hov 1
+ (str"fun " ++ Name.print na ++ str":" ++
+ debug_print t ++ str" =>" ++ spc() ++ debug_print c)
+ | LetIn (na,b,t,c) -> hov 0
+ (str"let " ++ Name.print na ++ str":=" ++ debug_print b ++
+ str":" ++ brk(1,2) ++ debug_print t ++ cut() ++
+ debug_print c)
+ | App (c,l) -> hov 1
+ (str"(" ++ debug_print c ++ spc() ++
+ prlist_with_sep spc debug_print (Array.to_list l) ++ str")")
+ | Evar (e,l) -> hov 1
+ (str"Evar#" ++ int (Evar.repr e) ++ str"{" ++
+ prlist_with_sep spc debug_print (Array.to_list l) ++str"}")
+ | Const (c,u) -> str"Cst(" ++ pr_puniverses (Constant.debug_print c) u ++ str")"
+ | Ind ((sp,i),u) -> str"Ind(" ++ pr_puniverses (MutInd.print sp ++ str"," ++ int i) u ++ str")"
+ | Construct (((sp,i),j),u) ->
+ str"Constr(" ++ pr_puniverses (MutInd.print sp ++ str"," ++ int i ++ str"," ++ int j) u ++ str")"
+ | Proj (p,c) -> str"Proj(" ++ Constant.debug_print (Projection.constant p) ++ str"," ++ bool (Projection.unfolded p) ++ debug_print c ++ str")"
+ | Case (_ci,p,c,bl) -> v 0
+ (hv 0 (str"<"++debug_print p++str">"++ cut() ++ str"Case " ++
+ debug_print c ++ str"of") ++ cut() ++
+ prlist_with_sep (fun _ -> brk(1,2)) debug_print (Array.to_list bl) ++
+ cut() ++ str"end")
+ | Fix f -> debug_print_fix debug_print f
+ | CoFix(i,(lna,tl,bl)) ->
+ let fixl = Array.mapi (fun i na -> (na,tl.(i),bl.(i))) lna in
+ hov 1
+ (str"cofix " ++ int i ++ spc() ++ str"{" ++
+ v 0 (prlist_with_sep spc (fun (na,ty,bd) ->
+ Name.print na ++ str":" ++ debug_print ty ++
+ cut() ++ str":=" ++ debug_print bd) (Array.to_list fixl)) ++
+ str"}")
diff --git a/kernel/constr.mli b/kernel/constr.mli
index 8753c20eac..1be1f63ff7 100644
--- a/kernel/constr.mli
+++ b/kernel/constr.mli
@@ -590,3 +590,6 @@ val case_info_hash : case_info -> int
(*********************************************************************)
val hcons : constr -> constr
+
+val debug_print : constr -> Pp.t
+val debug_print_fix : ('a -> Pp.t) -> ('a, 'a) pfixpoint -> Pp.t
diff --git a/kernel/indtypes.mli b/kernel/indtypes.mli
index cb09cfa827..a827c17683 100644
--- a/kernel/indtypes.mli
+++ b/kernel/indtypes.mli
@@ -34,6 +34,19 @@ type inductive_error =
exception InductiveError of inductive_error
+val infos_and_sort : env -> constr -> Univ.Universe.t
+
+val check_subtyping_arity_constructor : env -> (constr -> constr) -> types -> int -> bool -> unit
+
+val check_positivity : chkpos:bool ->
+ Names.MutInd.t ->
+ Environ.env ->
+ (Constr.constr, Constr.types) Context.Rel.pt ->
+ Declarations.recursivity_kind ->
+ ('a * Names.Id.t list * Constr.types array *
+ (('b, 'c) Context.Rel.pt * 'd))
+ array -> Int.t * Declarations.recarg Rtree.t array
+
(** The following function does checks on inductive declarations. *)
val check_inductive : env -> MutInd.t -> mutual_inductive_entry -> mutual_inductive_body
diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml
index 4b64cc6d11..8b11851bbb 100644
--- a/kernel/safe_typing.ml
+++ b/kernel/safe_typing.ml
@@ -1061,6 +1061,8 @@ type compiled_library = {
comp_natsymbs : Nativecode.symbols
}
+let module_of_library lib = lib.comp_mod
+
type native_library = Nativecode.global list
let get_library_native_symbols senv dir =
diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli
index 8fb33b04d4..7af773e3bc 100644
--- a/kernel/safe_typing.mli
+++ b/kernel/safe_typing.mli
@@ -141,6 +141,8 @@ val set_share_reduction : bool -> safe_transformer0
val set_VM : bool -> safe_transformer0
val set_native_compiler : bool -> safe_transformer0
+val check_engagement : Environ.env -> Declarations.set_predicativity -> unit
+
(** {6 Interactive module functions } *)
val start_module : Label.t -> ModPath.t safe_transformer
@@ -177,6 +179,8 @@ type compiled_library
type native_library = Nativecode.global list
+val module_of_library : compiled_library -> Declarations.module_body
+
val get_library_native_symbols : safe_environment -> DirPath.t -> Nativecode.symbols
val start_library : DirPath.t -> ModPath.t safe_transformer
diff --git a/kernel/sorts.ml b/kernel/sorts.ml
index a7bb08f5b6..566dce04c6 100644
--- a/kernel/sorts.ml
+++ b/kernel/sorts.ml
@@ -102,3 +102,13 @@ module Hsorts =
end)
let hcons = Hashcons.simple_hcons Hsorts.generate Hsorts.hcons hcons_univ
+
+let debug_print = function
+ | Set -> Pp.(str "Set")
+ | Prop -> Pp.(str "Prop")
+ | Type u -> Pp.(str "Type(" ++ Univ.Universe.pr u ++ str ")")
+
+let pr_sort_family = function
+ | InSet -> Pp.(str "Set")
+ | InProp -> Pp.(str "Prop")
+ | InType -> Pp.(str "Type")
diff --git a/kernel/sorts.mli b/kernel/sorts.mli
index cac6229b91..6c5ce4df80 100644
--- a/kernel/sorts.mli
+++ b/kernel/sorts.mli
@@ -41,3 +41,7 @@ end
val univ_of_sort : t -> Univ.Universe.t
val sort_of_univ : Univ.Universe.t -> t
+
+val debug_print : t -> Pp.t
+
+val pr_sort_family : family -> Pp.t
diff --git a/plugins/funind/g_indfun.mlg b/plugins/funind/g_indfun.mlg
index 857215751a..155df1c1e0 100644
--- a/plugins/funind/g_indfun.mlg
+++ b/plugins/funind/g_indfun.mlg
@@ -197,7 +197,7 @@ END
let pr_fun_scheme_arg (princ_name,fun_name,s) =
Names.Id.print princ_name ++ str " :=" ++ spc() ++ str "Induction for " ++
Libnames.pr_qualid fun_name ++ spc() ++ str "Sort " ++
- Termops.pr_sort_family s
+ Sorts.pr_sort_family s
}
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index 89f64d328a..bd321d5886 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -618,5 +618,5 @@ let lookup_eliminator ind_sp s =
(strbrk "Cannot find the elimination combinator " ++
Id.print id ++ strbrk ", the elimination of the inductive definition " ++
Nametab.pr_global_env Id.Set.empty (IndRef ind_sp) ++
- strbrk " on sort " ++ Termops.pr_sort_family s ++
+ strbrk " on sort " ++ Sorts.pr_sort_family s ++
strbrk " is probably not allowed.")
diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml
index 5d74b59b27..4faa753dfb 100644
--- a/pretyping/recordops.ml
+++ b/pretyping/recordops.ml
@@ -270,7 +270,7 @@ let pr_cs_pattern = function
Const_cs c -> Nametab.pr_global_env Id.Set.empty c
| Prod_cs -> str "_ -> _"
| Default_cs -> str "_"
- | Sort_cs s -> Termops.pr_sort_family s
+ | Sort_cs s -> Sorts.pr_sort_family s
let warn_redundant_canonical_projection =
CWarnings.create ~name:"redundant-canonical-projection" ~category:"typechecker"
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index aced97e910..17003cd1dd 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -353,7 +353,7 @@ struct
| Proj (p,cst) ->
str "ZProj(" ++ Constant.debug_print (Projection.constant p) ++ str ")"
| Fix (f,args,cst) ->
- str "ZFix(" ++ Termops.pr_fix pr_c f
+ str "ZFix(" ++ Constr.debug_print_fix pr_c f
++ pr_comma () ++ pr pr_c args ++ str ")"
| Cst (mem,curr,remains,params,cst_l) ->
str "ZCst(" ++ pr_cst_member pr_c mem ++ pr_comma () ++ int curr
diff --git a/vernac/himsg.ml b/vernac/himsg.ml
index ad6ca3a84e..ba31f73030 100644
--- a/vernac/himsg.ml
+++ b/vernac/himsg.ml
@@ -201,8 +201,8 @@ let explain_elim_arity env sigma ind sorts c pj okinds =
let pc = pr_leconstr_env env sigma c in
let msg = match okinds with
| Some(kp,ki,explanation) ->
- let pki = pr_sort_family ki in
- let pkp = pr_sort_family kp in
+ let pki = Sorts.pr_sort_family ki in
+ let pkp = Sorts.pr_sort_family kp in
let explanation = match explanation with
| NonInformativeToInformative ->
"proofs can be eliminated only to build proofs"
@@ -210,7 +210,7 @@ let explain_elim_arity env sigma ind sorts c pj okinds =
"strong elimination on non-small inductive types leads to paradoxes"
| WrongArity ->
"wrong arity" in
- let ppar = pr_disjunction (fun s -> quote (pr_sort_family s)) sorts in
+ let ppar = pr_disjunction (fun s -> quote (Sorts.pr_sort_family s)) sorts in
let ppt = pr_leconstr_env env sigma (snd (decompose_prod_assum sigma pj.uj_type)) in
hov 0
(str "the return type has sort" ++ spc () ++ ppt ++ spc () ++
diff --git a/vernac/ppvernac.ml b/vernac/ppvernac.ml
index a0e8f38c0b..1c1faca599 100644
--- a/vernac/ppvernac.ml
+++ b/vernac/ppvernac.ml
@@ -312,7 +312,7 @@ open Pputils
) ++
hov 0 ((if dep then keyword "Induction for" else keyword "Minimality for")
++ spc() ++ pr_smart_global ind) ++ spc() ++
- hov 0 (keyword "Sort" ++ spc() ++ Termops.pr_sort_family s)
+ hov 0 (keyword "Sort" ++ spc() ++ Sorts.pr_sort_family s)
| CaseScheme (dep,ind,s) ->
(match idop with
| Some id -> hov 0 (pr_lident id ++ str" :=") ++ spc()
@@ -320,7 +320,7 @@ open Pputils
) ++
hov 0 ((if dep then keyword "Elimination for" else keyword "Case for")
++ spc() ++ pr_smart_global ind) ++ spc() ++
- hov 0 (keyword "Sort" ++ spc() ++ Termops.pr_sort_family s)
+ hov 0 (keyword "Sort" ++ spc() ++ Sorts.pr_sort_family s)
| EqualityScheme ind ->
(match idop with
| Some id -> hov 0 (pr_lident id ++ str" :=") ++ spc()