aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--.gitlab-ci.yml4
-rw-r--r--CHANGES8
-rw-r--r--doc/stdlib/index-list.html.template1
-rw-r--r--engine/evarutil.ml6
-rw-r--r--engine/evd.ml77
-rw-r--r--engine/evd.mli19
-rw-r--r--engine/termops.ml3
-rw-r--r--engine/uState.ml12
-rw-r--r--engine/uState.mli5
-rw-r--r--lib/coqProject_file.ml487
-rw-r--r--lib/coqProject_file.mli42
-rw-r--r--man/coqdep.13
-rw-r--r--plugins/extraction/extract_env.ml60
-rw-r--r--plugins/extraction/extract_env.mli7
-rw-r--r--plugins/extraction/extraction.ml482
-rw-r--r--plugins/extraction/extraction.mli12
-rw-r--r--plugins/extraction/g_extraction.ml46
-rw-r--r--plugins/extraction/table.ml7
-rw-r--r--plugins/ltac/rewrite.ml2
-rw-r--r--plugins/ssr/ssrcommon.ml2
-rw-r--r--plugins/ssr/ssrequality.ml2
-rw-r--r--plugins/ssrmatching/ssrmatching.ml44
-rw-r--r--pretyping/univdecls.ml2
-rw-r--r--printing/prettyp.ml4
-rw-r--r--printing/printer.ml44
-rw-r--r--printing/printer.mli8
-rw-r--r--printing/printmod.ml4
-rw-r--r--proofs/proof_global.ml4
-rw-r--r--stm/stm.ml2
-rw-r--r--tactics/eqschemes.ml12
-rw-r--r--tactics/ind_tables.ml4
-rw-r--r--tactics/leminv.ml2
-rw-r--r--test-suite/output/Arguments_renaming.out2
-rw-r--r--test-suite/output/Existentials.out6
-rw-r--r--test-suite/output/Notations3.out8
-rw-r--r--test-suite/output/inference.out8
-rw-r--r--test-suite/success/Hints.v11
-rw-r--r--test-suite/success/ShowExtraction.v31
-rw-r--r--theories/Compat/Coq87.v1
-rw-r--r--theories/Compat/Coq88.v11
-rw-r--r--theories/Init/Peano.v4
-rw-r--r--tools/CoqMakefile.in9
-rw-r--r--tools/coq_makefile.ml71
-rw-r--r--tools/coqdep.ml19
-rw-r--r--vernac/auto_ind_decl.ml8
-rw-r--r--vernac/comAssumption.ml2
-rw-r--r--vernac/comDefinition.ml2
-rw-r--r--vernac/indschemes.ml2
-rw-r--r--vernac/obligations.ml2
-rw-r--r--vernac/vernacentries.ml6
50 files changed, 699 insertions, 441 deletions
diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml
index 5dd376079..04b75bfdf 100644
--- a/.gitlab-ci.yml
+++ b/.gitlab-ci.yml
@@ -26,7 +26,7 @@ variables:
COQIDE_PACKAGES: "libgtk2.0-dev libgtksourceview2.0-dev"
#COQIDE_PACKAGES_32BIT: "libgtk2.0-dev:i386 libgtksourceview2.0-dev:i386"
COQIDE_OPAM: "lablgtk-extras"
- COQIDE_OPAM_BE: "num lablgtk.2.18.6 lablgtk-extras.1.6"
+ COQIDE_OPAM_BE: "lablgtk.2.18.6 lablgtk-extras.1.6"
COQDOC_PACKAGES: "texlive-latex-base texlive-latex-recommended texlive-latex-extra texlive-math-extra texlive-fonts-recommended texlive-fonts-extra latex-xcolor ghostscript transfig imagemagick tipa"
COQDOC_OPAM: "hevea"
@@ -49,7 +49,7 @@ before_script:
- opam switch ${COMPILER}
- eval $(opam config env)
- opam config list
- - opam install -j ${NJOBS} -y camlp5.${CAMLP5_VER} ocamlfind ${EXTRA_OPAM}
+ - opam install -j ${NJOBS} -y camlp5.${CAMLP5_VER} ocamlfind num ${EXTRA_OPAM}
- rm -rf ~/.opam/log/
- opam list
diff --git a/CHANGES b/CHANGES
index 2b10aad68..1c7c53f29 100644
--- a/CHANGES
+++ b/CHANGES
@@ -75,6 +75,8 @@ Vernacular Commands
- For the Extraction Language command, "OCaml" is spelled correctly.
The older "Ocaml" is still accepted, but deprecated.
- Using “Require” inside a section is deprecated.
+- An experimental command "Show Extraction" allows to extract the content
+ of the current ongoing proof (grant wish #4129).
Universes
@@ -107,6 +109,12 @@ CoqIDE
- Find and Replace All report the number of occurrences found; Find indicates
when it wraps.
+coqdep
+
+- Learned to read -I, -Q, -R and filenames from _CoqProject files.
+ This is used by coq_makefile when generating dependencies for .v
+ files (but not other files).
+
Documentation
- The Coq FAQ, formerly located at https://coq.inria.fr/faq, has been
diff --git a/doc/stdlib/index-list.html.template b/doc/stdlib/index-list.html.template
index a2739e457..8c09b23a5 100644
--- a/doc/stdlib/index-list.html.template
+++ b/doc/stdlib/index-list.html.template
@@ -598,5 +598,6 @@ through the <tt>Require Import</tt> command.</p>
theories/Compat/AdmitAxiom.v
theories/Compat/Coq86.v
theories/Compat/Coq87.v
+ theories/Compat/Coq88.v
</dd>
</dl>
diff --git a/engine/evarutil.ml b/engine/evarutil.ml
index 7674cf67a..6b3ce048f 100644
--- a/engine/evarutil.ml
+++ b/engine/evarutil.ml
@@ -89,15 +89,15 @@ let nf_evars_universes evm =
(Evd.universe_subst evm)
let nf_evars_and_universes evm =
- let evm = Evd.nf_constraints evm in
+ let evm = Evd.minimize_universes evm in
evm, nf_evars_universes evm
let e_nf_evars_and_universes evdref =
- evdref := Evd.nf_constraints !evdref;
+ evdref := Evd.minimize_universes !evdref;
nf_evars_universes !evdref, Evd.universe_subst !evdref
let nf_evar_map_universes evm =
- let evm = Evd.nf_constraints evm in
+ let evm = Evd.minimize_universes evm in
let subst = Evd.universe_subst evm in
if Univ.LMap.is_empty subst then evm, nf_evar0 evm
else
diff --git a/engine/evd.ml b/engine/evd.ml
index e7d542d12..185cab101 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -253,21 +253,8 @@ let instantiate_evar_array info c args =
| [] -> c
| _ -> replace_vars inst c
-type evar_universe_context = UState.t
-type 'a in_evar_universe_context = 'a * evar_universe_context
-
-let empty_evar_universe_context = UState.empty
-let union_evar_universe_context = UState.union
-let evar_universe_context_set = UState.context_set
-let evar_universe_context_constraints = UState.constraints
-let evar_context_universe_context = UState.context
-let evar_universe_context_of = UState.of_context_set
-let evar_universe_context_subst = UState.subst
-let add_constraints_context = UState.add_constraints
-let add_universe_constraints_context = UState.add_universe_constraints
-let constrain_variables = UState.constrain_variables
-let evar_universe_context_of_binders = UState.of_binders
+type 'a in_evar_universe_context = 'a * UState.t
(*******************************************************************)
(* Metamaps *)
@@ -427,7 +414,7 @@ type evar_map = {
undf_evars : evar_info EvMap.t;
evar_names : EvNames.t;
(** Universes *)
- universes : evar_universe_context;
+ universes : UState.t;
(** Conversion problems *)
conv_pbs : evar_constraint list;
last_mods : Evar.Set.t;
@@ -558,10 +545,10 @@ let existential_type d (n, args) =
instantiate_evar_array info info.evar_concl args
let add_constraints d c =
- { d with universes = add_constraints_context d.universes c }
+ { d with universes = UState.add_constraints d.universes c }
let add_universe_constraints d c =
- { d with universes = add_universe_constraints_context d.universes c }
+ { d with universes = UState.add_universe_constraints d.universes c }
(*** /Lifting... ***)
@@ -586,7 +573,7 @@ let create_evar_defs sigma = { sigma with
let empty = {
defn_evars = EvMap.empty;
undf_evars = EvMap.empty;
- universes = empty_evar_universe_context;
+ universes = UState.empty;
conv_pbs = [];
last_mods = Evar.Set.empty;
metas = Metamap.empty;
@@ -609,14 +596,14 @@ let evars_reset_evd ?(with_conv_pbs=false) ?(with_univs=true) evd d =
let last_mods = if with_conv_pbs then evd.last_mods else d.last_mods in
let universes =
if not with_univs then evd.universes
- else union_evar_universe_context evd.universes d.universes
+ else UState.union evd.universes d.universes
in
{ evd with
metas = d.metas;
last_mods; conv_pbs; universes }
let merge_universe_context evd uctx' =
- { evd with universes = union_evar_universe_context evd.universes uctx' }
+ { evd with universes = UState.union evd.universes uctx' }
let set_universe_context evd uctx' =
{ evd with universes = uctx' }
@@ -798,16 +785,6 @@ let make_flexible_variable evd ~algebraic u =
{ evd with universes =
UState.make_flexible_variable evd.universes ~algebraic u }
-let make_evar_universe_context e l =
- let uctx = UState.make (Environ.universes e) in
- match l with
- | None -> uctx
- | Some us ->
- List.fold_left
- (fun uctx { CAst.loc; v = id } ->
- fst (UState.new_univ_variable ?loc univ_rigid (Some id) uctx))
- uctx us
-
(****************************************)
(* Operations on constants *)
(****************************************)
@@ -910,10 +887,6 @@ let check_eq evd s s' =
let check_leq evd s s' =
UGraph.check_leq (UState.ugraph evd.universes) s s'
-let normalize_evar_universe_context_variables = UState.normalize_variables
-
-let abstract_undefined_variables = UState.abstract_undefined_variables
-
let fix_undefined_variables evd =
{ evd with universes = UState.fix_undefined_variables evd.universes }
@@ -922,16 +895,14 @@ let refresh_undefined_universes evd =
let evd' = cmap (subst_univs_level_constr subst) {evd with universes = uctx'} in
evd', subst
-let normalize_evar_universe_context = UState.normalize
-
-let nf_univ_variables evd =
- let subst, uctx' = normalize_evar_universe_context_variables evd.universes in
+let nf_univ_variables evd =
+ let subst, uctx' = UState.normalize_variables evd.universes in
let evd' = {evd with universes = uctx'} in
evd', subst
-let nf_constraints evd =
- let subst, uctx' = normalize_evar_universe_context_variables evd.universes in
- let uctx' = normalize_evar_universe_context uctx' in
+let minimize_universes evd =
+ let subst, uctx' = UState.normalize_variables evd.universes in
+ let uctx' = UState.minimize uctx' in
{evd with universes = uctx'}
let universe_of_name evd s = UState.universe_of_name evd.universes s
@@ -1076,7 +1047,7 @@ let clear_metas evd = {evd with metas = Metamap.empty}
let meta_merge ?(with_univs = true) evd1 evd2 =
let metas = Metamap.fold Metamap.add evd1.metas evd2.metas in
let universes =
- if with_univs then union_evar_universe_context evd2.universes evd1.universes
+ if with_univs then UState.union evd2.universes evd1.universes
else evd2.universes
in
{evd2 with universes; metas; }
@@ -1176,3 +1147,25 @@ module Monad =
(* Failure explanation *)
type unsolvability_explanation = SeveralInstancesFound of int
+
+(** Deprecated *)
+type evar_universe_context = UState.t
+let empty_evar_universe_context = UState.empty
+let union_evar_universe_context = UState.union
+let evar_universe_context_set = UState.context_set
+let evar_universe_context_constraints = UState.constraints
+let evar_context_universe_context = UState.context
+let evar_universe_context_of = UState.of_context_set
+let evar_universe_context_subst = UState.subst
+let add_constraints_context = UState.add_constraints
+let constrain_variables = UState.constrain_variables
+let evar_universe_context_of_binders = UState.of_binders
+let make_evar_universe_context e l =
+ let g = Environ.universes e in
+ match l with
+ | None -> UState.make g
+ | Some l -> UState.make_with_initial_binders g l
+let normalize_evar_universe_context_variables = UState.normalize_variables
+let abstract_undefined_variables = UState.abstract_undefined_variables
+let normalize_evar_universe_context = UState.minimize
+let nf_constraints = minimize_universes
diff --git a/engine/evd.mli b/engine/evd.mli
index ed3316c16..49c953f6d 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -493,22 +493,31 @@ val univ_flexible_alg : rigid
type 'a in_evar_universe_context = 'a * UState.t
val evar_universe_context_set : UState.t -> Univ.ContextSet.t
+[@@ocaml.deprecated "Alias of UState.context_set"]
val evar_universe_context_constraints : UState.t -> Univ.Constraint.t
+[@@ocaml.deprecated "Alias of UState.constraints"]
val evar_context_universe_context : UState.t -> Univ.UContext.t
[@@ocaml.deprecated "alias of UState.context"]
val evar_universe_context_of : Univ.ContextSet.t -> UState.t
+[@@ocaml.deprecated "Alias of UState.of_context_set"]
val empty_evar_universe_context : UState.t
+[@@ocaml.deprecated "Alias of UState.empty"]
val union_evar_universe_context : UState.t -> UState.t ->
UState.t
+[@@ocaml.deprecated "Alias of UState.union"]
val evar_universe_context_subst : UState.t -> Universes.universe_opt_subst
+[@@ocaml.deprecated "Alias of UState.subst"]
val constrain_variables : Univ.LSet.t -> UState.t -> UState.t
+[@@ocaml.deprecated "Alias of UState.constrain_variables"]
val evar_universe_context_of_binders :
Universes.universe_binders -> UState.t
+[@@ocaml.deprecated "Alias of UState.of_binders"]
val make_evar_universe_context : env -> Misctypes.lident list option -> UState.t
+[@@ocaml.deprecated "Use UState.make or UState.make_with_initial_binders"]
val restrict_universe_context : evar_map -> Univ.LSet.t -> evar_map
(** Raises Not_found if not a name for a universe in this map. *)
val universe_of_name : evar_map -> Id.t -> Univ.Level.t
@@ -516,13 +525,15 @@ val universe_of_name : evar_map -> Id.t -> Univ.Level.t
val universe_binders : evar_map -> Universes.universe_binders
val add_constraints_context : UState.t ->
Univ.Constraint.t -> UState.t
+[@@ocaml.deprecated "Alias of UState.add_constraints"]
val normalize_evar_universe_context_variables : UState.t ->
Univ.universe_subst in_evar_universe_context
+[@@ocaml.deprecated "Alias of UState.normalize_variables"]
-val normalize_evar_universe_context : UState.t ->
- UState.t
+val normalize_evar_universe_context : UState.t -> UState.t
+[@@ocaml.deprecated "Alias of UState.minimize"]
val new_univ_level_variable : ?loc:Loc.t -> ?name:Id.t -> rigid -> evar_map -> evar_map * Univ.Level.t
val new_univ_variable : ?loc:Loc.t -> ?name:Id.t -> rigid -> evar_map -> evar_map * Univ.Universe.t
@@ -581,12 +592,16 @@ val with_context_set : ?loc:Loc.t -> rigid -> evar_map -> 'a Univ.in_universe_co
val nf_univ_variables : evar_map -> evar_map * Univ.universe_subst
val abstract_undefined_variables : UState.t -> UState.t
+[@@ocaml.deprecated "Alias of UState.abstract_undefined_variables"]
val fix_undefined_variables : evar_map -> evar_map
val refresh_undefined_universes : evar_map -> evar_map * Univ.universe_level_subst
+(** Universe minimization *)
+val minimize_universes : evar_map -> evar_map
val nf_constraints : evar_map -> evar_map
+[@@ocaml.deprecated "Alias of Evd.minimize_universes"]
val update_sigma_env : evar_map -> env -> evar_map
diff --git a/engine/termops.ml b/engine/termops.ml
index c615155d1..35258762a 100644
--- a/engine/termops.ml
+++ b/engine/termops.ml
@@ -294,12 +294,11 @@ let reference_of_level evd l = UState.reference_of_level (Evd.evar_universe_cont
let pr_evar_universe_context ctx =
let open UState in
- let open Evd in
let prl = pr_uctx_level ctx in
if UState.is_empty ctx then mt ()
else
(str"UNIVERSES:"++brk(0,1)++
- h 0 (Univ.pr_universe_context_set prl (evar_universe_context_set ctx)) ++ fnl () ++
+ h 0 (Univ.pr_universe_context_set prl (UState.context_set ctx)) ++ fnl () ++
str"ALGEBRAIC UNIVERSES:"++brk(0,1)++
h 0 (Univ.LSet.pr prl (UState.algebraics ctx)) ++ fnl() ++
str"UNDEFINED UNIVERSES:"++brk(0,1)++
diff --git a/engine/uState.ml b/engine/uState.ml
index 00825208b..e57afd743 100644
--- a/engine/uState.ml
+++ b/engine/uState.ml
@@ -476,6 +476,13 @@ let new_univ_variable ?loc rigid name
uctx_initial_universes = initial}
in uctx', u
+let make_with_initial_binders e us =
+ let uctx = make e in
+ List.fold_left
+ (fun uctx { CAst.loc; v = id } ->
+ fst (new_univ_variable ?loc univ_rigid (Some id) uctx))
+ uctx us
+
let add_global_univ uctx u =
let initial =
UGraph.add_universe u true uctx.uctx_initial_universes
@@ -578,7 +585,7 @@ let refresh_undefined_univ_variables uctx =
uctx_initial_universes = initial } in
uctx', subst
-let normalize uctx =
+let minimize uctx =
let ((vars',algs'), us') =
Universes.normalize_context_set uctx.uctx_local uctx.uctx_univ_variables
uctx.uctx_univ_algebraic
@@ -606,3 +613,6 @@ let update_sigma_env uctx env =
uctx_universes = univs }
in
merge true univ_rigid eunivs eunivs.uctx_local
+
+(** Deprecated *)
+let normalize = minimize
diff --git a/engine/uState.mli b/engine/uState.mli
index 68fe350c0..9a2bc706b 100644
--- a/engine/uState.mli
+++ b/engine/uState.mli
@@ -26,6 +26,8 @@ val empty : t
val make : UGraph.t -> t
+val make_with_initial_binders : UGraph.t -> Misctypes.lident list -> t
+
val is_empty : t -> bool
val union : t -> t -> t
@@ -131,7 +133,10 @@ val fix_undefined_variables : t -> t
val refresh_undefined_univ_variables : t -> t * Univ.universe_level_subst
+(** Universe minimization *)
+val minimize : t -> t
val normalize : t -> t
+[@@ocaml.deprecated "Alias of UState.minimize"]
type universe_decl =
(Misctypes.lident list, Univ.Constraint.t) Misctypes.gen_universe_decl
diff --git a/lib/coqProject_file.ml4 b/lib/coqProject_file.ml4
index 40945939f..d6c340f69 100644
--- a/lib/coqProject_file.ml4
+++ b/lib/coqProject_file.ml4
@@ -8,27 +8,32 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+type arg_source = CmdLine | ProjectFile
+
+type 'a sourced = { thing : 'a; source : arg_source }
+
type project = {
project_file : string option;
makefile : string option;
install_kind : install option;
use_ocamlopt : bool;
- v_files : string list;
- mli_files : string list;
- ml4_files : string list;
- ml_files : string list;
- mllib_files : string list;
- mlpack_files : string list;
-
- ml_includes : path list;
- r_includes : (path * logic_path) list;
- q_includes : (path * logic_path) list;
- extra_args : string list;
- defs : (string * string) list;
-
- extra_targets : extra_target list;
- subdirs : string list;
+ v_files : string sourced list;
+ mli_files : string sourced list;
+ ml4_files : string sourced list;
+ ml_files : string sourced list;
+ mllib_files : string sourced list;
+ mlpack_files : string sourced list;
+
+ ml_includes : path sourced list;
+ r_includes : (path * logic_path) sourced list;
+ q_includes : (path * logic_path) sourced list;
+ extra_args : string sourced list;
+ defs : (string * string) sourced list;
+
+ (* deprecated in favor of a Makefile.local using :: rules *)
+ extra_targets : extra_target sourced list;
+ subdirs : string sourced list;
}
and extra_target = {
target : string;
@@ -114,6 +119,7 @@ let exists_dir dir =
let process_cmd_line orig_dir proj args =
let parsing_project_file = ref (proj.project_file <> None) in
+ let sourced x = { thing = x; source = if !parsing_project_file then ProjectFile else CmdLine } in
let orig_dir = (* avoids turning foo.v in ./foo.v *)
if orig_dir = "." then "" else orig_dir in
let error s = Format.eprintf "@[%a]@@\n%!" Pp.pp_with Pp.(str (s^".")); exit 1 in
@@ -143,17 +149,17 @@ let process_cmd_line orig_dir proj args =
aux { proj with install_kind = Some install } r
| "-extra" :: target :: dependencies :: command :: r ->
let tgt = { target; dependencies; phony = false; command } in
- aux { proj with extra_targets = proj.extra_targets @ [tgt] } r
+ aux { proj with extra_targets = proj.extra_targets @ [sourced tgt] } r
| "-extra-phony" :: target :: dependencies :: command :: r ->
let tgt = { target; dependencies; phony = true; command } in
- aux { proj with extra_targets = proj.extra_targets @ [tgt] } r
+ aux { proj with extra_targets = proj.extra_targets @ [sourced tgt] } r
| "-Q" :: d :: lp :: r ->
- aux { proj with q_includes = proj.q_includes @ [mk_path d,lp] } r
+ aux { proj with q_includes = proj.q_includes @ [sourced (mk_path d,lp)] } r
| "-I" :: d :: r ->
- aux { proj with ml_includes = proj.ml_includes @ [mk_path d] } r
+ aux { proj with ml_includes = proj.ml_includes @ [sourced (mk_path d)] } r
| "-R" :: d :: lp :: r ->
- aux { proj with r_includes = proj.r_includes @ [mk_path d,lp] } r
+ aux { proj with r_includes = proj.r_includes @ [sourced (mk_path d,lp)] } r
| "-f" :: file :: r ->
if !parsing_project_file then
@@ -178,20 +184,21 @@ let process_cmd_line orig_dir proj args =
error "Option -o given more than once";
aux { proj with makefile = Some file } r
| v :: "=" :: def :: r ->
- aux { proj with defs = proj.defs @ [v,def] } r
+ aux { proj with defs = proj.defs @ [sourced (v,def)] } r
| "-arg" :: a :: r ->
- aux { proj with extra_args = proj.extra_args @ [a] } r
+ aux { proj with extra_args = proj.extra_args @ [sourced a] } r
| f :: r ->
let f = CUnix.correct_path f orig_dir in
let proj =
- if exists_dir f then { proj with subdirs = proj.subdirs @ [f] }
+ if exists_dir f then { proj with subdirs = proj.subdirs @ [sourced f] }
else match CUnix.get_extension f with
- | ".v" -> { proj with v_files = proj.v_files @ [f] }
- | ".ml" -> { proj with ml_files = proj.ml_files @ [f] }
- | ".ml4" -> { proj with ml4_files = proj.ml4_files @ [f] }
- | ".mli" -> { proj with mli_files = proj.mli_files @ [f] }
- | ".mllib" -> { proj with mllib_files = proj.mllib_files @ [f] }
- | ".mlpack" -> { proj with mlpack_files = proj.mlpack_files @ [f] }
+ | ".v" ->
+ { proj with v_files = proj.v_files @ [sourced f] }
+ | ".ml" -> { proj with ml_files = proj.ml_files @ [sourced f] }
+ | ".ml4" -> { proj with ml4_files = proj.ml4_files @ [sourced f] }
+ | ".mli" -> { proj with mli_files = proj.mli_files @ [sourced f] }
+ | ".mllib" -> { proj with mllib_files = proj.mllib_files @ [sourced f] }
+ | ".mlpack" -> { proj with mlpack_files = proj.mlpack_files @ [sourced f] }
| _ -> raise (Parsing_error ("Unknown option "^f)) in
aux proj r
in
@@ -215,16 +222,34 @@ let rec find_project_file ~from ~projfile_name =
else find_project_file ~from:newdir ~projfile_name
;;
+let all_files { v_files; ml_files; mli_files; ml4_files;
+ mllib_files; mlpack_files } =
+ v_files @ mli_files @ ml4_files @ ml_files @ mllib_files @ mlpack_files
+
+let map_sourced_list f l = List.map (fun x -> f x.thing) l
+;;
+
+let map_cmdline f l = CList.map_filter (function
+ | {thing=x; source=CmdLine} -> Some (f x)
+ | {source=ProjectFile} -> None) l
+
let coqtop_args_from_project
{ ml_includes; r_includes; q_includes; extra_args }
=
- let map = List.map in
+ let map = map_sourced_list in
let args =
map (fun { canonical_path = i } -> ["-I"; i]) ml_includes @
map (fun ({ canonical_path = i }, l) -> ["-Q"; i; l]) q_includes @
map (fun ({ canonical_path = p }, l) -> ["-R"; p; l]) r_includes @
- [extra_args] in
+ [map (fun x -> x) extra_args] in
List.flatten args
;;
+let filter_cmdline l = CList.map_filter
+ (function {thing; source=CmdLine} -> Some thing | {source=ProjectFile} -> None)
+ l
+;;
+
+let forget_source {thing} = thing
+
(* vim:set ft=ocaml: *)
diff --git a/lib/coqProject_file.mli b/lib/coqProject_file.mli
index 5a4dd3659..5780bb5d7 100644
--- a/lib/coqProject_file.mli
+++ b/lib/coqProject_file.mli
@@ -10,29 +10,32 @@
exception Parsing_error of string
+type arg_source = CmdLine | ProjectFile
+
+type 'a sourced = { thing : 'a; source : arg_source }
+
type project = {
project_file : string option;
makefile : string option;
install_kind : install option;
use_ocamlopt : bool;
- v_files : string list;
- mli_files : string list;
- ml4_files : string list;
- ml_files : string list;
- mllib_files : string list;
- mlpack_files : string list;
+ v_files : string sourced list;
+ mli_files : string sourced list;
+ ml4_files : string sourced list;
+ ml_files : string sourced list;
+ mllib_files : string sourced list;
+ mlpack_files : string sourced list;
- ml_includes : path list;
- r_includes : (path * logic_path) list;
- q_includes : (path * logic_path) list;
- extra_args : string list;
- defs : (string * string) list;
+ ml_includes : path sourced list;
+ r_includes : (path * logic_path) sourced list;
+ q_includes : (path * logic_path) sourced list;
+ extra_args : string sourced list;
+ defs : (string * string) sourced list;
(* deprecated in favor of a Makefile.local using :: rules *)
- extra_targets : extra_target list;
- subdirs : string list;
-
+ extra_targets : extra_target sourced list;
+ subdirs : string sourced list;
}
and extra_target = {
target : string;
@@ -52,3 +55,14 @@ val read_project_file : string -> project
val coqtop_args_from_project : project -> string list
val find_project_file : from:string -> projfile_name:string -> string option
+val all_files : project -> string sourced list
+
+val map_sourced_list : ('a -> 'b) -> 'a sourced list -> 'b list
+
+(** Only uses the elements with source=CmdLine *)
+val map_cmdline : ('a -> 'b) -> 'a sourced list -> 'b list
+
+(** Only uses the elements with source=CmdLine *)
+val filter_cmdline : 'a sourced list -> 'a list
+
+val forget_source : 'a sourced -> 'a
diff --git a/man/coqdep.1 b/man/coqdep.1
index ed727db7c..c417402c2 100644
--- a/man/coqdep.1
+++ b/man/coqdep.1
@@ -80,6 +80,9 @@ Prints the dependencies of Caml modules.
\" of each Coq file given as argument and complete (if needed)
\" the list of Caml modules. The new command is printed on
\" the standard output. No dependency is computed with this option.
+.TP
+.BI \-f \ file
+Read filenames and options -I, -R and -Q from a _CoqProject FILE.
.TP
.BI \-I/\-Q/\-R \ options
Have the same effects on load path and modules names as for other
diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml
index 722b3990c..a4e8c44cd 100644
--- a/plugins/extraction/extract_env.ml
+++ b/plugins/extraction/extract_env.ml
@@ -137,22 +137,25 @@ let check_arity env cb =
let t = cb.const_type in
if Reduction.is_arity env t then raise Impossible
-let check_fix env cb i =
+let get_body lbody =
+ EConstr.of_constr (Mod_subst.force_constr lbody)
+
+let check_fix env sg cb i =
match cb.const_body with
| Def lbody ->
- (match Constr.kind (Mod_subst.force_constr lbody) with
- | Fix ((_,j),recd) when Int.equal i j -> check_arity env cb; (true,recd)
+ (match EConstr.kind sg (get_body lbody) with
+ | Fix ((_,j),recd) when Int.equal i j -> check_arity env cb; (true,recd)
| CoFix (j,recd) when Int.equal i j -> check_arity env cb; (false,recd)
| _ -> raise Impossible)
| Undef _ | OpaqueDef _ -> raise Impossible
-let prec_declaration_equal (na1, ca1, ta1) (na2, ca2, ta2) =
+let prec_declaration_equal sg (na1, ca1, ta1) (na2, ca2, ta2) =
Array.equal Name.equal na1 na2 &&
- Array.equal Constr.equal ca1 ca2 &&
- Array.equal Constr.equal ta1 ta2
+ Array.equal (EConstr.eq_constr sg) ca1 ca2 &&
+ Array.equal (EConstr.eq_constr sg) ta1 ta2
-let factor_fix env l cb msb =
- let _,recd as check = check_fix env cb 0 in
+let factor_fix env sg l cb msb =
+ let _,recd as check = check_fix env sg cb 0 in
let n = Array.length (let fi,_,_ = recd in fi) in
if Int.equal n 1 then [|l|], recd, msb
else begin
@@ -163,9 +166,9 @@ let factor_fix env l cb msb =
(fun j ->
function
| (l,SFBconst cb') ->
- let check' = check_fix env cb' (j+1) in
- if not ((fst check : bool) == (fst check') &&
- prec_declaration_equal (snd check) (snd check'))
+ let check' = check_fix env sg cb' (j+1) in
+ if not ((fst check : bool) == (fst check') &&
+ prec_declaration_equal sg (snd check) (snd check'))
then raise Impossible;
labels.(j+1) <- l;
| _ -> raise Impossible) msb';
@@ -248,7 +251,9 @@ and extract_mexpr_spec env mp1 (me_struct_o,me_alg) = match me_alg with
let me_struct,delta = flatten_modtype env mp1 me' me_struct_o in
let env' = env_for_mtb_with_def env mp1 me_struct delta idl in
let mt = extract_mexpr_spec env mp1 (None,me') in
- (match extract_with_type env' c with (* cb may contain some kn *)
+ let sg = Evd.from_env env in
+ (match extract_with_type env' sg (EConstr.of_constr c) with
+ (* cb may contain some kn *)
| None -> mt
| Some (vl,typ) ->
type_iter_references Visit.add_ref typ;
@@ -299,12 +304,13 @@ let rec extract_structure env mp reso ~all = function
| [] -> []
| (l,SFBconst cb) :: struc ->
(try
- let vl,recd,struc = factor_fix env l cb struc in
+ let sg = Evd.from_env env in
+ let vl,recd,struc = factor_fix env sg l cb struc in
let vc = Array.map (make_cst reso mp) vl in
let ms = extract_structure env mp reso ~all struc in
let b = Array.exists Visit.needed_cst vc in
if all || b then
- let d = extract_fixpoint env vc recd in
+ let d = extract_fixpoint env sg vc recd in
if (not b) && (logical_decl d) then ms
else begin Visit.add_decl_deps d; (l,SEdecl d) :: ms end
else ms
@@ -572,8 +578,8 @@ let print_structure_to_file (fn,si,mo) dry struc =
let reset () =
Visit.reset (); reset_tables (); reset_renaming_tables Everything
-let init ?(compute=false) modular library =
- check_inside_section (); check_inside_module ();
+let init ?(compute=false) ?(inner=false) modular library =
+ if not inner then (check_inside_section (); check_inside_module ());
set_keywords (descr ()).keywords;
set_modular modular;
set_library library;
@@ -701,10 +707,9 @@ let flatten_structure struc =
and flatten_elems l = List.flatten (List.map flatten_elem l)
in flatten_elems (List.flatten (List.map snd struc))
-let structure_for_compute c =
+let structure_for_compute env sg c =
init false false ~compute:true;
- let env = Global.env () in
- let ast, mlt = Extraction.extract_constr env c in
+ let ast, mlt = Extraction.extract_constr env sg c in
let ast = Mlutil.normalize ast in
let refs = ref Refset.empty in
let add_ref r = refs := Refset.add r !refs in
@@ -744,3 +749,20 @@ let extract_and_compile l =
let base = Filename.chop_suffix f ".ml" in
let () = remove (base^".cmo"); remove (base^".cmi") in
Feedback.msg_notice (str "Extracted code successfully compiled")
+
+(* Show the extraction of the current ongoing proof *)
+
+let show_extraction () =
+ init ~inner:true false false;
+ let prf = Proof_global.give_me_the_proof () in
+ let sigma, env = Pfedit.get_current_context () in
+ let trms = Proof.partial_proof prf in
+ let extr_term t =
+ let ast, ty = extract_constr env sigma t in
+ let mp = Lib.current_mp () in
+ let l = Label.of_id (Proof_global.get_current_proof_name ()) in
+ let fake_ref = ConstRef (Constant.make2 mp l) in
+ let decl = Dterm (fake_ref, ast, ty) in
+ print_one_decl [] mp decl
+ in
+ Feedback.msg_notice (Pp.prlist_with_sep Pp.fnl extr_term trms)
diff --git a/plugins/extraction/extract_env.mli b/plugins/extraction/extract_env.mli
index 464f109be..591d3bb86 100644
--- a/plugins/extraction/extract_env.mli
+++ b/plugins/extraction/extract_env.mli
@@ -36,4 +36,9 @@ val print_one_decl :
(* Used by Extraction Compute *)
val structure_for_compute :
- Constr.t -> (Miniml.ml_decl list) * Miniml.ml_ast * Miniml.ml_type
+ Environ.env -> Evd.evar_map -> EConstr.t ->
+ Miniml.ml_decl list * Miniml.ml_ast * Miniml.ml_type
+
+(* Show the extraction of the current ongoing proof *)
+
+val show_extraction : unit -> unit
diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml
index ce4970056..f25f63624 100644
--- a/plugins/extraction/extraction.ml
+++ b/plugins/extraction/extraction.ml
@@ -13,7 +13,6 @@ open Util
open Names
open Term
open Constr
-open Vars
open Declarations
open Declareops
open Environ
@@ -36,20 +35,18 @@ exception I of inductive_kind
(* A set of all fixpoint functions currently being extracted *)
let current_fixpoints = ref ([] : Constant.t list)
-let none = Evd.empty
-
(* NB: In OCaml, [type_of] and [get_of] might raise
[SingletonInductiveBecomeProp]. This exception will be caught
in late wrappers around the exported functions of this file,
in order to display the location of the issue. *)
-let type_of env c =
+let type_of env sg c =
let polyprop = (lang() == Haskell) in
- EConstr.Unsafe.to_constr (Retyping.get_type_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c)))
+ Retyping.get_type_of ~polyprop env sg (strip_outer_cast sg c)
-let sort_of env c =
+let sort_of env sg c =
let polyprop = (lang() == Haskell) in
- Retyping.get_sort_family_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c))
+ Retyping.get_sort_family_of ~polyprop env sg (strip_outer_cast sg c)
(*S Generation of flags and signatures. *)
@@ -73,61 +70,91 @@ type scheme = TypeScheme | Default
type flag = info * scheme
-let whd_all env t =
- EConstr.Unsafe.to_constr (whd_all env none (EConstr.of_constr t))
-
-let whd_betaiotazeta t =
- EConstr.Unsafe.to_constr (whd_betaiotazeta none (EConstr.of_constr t))
-
(*s [flag_of_type] transforms a type [t] into a [flag].
Really important function. *)
-let rec flag_of_type env t : flag =
- let t = whd_all env t in
- match Constr.kind t with
- | Prod (x,t,c) -> flag_of_type (push_rel (LocalAssum (x,t)) env) c
- | Sort s when Sorts.is_prop s -> (Logic,TypeScheme)
+let rec flag_of_type env sg t : flag =
+ let t = whd_all env sg t in
+ match EConstr.kind sg t with
+ | Prod (x,t,c) -> flag_of_type (EConstr.push_rel (LocalAssum (x,t)) env) sg c
+ | Sort s when Sorts.is_prop (EConstr.ESorts.kind sg s) -> (Logic,TypeScheme)
| Sort _ -> (Info,TypeScheme)
- | _ -> if (sort_of env t) == InProp then (Logic,Default) else (Info,Default)
+ | _ -> if (sort_of env sg t) == InProp then (Logic,Default) else (Info,Default)
(*s Two particular cases of [flag_of_type]. *)
-let is_default env t = match flag_of_type env t with
+let is_default env sg t = match flag_of_type env sg t with
| (Info, Default) -> true
| _ -> false
exception NotDefault of kill_reason
-let check_default env t =
- match flag_of_type env t with
+let check_default env sg t =
+ match flag_of_type env sg t with
| _,TypeScheme -> raise (NotDefault Ktype)
| Logic,_ -> raise (NotDefault Kprop)
| _ -> ()
-let is_info_scheme env t = match flag_of_type env t with
+let is_info_scheme env sg t = match flag_of_type env sg t with
| (Info, TypeScheme) -> true
| _ -> false
let push_rel_assum (n, t) env =
- Environ.push_rel (LocalAssum (n, t)) env
+ EConstr.push_rel (LocalAssum (n, t)) env
+
+let push_rels_assum assums =
+ EConstr.push_rel_context (List.map (fun (x,t) -> LocalAssum (x,t)) assums)
+
+let get_body lconstr = EConstr.of_constr (Mod_subst.force_constr lconstr)
+
+let get_opaque env c =
+ EConstr.of_constr
+ (Opaqueproof.force_proof (Environ.opaque_tables env) c)
+
+let applistc c args = EConstr.mkApp (c, Array.of_list args)
+
+(* Same as [Environ.push_rec_types], but for [EConstr.t] *)
+let push_rec_types (lna,typarray,_) env =
+ let ctxt =
+ Array.map2_i
+ (fun i na t -> LocalAssum (na, EConstr.Vars.lift i t)) lna typarray
+ in
+ Array.fold_left (fun e assum -> EConstr.push_rel assum e) env ctxt
+
+(* Same as [Termops.nb_lam], but for [EConstr.t] *)
+let nb_lam sg c = List.length (fst (EConstr.decompose_lam sg c))
+
+(* Same as [Term.decompose_lam_n] but for [EConstr.t] *)
+let decompose_lam_n sg n =
+ let rec lamdec_rec l n c =
+ if n <= 0 then l,c
+ else match EConstr.kind sg c with
+ | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) (n-1) c
+ | Cast (c,_,_) -> lamdec_rec l n c
+ | _ -> raise Not_found
+ in
+ lamdec_rec [] n
(*s [type_sign] gernerates a signature aimed at treating a type application. *)
-let rec type_sign env c =
- match Constr.kind (whd_all env c) with
+let rec type_sign env sg c =
+ match EConstr.kind sg (whd_all env sg c) with
| Prod (n,t,d) ->
- (if is_info_scheme env t then Keep else Kill Kprop)
- :: (type_sign (push_rel_assum (n,t) env) d)
+ (if is_info_scheme env sg t then Keep else Kill Kprop)
+ :: (type_sign (push_rel_assum (n,t) env) sg d)
| _ -> []
-let rec type_scheme_nb_args env c =
- match Constr.kind (whd_all env c) with
+let rec type_scheme_nb_args env sg c =
+ match EConstr.kind sg (whd_all env sg c) with
| Prod (n,t,d) ->
- let n = type_scheme_nb_args (push_rel_assum (n,t) env) d in
- if is_info_scheme env t then n+1 else n
+ let n = type_scheme_nb_args (push_rel_assum (n,t) env) sg d in
+ if is_info_scheme env sg t then n+1 else n
| _ -> 0
-let _ = Hook.set type_scheme_nb_args_hook type_scheme_nb_args
+let type_scheme_nb_args' env c =
+ type_scheme_nb_args env (Evd.from_env env) (EConstr.of_constr c)
+
+let _ = Hook.set type_scheme_nb_args_hook type_scheme_nb_args'
(*s [type_sign_vl] does the same, plus a type var list. *)
@@ -147,19 +174,19 @@ let make_typvar n vl =
let vl = Id.Set.of_list vl in
next_ident_away id' vl
-let rec type_sign_vl env c =
- match Constr.kind (whd_all env c) with
+let rec type_sign_vl env sg c =
+ match EConstr.kind sg (whd_all env sg c) with
| Prod (n,t,d) ->
- let s,vl = type_sign_vl (push_rel_assum (n,t) env) d in
- if not (is_info_scheme env t) then Kill Kprop::s, vl
- else Keep::s, (make_typvar n vl) :: vl
+ let s,vl = type_sign_vl (push_rel_assum (n,t) env) sg d in
+ if not (is_info_scheme env sg t) then Kill Kprop::s, vl
+ else Keep::s, (make_typvar n vl) :: vl
| _ -> [],[]
-let rec nb_default_params env c =
- match Constr.kind (whd_all env c) with
+let rec nb_default_params env sg c =
+ match EConstr.kind sg (whd_all env sg c) with
| Prod (n,t,d) ->
- let n = nb_default_params (push_rel_assum (n,t) env) d in
- if is_default env t then n+1 else n
+ let n = nb_default_params (push_rel_assum (n,t) env) sg d in
+ if is_default env sg t then n+1 else n
| _ -> 0
(* Enriching a signature with implicit information *)
@@ -226,62 +253,62 @@ let parse_ind_args si args relmax =
generate ML type var anymore (in subterms for example). *)
-let rec extract_type env db j c args =
- match Constr.kind (whd_betaiotazeta c) with
+let rec extract_type env sg db j c args =
+ match EConstr.kind sg (whd_betaiotazeta sg c) with
| App (d, args') ->
- (* We just accumulate the arguments. *)
- extract_type env db j d (Array.to_list args' @ args)
+ (* We just accumulate the arguments. *)
+ extract_type env sg db j d (Array.to_list args' @ args)
| Lambda (_,_,d) ->
(match args with
| [] -> assert false (* A lambda cannot be a type. *)
- | a :: args -> extract_type env db j (subst1 a d) args)
+ | a :: args -> extract_type env sg db j (EConstr.Vars.subst1 a d) args)
| Prod (n,t,d) ->
assert (List.is_empty args);
let env' = push_rel_assum (n,t) env in
- (match flag_of_type env t with
+ (match flag_of_type env sg t with
| (Info, Default) ->
(* Standard case: two [extract_type] ... *)
- let mld = extract_type env' (0::db) j d [] in
+ let mld = extract_type env' sg (0::db) j d [] in
(match expand env mld with
| Tdummy d -> Tdummy d
- | _ -> Tarr (extract_type env db 0 t [], mld))
+ | _ -> Tarr (extract_type env sg db 0 t [], mld))
| (Info, TypeScheme) when j > 0 ->
(* A new type var. *)
- let mld = extract_type env' (j::db) (j+1) d [] in
+ let mld = extract_type env' sg (j::db) (j+1) d [] in
(match expand env mld with
| Tdummy d -> Tdummy d
| _ -> Tarr (Tdummy Ktype, mld))
| _,lvl ->
- let mld = extract_type env' (0::db) j d [] in
+ let mld = extract_type env' sg (0::db) j d [] in
(match expand env mld with
| Tdummy d -> Tdummy d
| _ ->
let reason = if lvl == TypeScheme then Ktype else Kprop in
Tarr (Tdummy reason, mld)))
| Sort _ -> Tdummy Ktype (* The two logical cases. *)
- | _ when sort_of env (applistc c args) == InProp -> Tdummy Kprop
+ | _ when sort_of env sg (applistc c args) == InProp -> Tdummy Kprop
| Rel n ->
- (match lookup_rel n env with
- | LocalDef (_,t,_) -> extract_type env db j (lift n t) args
+ (match EConstr.lookup_rel n env with
+ | LocalDef (_,t,_) ->
+ extract_type env sg db j (EConstr.Vars.lift n t) args
| _ ->
(* Asks [db] a translation for [n]. *)
if n > List.length db then Tunknown
else let n' = List.nth db (n-1) in
if Int.equal n' 0 then Tunknown else Tvar n')
- | Const (kn,u as c) ->
- let r = ConstRef kn in
- let cb = lookup_constant kn env in
- let typ = Typeops.type_of_constant_in env c in
- (match flag_of_type env typ with
+ | Const (kn,u) ->
+ let r = ConstRef kn in
+ let typ = type_of env sg (EConstr.mkConstU (kn,u)) in
+ (match flag_of_type env sg typ with
| (Logic,_) -> assert false (* Cf. logical cases above *)
| (Info, TypeScheme) ->
- let mlt = extract_type_app env db (r, type_sign env typ) args in
- (match cb.const_body with
+ let mlt = extract_type_app env sg db (r, type_sign env sg typ) args in
+ (match (lookup_constant kn env).const_body with
| Undef _ | OpaqueDef _ -> mlt
- | Def _ when is_custom r -> mlt
+ | Def _ when is_custom (ConstRef kn) -> mlt
| Def lbody ->
- let newc = applistc (Mod_subst.force_constr lbody) args in
- let mlt' = extract_type env db j newc [] in
+ let newc = applistc (get_body lbody) args in
+ let mlt' = extract_type env sg db j newc [] in
(* ML type abbreviations interact badly with Coq *)
(* reduction, so [mlt] and [mlt'] might be different: *)
(* The more precise is [mlt'], extracted after reduction *)
@@ -290,36 +317,51 @@ let rec extract_type env db j c args =
if eq_ml_type (expand env mlt) (expand env mlt') then mlt else mlt')
| (Info, Default) ->
(* Not an ML type, for example [(c:forall X, X->X) Type nat] *)
- (match cb.const_body with
+ (match (lookup_constant kn env).const_body with
| Undef _ | OpaqueDef _ -> Tunknown (* Brutal approx ... *)
| Def lbody ->
(* We try to reduce. *)
- let newc = applistc (Mod_subst.force_constr lbody) args in
- extract_type env db j newc []))
+ let newc = applistc (get_body lbody) args in
+ extract_type env sg db j newc []))
| Ind ((kn,i),u) ->
- let s = (extract_ind env kn).ind_packets.(i).ip_sign in
- extract_type_app env db (IndRef (kn,i),s) args
+ let s = (extract_ind env kn).ind_packets.(i).ip_sign in
+ extract_type_app env sg db (IndRef (kn,i),s) args
| Proj (p,t) ->
(* Let's try to reduce, if it hasn't already been done. *)
if Projection.unfolded p then Tunknown
- else extract_type env db j (mkProj (Projection.unfold p, t)) args
+ else
+ extract_type env sg db j (EConstr.mkProj (Projection.unfold p, t)) args
| Case _ | Fix _ | CoFix _ -> Tunknown
- | Var _ | Meta _ | Evar _ | Cast _ | LetIn _ | Construct _ -> assert false
+ | Evar _ | Meta _ -> Taxiom (* only possible during Show Extraction *)
+ | Var v ->
+ (* For Show Extraction *)
+ let open Context.Named.Declaration in
+ (match EConstr.lookup_named v env with
+ | LocalDef (_,body,_) ->
+ extract_type env sg db j (EConstr.applist (body,args)) []
+ | LocalAssum (_,ty) ->
+ let r = VarRef v in
+ (match flag_of_type env sg ty with
+ | (Logic,_) -> assert false (* Cf. logical cases above *)
+ | (Info, TypeScheme) ->
+ extract_type_app env sg db (r, type_sign env sg ty) args
+ | (Info, Default) -> Tunknown))
+ | Cast _ | LetIn _ | Construct _ -> assert false
(*s Auxiliary function dealing with type application.
Precondition: [r] is a type scheme represented by the signature [s],
and is completely applied: [List.length args = List.length s]. *)
-and extract_type_app env db (r,s) args =
+and extract_type_app env sg db (r,s) args =
let ml_args =
List.fold_right
(fun (b,c) a -> if b == Keep then
- let p = List.length (fst (splay_prod env none (EConstr.of_constr (type_of env c)))) in
+ let p = List.length (fst (splay_prod env sg (type_of env sg c))) in
let db = iterate (fun l -> 0 :: l) p db in
- (extract_type_scheme env db c p) :: a
+ (extract_type_scheme env sg db c p) :: a
else a)
(List.combine s args) []
- in Tglob (r, ml_args)
+ in Tglob (r, ml_args)
(*S Extraction of a type scheme. *)
@@ -330,19 +372,18 @@ and extract_type_app env db (r,s) args =
(* [db] is a context for translating Coq [Rel] into ML type [Tvar]. *)
-and extract_type_scheme env db c p =
- if Int.equal p 0 then extract_type env db 0 c []
+and extract_type_scheme env sg db c p =
+ if Int.equal p 0 then extract_type env sg db 0 c []
else
- let c = whd_betaiotazeta c in
- match Constr.kind c with
+ let c = whd_betaiotazeta sg c in
+ match EConstr.kind sg c with
| Lambda (n,t,d) ->
- extract_type_scheme (push_rel_assum (n,t) env) db d (p-1)
+ extract_type_scheme (push_rel_assum (n,t) env) sg db d (p-1)
| _ ->
- let rels = fst (splay_prod env none (EConstr.of_constr (type_of env c))) in
- let rels = List.map (on_snd EConstr.Unsafe.to_constr) rels in
+ let rels = fst (splay_prod env sg (type_of env sg c)) in
let env = push_rels_assum rels env in
- let eta_args = List.rev_map mkRel (List.interval 1 p) in
- extract_type env db 0 (lift p c) eta_args
+ let eta_args = List.rev_map EConstr.mkRel (List.interval 1 p) in
+ extract_type env sg db 0 (EConstr.Vars.lift p c) eta_args
(*S Extraction of an inductive type. *)
@@ -384,6 +425,7 @@ and extract_really_ind env kn mib =
let mip0 = mib.mind_packets.(0) in
let npar = mib.mind_nparams in
let epar = push_rel_context mib.mind_params_ctxt env in
+ let sg = Evd.from_env env in
(* First pass: we store inductive signatures together with *)
(* their type var list. *)
let packets =
@@ -391,8 +433,9 @@ and extract_really_ind env kn mib =
(fun i mip ->
let (_,u),_ = Universes.fresh_inductive_instance env (kn,i) in
let ar = Inductive.type_of_inductive env ((mib,mip),u) in
- let info = (fst (flag_of_type env ar) = Info) in
- let s,v = if info then type_sign_vl env ar else [],[] in
+ let ar = EConstr.of_constr ar in
+ let info = (fst (flag_of_type env sg ar) = Info) in
+ let s,v = if info then type_sign_vl env sg ar else [],[] in
let t = Array.make (Array.length mip.mind_nf_lc) [] in
{ ip_typename = mip.mind_typename;
ip_consnames = mip.mind_consnames;
@@ -424,7 +467,8 @@ and extract_really_ind env kn mib =
in
let dbmap = parse_ind_args p.ip_sign args (nprods + npar) in
let db = db_from_ind dbmap npar in
- p.ip_types.(j) <- extract_type_cons epar db dbmap t (npar+1)
+ p.ip_types.(j) <-
+ extract_type_cons epar sg db dbmap (EConstr.of_constr t) (npar+1)
done
done;
(* Third pass: we determine special cases. *)
@@ -477,10 +521,9 @@ and extract_really_ind env kn mib =
(* Is this record officially declared with its projections ? *)
(* If so, we use this information. *)
begin try
- let n = nb_default_params env
- (Inductive.type_of_inductive env ((mib,mip0),u))
- in
- let check_proj kn = if Cset.mem kn !projs then add_projection n kn ip
+ let ty = Inductive.type_of_inductive env ((mib,mip0),u) in
+ let n = nb_default_params env sg (EConstr.of_constr ty) in
+ let check_proj kn = if Cset.mem kn !projs then add_projection n kn ip
in
List.iter (Option.iter check_proj) (lookup_projections ip)
with Not_found -> ()
@@ -505,13 +548,13 @@ and extract_really_ind env kn mib =
- [i] is the rank of the current product (initially [params_nb+1])
*)
-and extract_type_cons env db dbmap c i =
- match Constr.kind (whd_all env c) with
+and extract_type_cons env sg db dbmap c i =
+ match EConstr.kind sg (whd_all env sg c) with
| Prod (n,t,d) ->
let env' = push_rel_assum (n,t) env in
let db' = (try Int.Map.find i dbmap with Not_found -> 0) :: db in
- let l = extract_type_cons env' db' dbmap d (i+1) in
- (extract_type env db 0 t []) :: l
+ let l = extract_type_cons env' sg db' dbmap d (i+1) in
+ (extract_type env sg db 0 t []) :: l
| _ -> []
(*s Recording the ML type abbreviation of a Coq type scheme constant. *)
@@ -526,16 +569,17 @@ and mlt_env env r = match r with
match lookup_typedef kn cb with
| Some _ as o -> o
| None ->
- let typ = cb.const_type
+ let sg = Evd.from_env env in
+ let typ = EConstr.of_constr cb.const_type
(* FIXME not sure if we should instantiate univs here *) in
- match flag_of_type env typ with
- | Info,TypeScheme ->
- let body = Mod_subst.force_constr l_body in
- let s = type_sign env typ in
- let db = db_from_sign s in
- let t = extract_type_scheme env db body (List.length s)
- in add_typedef kn cb t; Some t
- | _ -> None
+ match flag_of_type env sg typ with
+ | Info,TypeScheme ->
+ let body = get_body l_body in
+ let s = type_sign env sg typ in
+ let db = db_from_sign s in
+ let t = extract_type_scheme env sg db body (List.length s)
+ in add_typedef kn cb t; Some t
+ | _ -> None
and expand env = type_expand (mlt_env env)
and type2signature env = type_to_signature (mlt_env env)
@@ -545,16 +589,16 @@ let type_expunge_from_sign env = type_expunge_from_sign (mlt_env env)
(*s Extraction of the type of a constant. *)
-let record_constant_type env kn opt_typ =
+let record_constant_type env sg kn opt_typ =
let cb = lookup_constant kn env in
match lookup_cst_type kn cb with
| Some schema -> schema
| None ->
let typ = match opt_typ with
- | None -> cb.const_type
+ | None -> EConstr.of_constr cb.const_type
| Some typ -> typ
in
- let mlt = extract_type env [] 1 typ [] in
+ let mlt = extract_type env sg [] 1 typ [] in
let schema = (type_maxvar mlt, mlt) in
let () = add_cst_type kn cb schema in
schema
@@ -566,75 +610,86 @@ let record_constant_type env kn opt_typ =
(* [mle] is a ML environment [Mlenv.t]. *)
(* [mlt] is the ML type we want our extraction of [(c args)] to have. *)
-let rec extract_term env mle mlt c args =
- match Constr.kind c with
+let rec extract_term env sg mle mlt c args =
+ match EConstr.kind sg c with
| App (f,a) ->
- extract_term env mle mlt f (Array.to_list a @ args)
+ extract_term env sg mle mlt f (Array.to_list a @ args)
| Lambda (n, t, d) ->
let id = id_of_name n in
(match args with
| a :: l ->
(* We make as many [LetIn] as possible. *)
- let d' = mkLetIn (Name id,a,t,applistc d (List.map (lift 1) l))
- in extract_term env mle mlt d' []
+ let l' = List.map (EConstr.Vars.lift 1) l in
+ let d' = EConstr.mkLetIn (Name id,a,t,applistc d l') in
+ extract_term env sg mle mlt d' []
| [] ->
let env' = push_rel_assum (Name id, t) env in
let id, a =
- try check_default env t; Id id, new_meta()
- with NotDefault d -> Dummy, Tdummy d
+ try check_default env sg t; Id id, new_meta()
+ with NotDefault d -> Dummy, Tdummy d
in
let b = new_meta () in
(* If [mlt] cannot be unified with an arrow type, then magic! *)
let magic = needs_magic (mlt, Tarr (a, b)) in
- let d' = extract_term env' (Mlenv.push_type mle a) b d [] in
+ let d' = extract_term env' sg (Mlenv.push_type mle a) b d [] in
put_magic_if magic (MLlam (id, d')))
| LetIn (n, c1, t1, c2) ->
let id = id_of_name n in
- let env' = push_rel (LocalDef (Name id, c1, t1)) env in
+ let env' = EConstr.push_rel (LocalDef (Name id, c1, t1)) env in
(* We directly push the args inside the [LetIn].
TODO: the opt_let_app flag is supposed to prevent that *)
- let args' = List.map (lift 1) args in
+ let args' = List.map (EConstr.Vars.lift 1) args in
(try
- check_default env t1;
+ check_default env sg t1;
let a = new_meta () in
- let c1' = extract_term env mle a c1 [] in
+ let c1' = extract_term env sg mle a c1 [] in
(* The type of [c1'] is generalized and stored in [mle]. *)
let mle' =
if generalizable c1'
then Mlenv.push_gen mle a
else Mlenv.push_type mle a
in
- MLletin (Id id, c1', extract_term env' mle' mlt c2 args')
+ MLletin (Id id, c1', extract_term env' sg mle' mlt c2 args')
with NotDefault d ->
let mle' = Mlenv.push_std_type mle (Tdummy d) in
- ast_pop (extract_term env' mle' mlt c2 args'))
+ ast_pop (extract_term env' sg mle' mlt c2 args'))
| Const (kn,_) ->
- extract_cst_app env mle mlt kn args
+ extract_cst_app env sg mle mlt kn args
| Construct (cp,_) ->
- extract_cons_app env mle mlt cp args
+ extract_cons_app env sg mle mlt cp args
| Proj (p, c) ->
- let term = Retyping.expand_projection env (Evd.from_env env) p (EConstr.of_constr c) [] in
- let term = EConstr.Unsafe.to_constr term in
- extract_term env mle mlt term args
+ let term = Retyping.expand_projection env (Evd.from_env env) p c [] in
+ extract_term env sg mle mlt term args
| Rel n ->
(* As soon as the expected [mlt] for the head is known, *)
(* we unify it with an fresh copy of the stored type of [Rel n]. *)
let extract_rel mlt = put_magic (mlt, Mlenv.get mle n) (MLrel n)
- in extract_app env mle mlt extract_rel args
+ in extract_app env sg mle mlt extract_rel args
| Case ({ci_ind=ip},_,c0,br) ->
- extract_app env mle mlt (extract_case env mle (ip,c0,br)) args
+ extract_app env sg mle mlt (extract_case env sg mle (ip,c0,br)) args
| Fix ((_,i),recd) ->
- extract_app env mle mlt (extract_fix env mle i recd) args
+ extract_app env sg mle mlt (extract_fix env sg mle i recd) args
| CoFix (i,recd) ->
- extract_app env mle mlt (extract_fix env mle i recd) args
- | Cast (c,_,_) -> extract_term env mle mlt c args
- | Ind _ | Prod _ | Sort _ | Meta _ | Evar _ | Var _ -> assert false
+ extract_app env sg mle mlt (extract_fix env sg mle i recd) args
+ | Cast (c,_,_) -> extract_term env sg mle mlt c args
+ | Evar _ | Meta _ -> MLaxiom
+ | Var v ->
+ (* Only during Show Extraction *)
+ let open Context.Named.Declaration in
+ let ty = match EConstr.lookup_named v env with
+ | LocalAssum (_,ty) -> ty
+ | LocalDef (_,_,ty) -> ty
+ in
+ let vty = extract_type env sg [] 0 ty [] in
+ let extract_var mlt = put_magic (mlt,vty) (MLglob (VarRef v)) in
+ extract_app env sg mle mlt extract_var args
+ | Ind _ | Prod _ | Sort _ -> assert false
(*s [extract_maybe_term] is [extract_term] for usual terms, else [MLdummy] *)
-and extract_maybe_term env mle mlt c =
- try check_default env (type_of env c);
- extract_term env mle mlt c []
+and extract_maybe_term env sg mle mlt c =
+ try check_default env sg (type_of env sg c);
+ extract_term env sg mle mlt c []
with NotDefault d ->
put_magic (mlt, Tdummy d) (MLdummy d)
@@ -644,28 +699,28 @@ and extract_maybe_term env mle mlt c =
This gives us the expected type of the head. Then we use the
[mk_head] to produce the ML head from this type. *)
-and extract_app env mle mlt mk_head args =
+and extract_app env sg mle mlt mk_head args =
let metas = List.map new_meta args in
let type_head = type_recomp (metas, mlt) in
- let mlargs = List.map2 (extract_maybe_term env mle) metas args in
+ let mlargs = List.map2 (extract_maybe_term env sg mle) metas args in
mlapp (mk_head type_head) mlargs
(*s Auxiliary function used to extract arguments of constant or constructor. *)
-and make_mlargs env e s args typs =
+and make_mlargs env sg e s args typs =
let rec f = function
| [], [], _ -> []
- | a::la, t::lt, [] -> extract_maybe_term env e t a :: (f (la,lt,[]))
- | a::la, t::lt, Keep::s -> extract_maybe_term env e t a :: (f (la,lt,s))
+ | a::la, t::lt, [] -> extract_maybe_term env sg e t a :: (f (la,lt,[]))
+ | a::la, t::lt, Keep::s -> extract_maybe_term env sg e t a :: (f (la,lt,s))
| _::la, _::lt, _::s -> f (la,lt,s)
| _ -> assert false
in f (args,typs,s)
(*s Extraction of a constant applied to arguments. *)
-and extract_cst_app env mle mlt kn args =
+and extract_cst_app env sg mle mlt kn args =
(* First, the [ml_schema] of the constant, in expanded version. *)
- let nb,t = record_constant_type env kn None in
+ let nb,t = record_constant_type env sg kn None in
let schema = nb, expand env t in
(* Can we instantiate types variables for this constant ? *)
(* In Ocaml, inside the definition of this constant, the answer is no. *)
@@ -691,7 +746,7 @@ and extract_cst_app env mle mlt kn args =
let ls = List.length s in
let la = List.length args in
(* The ml arguments, already expunged from known logical ones *)
- let mla = make_mlargs env mle s args metas in
+ let mla = make_mlargs env sg mle s args metas in
let mla =
if magic1 || lang () != Ocaml then mla
else
@@ -736,7 +791,7 @@ and extract_cst_app env mle mlt kn args =
they are fixed, and thus are not used for the computation.
\end{itemize} *)
-and extract_cons_app env mle mlt (((kn,i) as ip,j) as cp) args =
+and extract_cons_app env sg mle mlt (((kn,i) as ip,j) as cp) args =
(* First, we build the type of the constructor, stored in small pieces. *)
let mi = extract_ind env kn in
let params_nb = mi.ind_nparams in
@@ -777,7 +832,7 @@ and extract_cons_app env mle mlt (((kn,i) as ip,j) as cp) args =
put_magic_if magic2
(dummy_lams (anonym_or_dummy_lams head' s) (params_nb - la))
else
- let mla = make_mlargs env mle s args' metas in
+ let mla = make_mlargs env sg mle s args' metas in
if Int.equal la (ls + params_nb)
then put_magic_if (magic2 && not magic1) (head mla)
else (* [ params_nb <= la <= ls + params_nb ] *)
@@ -788,7 +843,7 @@ and extract_cons_app env mle mlt (((kn,i) as ip,j) as cp) args =
(*S Extraction of a case. *)
-and extract_case env mle ((kn,i) as ip,c,br) mlt =
+and extract_case env sg mle ((kn,i) as ip,c,br) mlt =
(* [br]: bodies of each branch (in functional form) *)
(* [ni]: number of arguments without parameters in each branch *)
let ni = constructors_nrealargs_env env ip in
@@ -799,9 +854,9 @@ and extract_case env mle ((kn,i) as ip,c,br) mlt =
MLexn "absurd case"
end else
(* [c] has an inductive type, and is not a type scheme type. *)
- let t = type_of env c in
+ let t = type_of env sg c in
(* The only non-informative case: [c] is of sort [Prop] *)
- if (sort_of env t) == InProp then
+ if (sort_of env sg t) == InProp then
begin
add_recursors env kn; (* May have passed unseen if logical ... *)
(* Logical singleton case: *)
@@ -809,7 +864,7 @@ and extract_case env mle ((kn,i) as ip,c,br) mlt =
assert (Int.equal br_size 1);
let s = iterate (fun l -> Kill Kprop :: l) ni.(0) [] in
let mlt = iterate (fun t -> Tarr (Tdummy Kprop, t)) ni.(0) mlt in
- let e = extract_maybe_term env mle mlt br.(0) in
+ let e = extract_maybe_term env sg mle mlt br.(0) in
snd (case_expunge s e)
end
else
@@ -818,7 +873,7 @@ and extract_case env mle ((kn,i) as ip,c,br) mlt =
let metas = Array.init (List.length oi.ip_vars) new_meta in
(* The extraction of the head. *)
let type_head = Tglob (IndRef ip, Array.to_list metas) in
- let a = extract_term env mle type_head c [] in
+ let a = extract_term env sg mle type_head c [] in
(* The extraction of each branch. *)
let extract_branch i =
let r = ConstructRef (ip,i+1) in
@@ -829,7 +884,7 @@ and extract_case env mle ((kn,i) as ip,c,br) mlt =
let s = List.map (type2sign env) oi.ip_types.(i) in
let s = sign_with_implicits r s mi.ind_nparams in
(* Extraction of the branch (in functional form). *)
- let e = extract_maybe_term env mle (type_recomp (l,mlt)) br.(i) in
+ let e = extract_maybe_term env sg mle (type_recomp (l,mlt)) br.(i) in
(* We suppress dummy arguments according to signature. *)
let ids,e = case_expunge s e in
(List.rev ids, Pusual r, e)
@@ -851,12 +906,12 @@ and extract_case env mle ((kn,i) as ip,c,br) mlt =
(*s Extraction of a (co)-fixpoint. *)
-and extract_fix env mle i (fi,ti,ci as recd) mlt =
+and extract_fix env sg mle i (fi,ti,ci as recd) mlt =
let env = push_rec_types recd env in
let metas = Array.map new_meta fi in
metas.(i) <- mlt;
let mle = Array.fold_left Mlenv.push_type mle metas in
- let ei = Array.map2 (extract_maybe_term env mle) metas ci in
+ let ei = Array.map2 (extract_maybe_term env sg mle) metas ci in
MLfix (i, Array.map id_of_name fi, ei)
(*S ML declarations. *)
@@ -864,34 +919,34 @@ and extract_fix env mle i (fi,ti,ci as recd) mlt =
(* [decomp_lams_eta env c t] finds the number [n] of products in the type [t],
and decompose the term [c] in [n] lambdas, with eta-expansion if needed. *)
-let decomp_lams_eta_n n m env c t =
- let rels = fst (splay_prod_n env none n (EConstr.of_constr t)) in
- let rels = List.map (fun (LocalAssum (id,c) | LocalDef (id,_,c)) -> (id,EConstr.Unsafe.to_constr c)) rels in
- let rels',c = decompose_lam c in
+let decomp_lams_eta_n n m env sg c t =
+ let rels = fst (splay_prod_n env sg n t) in
+ let rels = List.map (fun (LocalAssum (id,c) | LocalDef (id,_,c)) -> (id,c)) rels in
+ let rels',c = EConstr.decompose_lam sg c in
let d = n - m in
(* we'd better keep rels' as long as possible. *)
let rels = (List.firstn d rels) @ rels' in
- let eta_args = List.rev_map mkRel (List.interval 1 d) in
- rels, applistc (lift d c) eta_args
+ let eta_args = List.rev_map EConstr.mkRel (List.interval 1 d) in
+ rels, applistc (EConstr.Vars.lift d c) eta_args
(* Let's try to identify some situation where extracted code
will allow generalisation of type variables *)
-let rec gentypvar_ok c = match Constr.kind c with
+let rec gentypvar_ok sg c = match EConstr.kind sg c with
| Lambda _ | Const _ -> true
| App (c,v) ->
(* if all arguments are variables, these variables will
disappear after extraction (see [empty_s] below) *)
- Array.for_all isRel v && gentypvar_ok c
- | Cast (c,_,_) -> gentypvar_ok c
+ Array.for_all (EConstr.isRel sg) v && gentypvar_ok sg c
+ | Cast (c,_,_) -> gentypvar_ok sg c
| _ -> false
(*s From a constant to a ML declaration. *)
-let extract_std_constant env kn body typ =
+let extract_std_constant env sg kn body typ =
reset_meta_count ();
(* The short type [t] (i.e. possibly with abbreviations). *)
- let t = snd (record_constant_type env kn (Some typ)) in
+ let t = snd (record_constant_type env sg kn (Some typ)) in
(* The real type [t']: without head products, expanded, *)
(* and with [Tvar] translated to [Tvar'] (not instantiable). *)
let l,t' = type_decomp (expand env (var2var' t)) in
@@ -906,14 +961,14 @@ let extract_std_constant env kn body typ =
break user's clever let-ins and partial applications). *)
let rels, c =
let n = List.length s
- and m = nb_lam Evd.empty (EConstr.of_constr body) (** FIXME *) in
- if n <= m then decompose_lam_n n body
+ and m = nb_lam sg body in
+ if n <= m then decompose_lam_n sg n body
else
let s,s' = List.chop m s in
if List.for_all ((==) Keep) s' &&
(lang () == Haskell || sign_kind s != UnsafeLogicalSig)
- then decompose_lam_n m body
- else decomp_lams_eta_n n m env body typ
+ then decompose_lam_n sg m body
+ else decomp_lams_eta_n n m env sg body typ
in
(* Should we do one eta-expansion to avoid non-generalizable '_a ? *)
let rels, c =
@@ -921,9 +976,9 @@ let extract_std_constant env kn body typ =
let s,s' = List.chop n s in
let k = sign_kind s in
let empty_s = (k == EmptySig || k == SafeLogicalSig) in
- if lang () == Ocaml && empty_s && not (gentypvar_ok c)
+ if lang () == Ocaml && empty_s && not (gentypvar_ok sg c)
&& not (List.is_empty s') && not (Int.equal (type_maxvar t) 0)
- then decomp_lams_eta_n (n+1) n env body typ
+ then decomp_lams_eta_n (n+1) n env sg body typ
else rels,c
in
let n = List.length rels in
@@ -937,16 +992,16 @@ let extract_std_constant env kn body typ =
(* The according Coq environment. *)
let env = push_rels_assum rels env in
(* The real extraction: *)
- let e = extract_term env mle t' c [] in
+ let e = extract_term env sg mle t' c [] in
(* Expunging term and type from dummy lambdas. *)
let trm = term_expunge s (ids,e) in
trm, type_expunge_from_sign env s t
(* Extracts the type of an axiom, honors the Extraction Implicit declaration. *)
-let extract_axiom env kn typ =
+let extract_axiom env sg kn typ =
reset_meta_count ();
(* The short type [t] (i.e. possibly with abbreviations). *)
- let t = snd (record_constant_type env kn (Some typ)) in
+ let t = snd (record_constant_type env sg kn (Some typ)) in
(* The real type [t']: without head products, expanded, *)
(* and with [Tvar] translated to [Tvar'] (not instantiable). *)
let l,_ = type_decomp (expand env (var2var' t)) in
@@ -955,18 +1010,19 @@ let extract_axiom env kn typ =
let s = sign_with_implicits (ConstRef kn) s 0 in
type_expunge_from_sign env s t
-let extract_fixpoint env vkn (fi,ti,ci) =
+let extract_fixpoint env sg vkn (fi,ti,ci) =
let n = Array.length vkn in
let types = Array.make n (Tdummy Kprop)
and terms = Array.make n (MLdummy Kprop) in
let kns = Array.to_list vkn in
current_fixpoints := kns;
(* for replacing recursive calls [Rel ..] by the corresponding [Const]: *)
- let sub = List.rev_map mkConst kns in
+ let sub = List.rev_map EConstr.mkConst kns in
for i = 0 to n-1 do
- if sort_of env ti.(i) != InProp then
+ if sort_of env sg ti.(i) != InProp then
try
- let e,t = extract_std_constant env vkn.(i) (substl sub ci.(i)) ti.(i) in
+ let e,t = extract_std_constant env sg vkn.(i)
+ (EConstr.Vars.substl sub ci.(i)) ti.(i) in
terms.(i) <- e;
types.(i) <- t;
with SingletonInductiveBecomesProp id ->
@@ -976,32 +1032,33 @@ let extract_fixpoint env vkn (fi,ti,ci) =
Dfix (Array.map (fun kn -> ConstRef kn) vkn, terms, types)
let extract_constant env kn cb =
+ let sg = Evd.from_env env in
let r = ConstRef kn in
- let typ = cb.const_type in
+ let typ = EConstr.of_constr cb.const_type in
let warn_info () = if not (is_custom r) then add_info_axiom r in
let warn_log () = if not (constant_has_body cb) then add_log_axiom r
in
let mk_typ_ax () =
- let n = type_scheme_nb_args env typ in
+ let n = type_scheme_nb_args env sg typ in
let ids = iterate (fun l -> anonymous_name::l) n [] in
Dtype (r, ids, Taxiom)
in
let mk_typ c =
- let s,vl = type_sign_vl env typ in
+ let s,vl = type_sign_vl env sg typ in
let db = db_from_sign s in
- let t = extract_type_scheme env db c (List.length s)
+ let t = extract_type_scheme env sg db c (List.length s)
in Dtype (r, vl, t)
in
let mk_ax () =
- let t = extract_axiom env kn typ in
+ let t = extract_axiom env sg kn typ in
Dterm (r, MLaxiom, t)
in
let mk_def c =
- let e,t = extract_std_constant env kn c typ in
+ let e,t = extract_std_constant env sg kn c typ in
Dterm (r,e,t)
in
try
- match flag_of_type env typ with
+ match flag_of_type env sg typ with
| (Logic,TypeScheme) -> warn_log (); Dtype (r, [], Tdummy Ktype)
| (Logic,Default) -> warn_log (); Dterm (r, MLdummy Kprop, Tdummy Kprop)
| (Info,TypeScheme) ->
@@ -1009,73 +1066,72 @@ let extract_constant env kn cb =
| Undef _ -> warn_info (); mk_typ_ax ()
| Def c ->
(match cb.const_proj with
- | None -> mk_typ (Mod_subst.force_constr c)
- | Some pb -> mk_typ pb.proj_body)
+ | None -> mk_typ (get_body c)
+ | Some pb -> mk_typ (EConstr.of_constr pb.proj_body))
| OpaqueDef c ->
add_opaque r;
- if access_opaque () then
- mk_typ (Opaqueproof.force_proof (Environ.opaque_tables env) c)
+ if access_opaque () then mk_typ (get_opaque env c)
else mk_typ_ax ())
| (Info,Default) ->
(match cb.const_body with
| Undef _ -> warn_info (); mk_ax ()
| Def c ->
(match cb.const_proj with
- | None -> mk_def (Mod_subst.force_constr c)
- | Some pb -> mk_def pb.proj_body)
+ | None -> mk_def (get_body c)
+ | Some pb -> mk_def (EConstr.of_constr pb.proj_body))
| OpaqueDef c ->
add_opaque r;
- if access_opaque () then
- mk_def (Opaqueproof.force_proof (Environ.opaque_tables env) c)
+ if access_opaque () then mk_def (get_opaque env c)
else mk_ax ())
with SingletonInductiveBecomesProp id ->
error_singleton_become_prop id (Some (ConstRef kn))
let extract_constant_spec env kn cb =
+ let sg = Evd.from_env env in
let r = ConstRef kn in
- let typ = cb.const_type in
+ let typ = EConstr.of_constr cb.const_type in
try
- match flag_of_type env typ with
+ match flag_of_type env sg typ with
| (Logic, TypeScheme) -> Stype (r, [], Some (Tdummy Ktype))
| (Logic, Default) -> Sval (r, Tdummy Kprop)
| (Info, TypeScheme) ->
- let s,vl = type_sign_vl env typ in
+ let s,vl = type_sign_vl env sg typ in
(match cb.const_body with
| Undef _ | OpaqueDef _ -> Stype (r, vl, None)
| Def body ->
let db = db_from_sign s in
- let body = Mod_subst.force_constr body in
- let t = extract_type_scheme env db body (List.length s)
- in Stype (r, vl, Some t))
+ let body = get_body body in
+ let t = extract_type_scheme env sg db body (List.length s)
+ in Stype (r, vl, Some t))
| (Info, Default) ->
- let t = snd (record_constant_type env kn (Some typ)) in
- Sval (r, type_expunge env t)
+ let t = snd (record_constant_type env sg kn (Some typ)) in
+ Sval (r, type_expunge env t)
with SingletonInductiveBecomesProp id ->
error_singleton_become_prop id (Some (ConstRef kn))
-let extract_with_type env c =
+let extract_with_type env sg c =
try
- let typ = type_of env c in
- match flag_of_type env typ with
+ let typ = type_of env sg c in
+ match flag_of_type env sg typ with
| (Info, TypeScheme) ->
- let s,vl = type_sign_vl env typ in
- let db = db_from_sign s in
- let t = extract_type_scheme env db c (List.length s) in
- Some (vl, t)
+ let s,vl = type_sign_vl env sg typ in
+ let db = db_from_sign s in
+ let t = extract_type_scheme env sg db c (List.length s) in
+ Some (vl, t)
| _ -> None
with SingletonInductiveBecomesProp id ->
error_singleton_become_prop id None
-let extract_constr env c =
+let extract_constr env sg c =
reset_meta_count ();
try
- let typ = type_of env c in
- match flag_of_type env typ with
+ let typ = type_of env sg c in
+ match flag_of_type env sg typ with
| (_,TypeScheme) -> MLdummy Ktype, Tdummy Ktype
| (Logic,_) -> MLdummy Kprop, Tdummy Kprop
| (Info,Default) ->
- let mlt = extract_type env [] 1 typ [] in
- extract_term env Mlenv.empty mlt c [], mlt
+ let mlt = extract_type env sg [] 1 typ [] in
+ extract_term env sg Mlenv.empty mlt c [], mlt
with SingletonInductiveBecomesProp id ->
error_singleton_become_prop id None
diff --git a/plugins/extraction/extraction.mli b/plugins/extraction/extraction.mli
index a0f2885a4..d27c79cb6 100644
--- a/plugins/extraction/extraction.mli
+++ b/plugins/extraction/extraction.mli
@@ -11,9 +11,9 @@
(*s Extraction from Coq terms to Miniml. *)
open Names
-open Constr
open Declarations
open Environ
+open Evd
open Miniml
val extract_constant : env -> Constant.t -> constant_body -> ml_decl
@@ -22,16 +22,18 @@ val extract_constant_spec : env -> Constant.t -> constant_body -> ml_spec
(** For extracting "module ... with ..." declaration *)
-val extract_with_type : env -> constr -> ( Id.t list * ml_type ) option
+val extract_with_type :
+ env -> evar_map -> EConstr.t -> ( Id.t list * ml_type ) option
val extract_fixpoint :
- env -> Constant.t array -> (constr, types) prec_declaration -> ml_decl
+ env -> evar_map -> Constant.t array ->
+ (EConstr.t, EConstr.types) Constr.prec_declaration -> ml_decl
val extract_inductive : env -> MutInd.t -> ml_ind
-(** For extraction compute *)
+(** For Extraction Compute and Show Extraction *)
-val extract_constr : env -> constr -> ml_ast * ml_type
+val extract_constr : env -> evar_map -> EConstr.t -> ml_ast * ml_type
(*s Is a [ml_decl] or a [ml_spec] logical ? *)
diff --git a/plugins/extraction/g_extraction.ml4 b/plugins/extraction/g_extraction.ml4
index 468f2fe8c..93909f3e6 100644
--- a/plugins/extraction/g_extraction.ml4
+++ b/plugins/extraction/g_extraction.ml4
@@ -160,3 +160,9 @@ VERNAC COMMAND EXTEND ExtractionInductive CLASSIFIED AS SIDEFF
mlname(id) "[" mlname_list(idl) "]" string_opt(o) ]
-> [ extract_inductive x id idl o ]
END
+(* Show the extraction of the current proof *)
+
+VERNAC COMMAND EXTEND ShowExtraction CLASSIFIED AS QUERY
+| [ "Show" "Extraction" ]
+ -> [ show_extraction () ]
+END
diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml
index 0bcda69d4..6c421491f 100644
--- a/plugins/extraction/table.ml
+++ b/plugins/extraction/table.ml
@@ -38,14 +38,13 @@ module Refset' = Refset_env
let occur_kn_in_ref kn = function
| IndRef (kn',_)
| ConstructRef ((kn',_),_) -> MutInd.equal kn kn'
- | ConstRef _ -> false
- | VarRef _ -> assert false
+ | ConstRef _ | VarRef _ -> false
let repr_of_r = function
| ConstRef kn -> Constant.repr3 kn
| IndRef (kn,_)
| ConstructRef ((kn,_),_) -> MutInd.repr3 kn
- | VarRef _ -> assert false
+ | VarRef v -> KerName.repr (Lib.make_kn v)
let modpath_of_r r =
let mp,_,_ = repr_of_r r in mp
@@ -279,7 +278,7 @@ let safe_basename_of_global r =
| ConstructRef ((kn,i),j) ->
(try (unsafe_lookup_ind kn).ind_packets.(i).ip_consnames.(j-1)
with Not_found -> last_chance r)
- | VarRef _ -> assert false
+ | VarRef v -> v
let string_of_global r =
try string_of_qualid (Nametab.shortest_qualid_of_global Id.Set.empty r)
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index 6e38b4641..e0368153e 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -1922,7 +1922,7 @@ let build_morphism_signature env sigma m =
in
let morph = e_app_poly env evd PropGlobal.proper_type [| t; sig_; m |] in
let evd = solve_constraints env !evd in
- let evd = Evd.nf_constraints evd in
+ let evd = Evd.minimize_universes evd in
let m = Evarutil.nf_evars_universes evd (EConstr.Unsafe.to_constr morph) in
Pretyping.check_evars env Evd.empty evd (EConstr.of_constr m);
Evd.evar_universe_context evd, m
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index 43c29a08a..f049963f1 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -1193,7 +1193,7 @@ let pf_interp_gen_aux gl to_ind ((oclr, occ), t) =
else let gl, ccl = pf_mkprod gl c cl in false, pat, ccl, c, clr,ucst,gl
else if to_ind && occ = None then
let nv, p, _, ucst' = pf_abs_evars gl (fst pat, c) in
- let ucst = Evd.union_evar_universe_context ucst ucst' in
+ let ucst = UState.union ucst ucst' in
if nv = 0 then anomaly "occur_existential but no evars" else
let gl, pty = pfe_type_of gl p in
false, pat, EConstr.mkProd (constr_name (project gl) c, pty, Tacmach.pf_concl gl), p, clr,ucst,gl
diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml
index 859513d5c..57635edac 100644
--- a/plugins/ssr/ssrequality.ml
+++ b/plugins/ssr/ssrequality.ml
@@ -216,7 +216,7 @@ let same_proj sigma t1 t2 =
let all_ok _ _ = true
let fake_pmatcher_end () =
- mkProp, L2R, (Evd.empty, Evd.empty_evar_universe_context, mkProp)
+ mkProp, L2R, (Evd.empty, UState.empty, mkProp)
let unfoldintac occ rdx t (kt,_) gl =
let fs sigma x = Reductionops.nf_evar sigma x in
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index 62c35d6df..33b18001c 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -361,7 +361,7 @@ let unif_end env sigma0 ise0 pt ok =
if ise2 == ise1 then (s, uc, t)
else
let s, uc', t = nf_open_term sigma0 ise2 t in
- s, Evd.union_evar_universe_context uc uc', t
+ s, UState.union uc uc', t
let unify_HO env sigma0 t1 t2 =
let sigma = unif_HO env sigma0 t1 t2 in
@@ -1268,7 +1268,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst =
let sigma,pat= mk_tpattern ?hack env sigma0 (sigma,p) ok L2R (fs sigma t) in
sigma, [pat] in
match pattern with
- | None -> do_subst env0 concl0 concl0 1, Evd.empty_evar_universe_context
+ | None -> do_subst env0 concl0 concl0 1, UState.empty
| Some (sigma, (T rp | In_T rp)) ->
let rp = fs sigma rp in
let ise = create_evar_defs sigma in
diff --git a/pretyping/univdecls.ml b/pretyping/univdecls.ml
index d16f046fa..8864be576 100644
--- a/pretyping/univdecls.ml
+++ b/pretyping/univdecls.ml
@@ -38,7 +38,7 @@ let interp_univ_constraints env evd cstrs =
let interp_univ_decl env decl =
let open Misctypes in
let pl : lident list = decl.univdecl_instance in
- let evd = Evd.from_ctx (Evd.make_evar_universe_context env (Some pl)) in
+ let evd = Evd.from_ctx (UState.make_with_initial_binders (Environ.universes env) pl) in
let evd, cstrs = interp_univ_constraints env evd decl.univdecl_constraints in
let decl = { univdecl_instance = pl;
univdecl_extensible_instance = decl.univdecl_extensible_instance;
diff --git a/printing/prettyp.ml b/printing/prettyp.ml
index 3682b7c25..9da94e42a 100644
--- a/printing/prettyp.ml
+++ b/printing/prettyp.ml
@@ -94,7 +94,7 @@ let print_ref reduce ref udecl =
let env = Global.env () in
let bl = Universes.universe_binders_with_opt_names ref
(Array.to_list (Univ.Instance.to_array inst)) udecl in
- let sigma = Evd.from_ctx (Evd.evar_universe_context_of_binders bl) in
+ let sigma = Evd.from_ctx (UState.of_binders bl) in
let inst =
if Global.is_polymorphic ref then Printer.pr_universe_instance sigma univs
else mt ()
@@ -593,7 +593,7 @@ let print_constant with_values sep sp udecl =
Array.to_list (Instance.to_array inst)
in
let ctx =
- Evd.evar_universe_context_of_binders
+ UState.of_binders
(Universes.universe_binders_with_opt_names (ConstRef sp) ulist udecl)
in
let env = Global.env () and sigma = Evd.from_ctx ctx in
diff --git a/printing/printer.ml b/printing/printer.ml
index 917ee2021..e50d302b3 100644
--- a/printing/printer.ml
+++ b/printing/printer.ml
@@ -539,7 +539,7 @@ let pr_evgl_sign sigma evi =
let ids = List.rev_map NamedDecl.get_id l in
let warn =
if List.is_empty ids then mt () else
- (str "(" ++ prlist_with_sep pr_comma pr_id ids ++ str " cannot be used)")
+ (str " (" ++ prlist_with_sep pr_comma pr_id ids ++ str " cannot be used)")
in
let pc = pr_lconstr_env env sigma evi.evar_concl in
let candidates =
@@ -551,7 +551,7 @@ let pr_evgl_sign sigma evi =
mt ()
in
hov 0 (str"[" ++ ps ++ spc () ++ str"|- " ++ pc ++ str"]" ++
- candidates ++ spc () ++ warn)
+ candidates ++ warn)
(* Print an existential variable *)
@@ -560,15 +560,25 @@ let pr_evar sigma (evk, evi) =
hov 0 (pr_existential_key sigma evk ++ str " : " ++ pegl)
(* Print an enumerated list of existential variables *)
-let rec pr_evars_int_hd head sigma i = function
+let rec pr_evars_int_hd pr sigma i = function
| [] -> mt ()
| (evk,evi)::rest ->
- (hov 0 (head i ++ pr_evar sigma (evk,evi))) ++
- (match rest with [] -> mt () | _ -> fnl () ++ pr_evars_int_hd head sigma (i+1) rest)
-
-let pr_evars_int sigma i evs = pr_evars_int_hd (fun i -> str "Existential " ++ int i ++ str " =" ++ spc ()) sigma i (Evar.Map.bindings evs)
-
-let pr_evars sigma evs = pr_evars_int_hd (fun i -> mt ()) sigma 1 (Evar.Map.bindings evs)
+ (hov 0 (pr i evk evi)) ++
+ (match rest with [] -> mt () | _ -> fnl () ++ pr_evars_int_hd pr sigma (i+1) rest)
+
+let pr_evars_int sigma ~shelf ~givenup i evs =
+ let pr_status i =
+ if List.mem i shelf then str " (shelved)"
+ else if List.mem i givenup then str " (given up)"
+ else mt () in
+ pr_evars_int_hd
+ (fun i evk evi ->
+ str "Existential " ++ int i ++ str " =" ++
+ spc () ++ pr_evar sigma (evk,evi) ++ pr_status evk)
+ sigma i (Evar.Map.bindings evs)
+
+let pr_evars sigma evs =
+ pr_evars_int_hd (fun i evk evi -> pr_evar sigma (evk,evi)) sigma 1 (Evar.Map.bindings evs)
(* Display a list of evars given by their name, with a prefix *)
let pr_ne_evar_set hd tl sigma l =
@@ -686,7 +696,7 @@ let print_dependent_evars gl sigma seeds =
(* spiwack: [pr_first] is true when the first goal must be singled out
and printed in its entirety. *)
let default_pr_subgoals ?(pr_first=true)
- close_cmd sigma seeds shelf stack unfocused goals =
+ close_cmd sigma ~seeds ~shelf ~stack ~unfocused ~goals =
(** Printing functions for the extra informations. *)
let rec print_stack a = function
| [] -> Pp.int a
@@ -748,7 +758,7 @@ let default_pr_subgoals ?(pr_first=true)
if Evar.Map.is_empty exl then
(str"No more subgoals." ++ print_dependent_evars None sigma seeds)
else
- let pei = pr_evars_int sigma 1 exl in
+ let pei = pr_evars_int sigma ~shelf ~givenup:[] 1 exl in
v 0 ((str "No more subgoals,"
++ str " but there are non-instantiated existential variables:"
++ cut () ++ (hov 0 pei)
@@ -775,7 +785,7 @@ let default_pr_subgoals ?(pr_first=true)
type printer_pr = {
- pr_subgoals : ?pr_first:bool -> Pp.t option -> evar_map -> Evar.t list -> Goal.goal list -> int list -> goal list -> goal list -> Pp.t;
+ pr_subgoals : ?pr_first:bool -> Pp.t option -> evar_map -> seeds:goal list -> shelf:goal list -> stack:int list -> unfocused:goal list -> goals:goal list -> Pp.t;
pr_subgoal : int -> evar_map -> goal list -> Pp.t;
pr_goal : goal sigma -> Pp.t;
}
@@ -809,16 +819,16 @@ let pr_open_subgoals ~proof =
begin match goals with
| [] -> let { Evd.it = bgoals ; sigma = bsigma } = Proof.V82.background_subgoals p in
begin match bgoals,shelf,given_up with
- | [] , [] , [] -> pr_subgoals None sigma seeds shelf stack [] goals
+ | [] , [] , [] -> pr_subgoals None sigma ~seeds ~shelf ~stack ~unfocused:[] ~goals
| [] , [] , _ ->
Feedback.msg_info (str "No more subgoals, but there are some goals you gave up:");
fnl ()
- ++ pr_subgoals ~pr_first:false None bsigma seeds [] [] [] given_up
+ ++ pr_subgoals ~pr_first:false None bsigma ~seeds ~shelf:[] ~stack:[] ~unfocused:[] ~goals:given_up
++ fnl () ++ str "You need to go back and solve them."
| [] , _ , _ ->
Feedback.msg_info (str "All the remaining goals are on the shelf.");
fnl ()
- ++ pr_subgoals ~pr_first:false None bsigma seeds [] [] [] shelf
+ ++ pr_subgoals ~pr_first:false None bsigma ~seeds ~shelf:[] ~stack:[] ~unfocused:[] ~goals:shelf
| _ , _, _ ->
let end_cmd =
str "This subproof is complete, but there are some unfocused goals." ++
@@ -826,13 +836,13 @@ let pr_open_subgoals ~proof =
if Pp.ismt s then s else fnl () ++ s) ++
fnl ()
in
- pr_subgoals ~pr_first:false (Some end_cmd) bsigma seeds shelf [] [] bgoals
+ pr_subgoals ~pr_first:false (Some end_cmd) bsigma ~seeds ~shelf ~stack:[] ~unfocused:[] ~goals:bgoals
end
| _ ->
let { Evd.it = bgoals ; sigma = bsigma } = Proof.V82.background_subgoals p in
let bgoals_focused, bgoals_unfocused = List.partition (fun x -> List.mem x goals) bgoals in
let unfocused_if_needed = if should_unfoc() then bgoals_unfocused else [] in
- pr_subgoals ~pr_first:true None bsigma seeds shelf [] unfocused_if_needed bgoals_focused
+ pr_subgoals ~pr_first:true None bsigma ~seeds ~shelf ~stack:[] ~unfocused:unfocused_if_needed ~goals:bgoals_focused
end
let pr_nth_open_subgoal ~proof n =
diff --git a/printing/printer.mli b/printing/printer.mli
index e32cb0921..41843680b 100644
--- a/printing/printer.mli
+++ b/printing/printer.mli
@@ -183,8 +183,7 @@ val pr_goal : goal sigma -> Pp.t
focused goals unless the conrresponding option
[enable_unfocused_goal_printing] is set. [seeds] is for printing
dependent evars (mainly for emacs proof tree mode). *)
-val pr_subgoals : ?pr_first:bool -> Pp.t option -> evar_map -> Evar.t list -> Goal.goal list -> int list
- -> goal list -> goal list -> Pp.t
+val pr_subgoals : ?pr_first:bool -> Pp.t option -> evar_map -> seeds:goal list -> shelf:goal list -> stack:int list -> unfocused:goal list -> goals:goal list -> Pp.t
val pr_subgoal : int -> evar_map -> goal list -> Pp.t
val pr_concl : int -> evar_map -> goal -> Pp.t
@@ -192,7 +191,7 @@ val pr_concl : int -> evar_map -> goal -> Pp.t
val pr_open_subgoals : proof:Proof.t -> Pp.t
val pr_nth_open_subgoal : proof:Proof.t -> int -> Pp.t
val pr_evar : evar_map -> (Evar.t * evar_info) -> Pp.t
-val pr_evars_int : evar_map -> int -> evar_info Evar.Map.t -> Pp.t
+val pr_evars_int : evar_map -> shelf:goal list -> givenup:goal list -> int -> evar_info Evar.Map.t -> Pp.t
val pr_evars : evar_map -> evar_info Evar.Map.t -> Pp.t
val pr_ne_evar_set : Pp.t -> Pp.t -> evar_map ->
Evar.Set.t -> Pp.t
@@ -225,7 +224,8 @@ val pr_assumptionset : env -> evar_map -> types ContextObjectMap.t -> Pp.t
val pr_goal_by_id : proof:Proof.t -> Id.t -> Pp.t
type printer_pr = {
- pr_subgoals : ?pr_first:bool -> Pp.t option -> evar_map -> Evar.t list -> Goal.goal list -> int list -> goal list -> goal list -> Pp.t;
+ pr_subgoals : ?pr_first:bool -> Pp.t option -> evar_map -> seeds:goal list -> shelf:goal list -> stack:int list -> unfocused:goal list -> goals:goal list -> Pp.t;
+
pr_subgoal : int -> evar_map -> goal list -> Pp.t;
pr_goal : goal sigma -> Pp.t;
}
diff --git a/printing/printmod.ml b/printing/printmod.ml
index 43796ec61..e076c10f3 100644
--- a/printing/printmod.ml
+++ b/printing/printmod.ml
@@ -141,7 +141,7 @@ let print_mutual_inductive env mind mib udecl =
else []
in
let bl = Universes.universe_binders_with_opt_names (IndRef (mind, 0)) univs udecl in
- let sigma = Evd.from_ctx (Evd.evar_universe_context_of_binders bl) in
+ let sigma = Evd.from_ctx (UState.of_binders bl) in
hov 0 (Printer.pr_polymorphic (Declareops.inductive_is_polymorphic mib) ++
Printer.pr_cumulative
(Declareops.inductive_is_polymorphic mib)
@@ -185,7 +185,7 @@ let print_record env mind mib udecl =
let envpar = push_rel_context params env in
let bl = Universes.universe_binders_with_opt_names (IndRef (mind,0))
(Array.to_list (Univ.Instance.to_array u)) udecl in
- let sigma = Evd.from_ctx (Evd.evar_universe_context_of_binders bl) in
+ let sigma = Evd.from_ctx (UState.of_binders bl) in
let keyword =
let open Declarations in
match mib.mind_finite with
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
index d713b0999..8b5b739a3 100644
--- a/proofs/proof_global.ml
+++ b/proofs/proof_global.ml
@@ -341,7 +341,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now
let subst_evar k =
Proof.in_proof proof (fun m -> Evd.existential_opt_value m k) in
let nf = Universes.nf_evars_and_universes_opt_subst subst_evar
- (Evd.evar_universe_context_subst universes) in
+ (UState.subst universes) in
let make_body =
if poly || now then
let make_body t (c, eff) =
@@ -436,7 +436,7 @@ let return_proof ?(allow_partial=false) () =
| Proof.HasUnresolvedEvar->
error(strbrk"Attempt to save a proof with existential variables still non-instantiated") in
let eff = Evd.eval_side_effects evd in
- let evd = Evd.nf_constraints evd in
+ let evd = Evd.minimize_universes evd in
(** ppedrot: FIXME, this is surely wrong. There is no reason to duplicate
side-effects... This may explain why one need to uniquize side-effects
thereafter... *)
diff --git a/stm/stm.ml b/stm/stm.ml
index d878bbb30..b3da97c6e 100644
--- a/stm/stm.ml
+++ b/stm/stm.ml
@@ -2482,7 +2482,7 @@ let known_state ?(redefine_qed=false) ~cache id =
match keep with
| VtDrop -> None
| VtKeepAsAxiom ->
- let ctx = Evd.empty_evar_universe_context in
+ let ctx = UState.empty in
let fp = Future.from_val ([],ctx) in
qed.fproof <- Some (fp, ref false); None
| VtKeep ->
diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml
index 45926551b..477de6452 100644
--- a/tactics/eqschemes.ml
+++ b/tactics/eqschemes.ml
@@ -214,7 +214,7 @@ let build_sym_scheme env ind =
rel_vect (2*nrealargs+2) nrealargs])),
mkRel 1 (* varH *),
[|cstr (nrealargs+1)|]))))
- in c, Evd.evar_universe_context_of ctx
+ in c, UState.of_context_set ctx
let sym_scheme_kind =
declare_individual_scheme_object "_sym_internal"
@@ -285,7 +285,7 @@ let build_sym_involutive_scheme env ind =
mkRel 1|])),
mkRel 1 (* varH *),
[|mkApp(eqrefl,[|applied_ind_C;cstr (nrealargs+1)|])|]))))
- in (c, Evd.evar_universe_context_of ctx), eff
+ in (c, UState.of_context_set ctx), eff
let sym_involutive_scheme_kind =
declare_individual_scheme_object "_sym_involutive"
@@ -439,7 +439,7 @@ let build_l2r_rew_scheme dep env ind kind =
[|main_body|])
else
main_body))))))
- in (c, Evd.evar_universe_context_of ctx),
+ in (c, UState.of_context_set ctx),
Safe_typing.concat_private eff' eff
(**********************************************************************)
@@ -528,7 +528,7 @@ let build_l2r_forward_rew_scheme dep env ind kind =
(if dep then realsign_ind_P 1 applied_ind_P' else realsign_P 2) s)
(mkNamedLambda varHC applied_PC'
(mkVar varHC))|])))))
- in c, Evd.evar_universe_context_of ctx
+ in c, UState.of_context_set ctx
(**********************************************************************)
(* Build the right-to-left rewriting lemma for hypotheses associated *)
@@ -601,7 +601,7 @@ let build_r2l_forward_rew_scheme dep env ind kind =
lift (nrealargs+3) applied_PC,
mkRel 1)|]),
[|mkVar varHC|]))))))
- in c, Evd.evar_universe_context_of ctx
+ in c, UState.of_context_set ctx
(**********************************************************************)
(* This function "repairs" the non-dependent r2l forward rewriting *)
@@ -808,7 +808,7 @@ let build_congr env (eq,refl,ctx) ind =
[|mkApp (refl,
[|mkVar varB;
mkApp (mkVar varf, [|lift (mip.mind_nrealargs+3) b|])|])|]))))))
- in c, Evd.evar_universe_context_of ctx
+ in c, UState.of_context_set ctx
let congr_scheme_kind = declare_individual_scheme_object "_congr"
(fun _ ind ->
diff --git a/tactics/ind_tables.ml b/tactics/ind_tables.ml
index b960a845c..62ead57f3 100644
--- a/tactics/ind_tables.ml
+++ b/tactics/ind_tables.ml
@@ -122,8 +122,8 @@ let compute_name internal id =
let define internal id c p univs =
let fd = declare_constant ~internal in
let id = compute_name internal id in
- let ctx = Evd.normalize_evar_universe_context univs in
- let c = Universes.subst_opt_univs_constr (Evd.evar_universe_context_subst ctx) c in
+ let ctx = UState.minimize univs in
+ let c = Universes.subst_opt_univs_constr (UState.subst ctx) c in
let univs =
if p then Polymorphic_const_entry (UState.context ctx)
else Monomorphic_const_entry (UState.context_set ctx)
diff --git a/tactics/leminv.ml b/tactics/leminv.ml
index 655283c20..a4cdc1592 100644
--- a/tactics/leminv.ml
+++ b/tactics/leminv.ml
@@ -218,7 +218,7 @@ let inversion_scheme env sigma t sort dep_option inv_op =
end in
let avoid = ref Id.Set.empty in
let _,_,_,_,sigma = Proof.proof pf in
- let sigma = Evd.nf_constraints sigma in
+ let sigma = Evd.minimize_universes sigma in
let rec fill_holes c =
match EConstr.kind sigma c with
| Evar (e,args) ->
diff --git a/test-suite/output/Arguments_renaming.out b/test-suite/output/Arguments_renaming.out
index 4df21ae35..e73312c67 100644
--- a/test-suite/output/Arguments_renaming.out
+++ b/test-suite/output/Arguments_renaming.out
@@ -11,7 +11,7 @@ notation scopes add ': clear scopes' [arguments-assert,vernacular]
eq_refl
: ?y = ?y
where
-?y : [ |- nat]
+?y : [ |- nat]
Inductive eq (A : Type) (x : A) : A -> Prop := eq_refl : x = x
For eq_refl: Arguments are renamed to B, y
diff --git a/test-suite/output/Existentials.out b/test-suite/output/Existentials.out
index 9680d2bbf..18f5d89f6 100644
--- a/test-suite/output/Existentials.out
+++ b/test-suite/output/Existentials.out
@@ -1,4 +1,4 @@
-Existential 1 = ?Goal : [p : nat q := S p : nat n : nat m : nat |- ?y = m]
+Existential 1 = ?Goal : [p : nat q := S p : nat n : nat m : nat |- ?y = m]
Existential 2 =
-?y : [p : nat q := S p : nat n : nat m : nat |- nat] (p, q cannot be used)
-Existential 3 = ?Goal0 : [q : nat n : nat m : nat |- n = ?y]
+?y : [p : nat q := S p : nat n : nat m : nat |- nat] (p, q cannot be used) (shelved)
+Existential 3 = ?Goal0 : [q : nat n : nat m : nat |- n = ?y]
diff --git a/test-suite/output/Notations3.out b/test-suite/output/Notations3.out
index e6a6e0288..864b6151a 100644
--- a/test-suite/output/Notations3.out
+++ b/test-suite/output/Notations3.out
@@ -33,24 +33,24 @@ fun f : forall x : nat * (bool * unit), ?T => CURRY (x : nat) (y : bool), f
: (forall x : nat * (bool * unit), ?T) ->
forall (x : nat) (y : bool), ?T@{x:=(x, (y, tt))}
where
-?T : [x : nat * (bool * unit) |- Type]
+?T : [x : nat * (bool * unit) |- Type]
fun f : forall x : bool * (nat * unit), ?T =>
CURRYINV (x : nat) (y : bool), f
: (forall x : bool * (nat * unit), ?T) ->
forall (x : nat) (y : bool), ?T@{x:=(y, (x, tt))}
where
-?T : [x : bool * (nat * unit) |- Type]
+?T : [x : bool * (nat * unit) |- Type]
fun f : forall x : unit * nat * bool, ?T => CURRYLEFT (x : nat) (y : bool), f
: (forall x : unit * nat * bool, ?T) ->
forall (x : nat) (y : bool), ?T@{x:=(tt, x, y)}
where
-?T : [x : unit * nat * bool |- Type]
+?T : [x : unit * nat * bool |- Type]
fun f : forall x : unit * bool * nat, ?T =>
CURRYINVLEFT (x : nat) (y : bool), f
: (forall x : unit * bool * nat, ?T) ->
forall (x : nat) (y : bool), ?T@{x:=(tt, y, x)}
where
-?T : [x : unit * bool * nat |- Type]
+?T : [x : unit * bool * nat |- Type]
forall n : nat, {#n | 1 > n}
: Prop
forall x : nat, {|x | x > 0|}
diff --git a/test-suite/output/inference.out b/test-suite/output/inference.out
index d28ee4276..5e9eff048 100644
--- a/test-suite/output/inference.out
+++ b/test-suite/output/inference.out
@@ -9,10 +9,10 @@ fun (m n p : nat) (H : S m <= S n + p) => le_S_n m (n + p) H
fun n : nat => let y : T n := A n in ?t ?x : T n
: forall n : nat, T n
where
-?t : [n : nat y := A n : T n |- ?T -> T n]
-?x : [n : nat y := A n : T n |- ?T]
+?t : [n : nat y := A n : T n |- ?T -> T n]
+?x : [n : nat y := A n : T n |- ?T]
fun n : nat => ?t ?x : T n
: forall n : nat, T n
where
-?t : [n : nat |- ?T -> T n]
-?x : [n : nat |- ?T]
+?t : [n : nat |- ?T -> T n]
+?x : [n : nat |- ?T]
diff --git a/test-suite/success/Hints.v b/test-suite/success/Hints.v
index 6962e43e7..8d08f5975 100644
--- a/test-suite/success/Hints.v
+++ b/test-suite/success/Hints.v
@@ -172,3 +172,14 @@ Hint Cut [_* (a_is_b | b_is_c | c_is_d | d_is_e)
Timeout 1 Fail apply _. (* 0.06s *)
Abort.
End HintCut.
+
+
+(* Check that auto-like tactics do not prefer "eq_refl" over more complex solutions, *)
+(* e.g. those tactics when considering a goal with existential varibles *)
+(* like "m = ?n" won't pick "plus_n_O" hint over "eq_refl" hint. *)
+(* See this Coq club post for more detail: *)
+(* https://sympa.inria.fr/sympa/arc/coq-club/2017-12/msg00103.html *)
+
+Goal forall (m : nat), exists n, m = n /\ m = n.
+ intros m; eexists; split; [trivial | reflexivity].
+Qed.
diff --git a/test-suite/success/ShowExtraction.v b/test-suite/success/ShowExtraction.v
new file mode 100644
index 000000000..e34c240c5
--- /dev/null
+++ b/test-suite/success/ShowExtraction.v
@@ -0,0 +1,31 @@
+
+Require Extraction.
+Require Import List.
+
+Section Test.
+Variable A : Type.
+Variable decA : forall (x y:A), {x=y}+{x<>y}.
+
+(** Should fail when no proofs are started *)
+Fail Show Extraction.
+
+Lemma decListA : forall (xs ys : list A), {xs=ys}+{xs<>ys}.
+Proof.
+Show Extraction.
+fix 1.
+destruct xs as [|x xs], ys as [|y ys].
+Show Extraction.
+- now left.
+- now right.
+- now right.
+- Show Extraction.
+ destruct (decA x y).
+ + destruct (decListA xs ys).
+ * left; now f_equal.
+ * Show Extraction.
+ right. congruence.
+ + right. congruence.
+Show Extraction.
+Defined.
+
+End Test.
diff --git a/theories/Compat/Coq87.v b/theories/Compat/Coq87.v
index 89556ee75..dc1397aff 100644
--- a/theories/Compat/Coq87.v
+++ b/theories/Compat/Coq87.v
@@ -9,6 +9,7 @@
(************************************************************************)
(** Compatibility file for making Coq act similar to Coq v8.7 *)
+Require Export Coq.Compat.Coq88.
(* In 8.7, omega wasn't taking advantage of local abbreviations,
see bug 148 and PR#768. For adjusting this flag, we're forced to
diff --git a/theories/Compat/Coq88.v b/theories/Compat/Coq88.v
new file mode 100644
index 000000000..4142af05d
--- /dev/null
+++ b/theories/Compat/Coq88.v
@@ -0,0 +1,11 @@
+(************************************************************************)
+(* * 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) *)
+(************************************************************************)
+
+(** Compatibility file for making Coq act similar to Coq v8.8 *)
diff --git a/theories/Init/Peano.v b/theories/Init/Peano.v
index 73c8c5ef4..d5322d094 100644
--- a/theories/Init/Peano.v
+++ b/theories/Init/Peano.v
@@ -92,7 +92,9 @@ Lemma plus_n_O : forall n:nat, n = n + 0.
Proof.
induction n; simpl; auto.
Qed.
-Hint Resolve plus_n_O: core.
+
+Remove Hints eq_refl : core.
+Hint Resolve plus_n_O eq_refl: core. (* We want eq_refl to have higher priority than plus_n_O *)
Lemma plus_O_n : forall n:nat, 0 + n = n.
Proof.
diff --git a/tools/CoqMakefile.in b/tools/CoqMakefile.in
index 727fd3ec3..e9f64542c 100644
--- a/tools/CoqMakefile.in
+++ b/tools/CoqMakefile.in
@@ -23,6 +23,7 @@ MLFILES := $(COQMF_MLFILES)
ML4FILES := $(COQMF_ML4FILES)
MLPACKFILES := $(COQMF_MLPACKFILES)
MLLIBFILES := $(COQMF_MLLIBFILES)
+CMDLINE_VFILES := $(COQMF_CMDLINE_VFILES)
INSTALLCOQDOCROOT := $(COQMF_INSTALLCOQDOCROOT)
OTHERFLAGS := $(COQMF_OTHERFLAGS)
COQ_SRC_SUBDIRS := $(COQMF_COQ_SRC_SUBDIRS)
@@ -30,6 +31,7 @@ OCAMLLIBS := $(COQMF_OCAMLLIBS)
SRC_SUBDIRS := $(COQMF_SRC_SUBDIRS)
COQLIBS := $(COQMF_COQLIBS)
COQLIBS_NOML := $(COQMF_COQLIBS_NOML)
+CMDLINE_COQLIBS := $(COQMF_CMDLINE_COQLIBS)
LOCAL := $(COQMF_LOCAL)
COQLIB := $(COQMF_COQLIB)
DOCDIR := $(COQMF_DOCDIR)
@@ -724,9 +726,14 @@ $(addsuffix .d,$(MLPACKFILES)): %.mlpack.d: %.mlpack
$(SHOW)'COQDEP $<'
$(HIDE)$(COQDEP) $(OCAMLLIBS) -c "$<" $(redir_if_ok)
+# If this makefile is created using a _CoqProject we have coqdep get
+# options from it. This avoids argument length limits for pathological
+# projects. Note that extra options might be on the command line.
+VDFILE_FLAGS:=$(if @PROJECT_FILE@,-f @PROJECT_FILE@,) $(CMDLINE_COQLIBS) $(CMDLINE_VFILES)
+
$(VDFILE).d: $(VFILES)
$(SHOW)'COQDEP VFILES'
- $(HIDE)$(COQDEP) $(COQLIBS) -dyndep var -c $(VFILES) $(redir_if_ok)
+ $(HIDE)$(COQDEP) -dyndep var $(VDFILE_FLAGS) $(redir_if_ok)
# Misc ########################################################################
diff --git a/tools/coq_makefile.ml b/tools/coq_makefile.ml
index ef4428755..6cd520d60 100644
--- a/tools/coq_makefile.ml
+++ b/tools/coq_makefile.ml
@@ -13,6 +13,8 @@
open CoqProject_file
open Printf
+let (>) f g = fun x -> g (f x)
+
let output_channel = ref stdout
let makefile_name = ref "Makefile"
let make_name = ref ""
@@ -175,21 +177,22 @@ let generate_conf_extra_target oc sps =
in
if sps <> [] then
section oc "Extra targets. (-extra and -extra-phony, DEPRECATED)";
- List.iter pr_path sps
+ List.iter (forget_source > pr_path) sps
let generate_conf_subdirs oc sds =
if sds <> [] then section oc "Subdirectories. (DEPRECATED)";
- List.iter (fprintf oc ".PHONY:%s\n") sds;
- List.iter (fprintf oc "post-all::\n\tcd \"%s\" && $(MAKE) all\n") sds;
- List.iter (fprintf oc "clean::\n\tcd \"%s\" && $(MAKE) clean\n") sds;
- List.iter (fprintf oc "archclean::\n\tcd \"%s\" && $(MAKE) archclean\n") sds;
- List.iter (fprintf oc "install-extra::\n\tcd \"%s\" && $(MAKE) install\n") sds
+ let iter f = List.iter (forget_source > f) in
+ iter (fprintf oc ".PHONY:%s\n") sds;
+ iter (fprintf oc "post-all::\n\tcd \"%s\" && $(MAKE) all\n") sds;
+ iter (fprintf oc "clean::\n\tcd \"%s\" && $(MAKE) clean\n") sds;
+ iter (fprintf oc "archclean::\n\tcd \"%s\" && $(MAKE) archclean\n") sds;
+ iter (fprintf oc "install-extra::\n\tcd \"%s\" && $(MAKE) install\n") sds
let generate_conf_includes oc { ml_includes; r_includes; q_includes } =
section oc "Path directives (-I, -R, -Q).";
let module S = String in
- let open List in
+ let map = map_sourced_list in
let dash1 opt v = sprintf "-%s %s" opt (quote v) in
let dash2 opt v1 v2 = sprintf "-%s %s %s" opt (quote v1) (quote v2) in
fprintf oc "COQMF_OCAMLLIBS = %s\n"
@@ -202,7 +205,11 @@ let generate_conf_includes oc { ml_includes; r_includes; q_includes } =
(S.concat " " (map (fun ({ path },l) -> dash2 "R" path l) r_includes));
fprintf oc "COQMF_COQLIBS_NOML = %s %s\n"
(S.concat " " (map (fun ({ path },l) -> dash2 "Q" path l) q_includes))
- (S.concat " " (map (fun ({ path },l) -> dash2 "R" path l) r_includes))
+ (S.concat " " (map (fun ({ path },l) -> dash2 "R" path l) r_includes));
+ fprintf oc "COQMF_CMDLINE_COQLIBS = %s %s %s\n"
+ (S.concat " " (map_cmdline (fun { path } -> dash1 "I" path) ml_includes))
+ (S.concat " " (map_cmdline (fun ({ path },l) -> dash2 "Q" path l) q_includes))
+ (S.concat " " (map_cmdline (fun ({ path },l) -> dash2 "R" path l) r_includes));
;;
let windrive s =
@@ -219,10 +226,10 @@ let generate_conf_coq_config oc args =
;;
let generate_conf_files oc
- { v_files; mli_files; ml4_files; ml_files; mllib_files; mlpack_files }
+ { v_files; mli_files; ml4_files; ml_files; mllib_files; mlpack_files; }
=
let module S = String in
- let open List in
+ let map = map_sourced_list in
section oc "Project files.";
fprintf oc "COQMF_VFILES = %s\n" (S.concat " " (map quote v_files));
fprintf oc "COQMF_MLIFILES = %s\n" (S.concat " " (map quote mli_files));
@@ -230,6 +237,8 @@ let generate_conf_files oc
fprintf oc "COQMF_ML4FILES = %s\n" (S.concat " " (map quote ml4_files));
fprintf oc "COQMF_MLPACKFILES = %s\n" (S.concat " " (map quote mlpack_files));
fprintf oc "COQMF_MLLIBFILES = %s\n" (S.concat " " (map quote mllib_files));
+ let cmdline_vfiles = filter_cmdline v_files in
+ fprintf oc "COQMF_CMDLINE_VFILES = %s\n" (S.concat " " (List.map quote cmdline_vfiles));
;;
let rec all_start_with prefix = function
@@ -246,12 +255,12 @@ let rec logic_gcd acc = function
else acc
let generate_conf_doc oc { defs; q_includes; r_includes } =
- let includes = List.map snd (q_includes @ r_includes) in
+ let includes = List.map (forget_source > snd) (q_includes @ r_includes) in
let logpaths = List.map (CString.split '.') includes in
let gcd = logic_gcd [] logpaths in
let root =
if gcd = [] then
- if not (List.mem_assoc "INSTALLDEFAULTROOT" defs) then begin
+ if not (List.exists (fun x -> fst x.thing = "INSTALLDEFAULTROOT") defs) then begin
let destination = "orphan_" ^ (String.concat "_" includes) in
eprintf "Warning: no common logical root\n";
eprintf "Warning: in such case INSTALLDEFAULTROOT must be defined\n";
@@ -264,9 +273,9 @@ let generate_conf_doc oc { defs; q_includes; r_includes } =
let generate_conf_defs oc { defs; extra_args } =
section oc "Extra variables.";
- List.iter (fun (k,v) -> Printf.fprintf oc "%s = %s\n" k v) defs;
+ List.iter (forget_source > (fun (k,v) -> Printf.fprintf oc "%s = %s\n" k v)) defs;
Printf.fprintf oc "COQMF_OTHERFLAGS = %s\n"
- (String.concat " " extra_args)
+ (String.concat " " (List.map forget_source extra_args))
let generate_conf oc project args =
fprintf oc "# This configuration file was generated by running:\n";
@@ -284,10 +293,10 @@ let ensure_root_dir
({ ml_includes; r_includes; q_includes;
v_files; ml_files; mli_files; ml4_files;
mllib_files; mlpack_files } as project)
-=
- let open List in
+ =
+ let exists f = List.exists (forget_source > f) in
let here = Sys.getcwd () in
- let not_tops = List.for_all (fun s -> s <> Filename.basename s) in
+ let not_tops = List.for_all (fun s -> s.thing <> Filename.basename s.thing) in
if exists (fun { canonical_path = x } -> x = here) ml_includes
|| exists (fun ({ canonical_path = x },_) -> is_prefix x here) r_includes
|| exists (fun ({ canonical_path = x },_) -> is_prefix x here) q_includes
@@ -297,29 +306,27 @@ let ensure_root_dir
then
project
else
+ let source x = {thing=x; source=CmdLine} in
let here_path = { path = "."; canonical_path = here } in
{ project with
- ml_includes = here_path :: ml_includes;
- r_includes = (here_path, "Top") :: r_includes }
+ ml_includes = source here_path :: ml_includes;
+ r_includes = source (here_path, "Top") :: r_includes }
;;
let warn_install_at_root_directory
- { q_includes; r_includes;
- v_files; ml_files; mli_files; ml4_files;
- mllib_files; mlpack_files }
+ ({ q_includes; r_includes; } as project)
=
let open CList in
let inc_top_p =
map_filter
- (fun ({ path } ,ldir) -> if ldir = "" then Some path else None)
+ (fun {thing=({ path } ,ldir)} -> if ldir = "" then Some path else None)
(r_includes @ q_includes) in
- let files =
- v_files @ mli_files @ ml4_files @ ml_files @ mllib_files @ mlpack_files in
- let bad = filter (fun f -> mem (Filename.dirname f) inc_top_p) files in
+ let files = all_files project in
+ let bad = filter (fun f -> mem (Filename.dirname f.thing) inc_top_p) files in
if bad <> [] then begin
eprintf "Warning: No file should be installed at the root of Coq's library.\n";
eprintf "Warning: No logical path (-R, -Q) applies to these files:\n";
- List.iter (fun x -> eprintf "Warning: %s\n" x) bad;
+ List.iter (fun x -> eprintf "Warning: %s\n" x.thing) bad;
eprintf "\n";
end
;;
@@ -328,10 +335,10 @@ let check_overlapping_include { q_includes; r_includes } =
let pwd = Sys.getcwd () in
let aux = function
| [] -> ()
- | ({ path; canonical_path }, _) :: l ->
+ | {thing = { path; canonical_path }, _} :: l ->
if not (is_prefix pwd canonical_path) then
eprintf "Warning: %s (used in -R or -Q) is not a subdirectory of the current directory\n\n" path;
- List.iter (fun ({ path = p; canonical_path = cp }, _) ->
+ List.iter (fun {thing={ path = p; canonical_path = cp }, _} ->
if is_prefix canonical_path cp || is_prefix cp canonical_path then
eprintf "Warning: %s and %s overlap (used in -R or -Q)\n\n"
path p) l
@@ -354,7 +361,7 @@ let destination_of { ml_includes; q_includes; r_includes; } file =
clean_path (physical_dir_of_logical_dir logic ^ "/" ^
chop_prefix canonical_path file_dir ^ "/") in
let candidates =
- CList.map_filter (fun ({ canonical_path }, logic) ->
+ CList.map_filter (fun {thing={ canonical_path }, logic} ->
if is_prefix canonical_path file_dir then
Some(mk_destination logic canonical_path)
else None) includes
@@ -364,10 +371,10 @@ let destination_of { ml_includes; q_includes; r_includes; } file =
(* BACKWARD COMPATIBILITY: -I into the only logical root *)
begin match
r_includes,
- List.find (fun { canonical_path = p } -> is_prefix p file_dir)
+ List.find (fun {thing={ canonical_path = p }} -> is_prefix p file_dir)
ml_includes
with
- | [{ canonical_path }, logic], { canonical_path = p } ->
+ | [{thing={ canonical_path }, logic}], {thing={ canonical_path = p }} ->
let destination =
clean_path (physical_dir_of_logical_dir logic ^ "/" ^
chop_prefix p file_dir ^ "/") in
diff --git a/tools/coqdep.ml b/tools/coqdep.ml
index f4f143b3b..12b5cab0a 100644
--- a/tools/coqdep.ml
+++ b/tools/coqdep.ml
@@ -446,6 +446,7 @@ let usage () =
eprintf " -boot : For coq developers, prints dependencies over coq library files (omitted by default).\n";
eprintf " -sort : output the given file name ordered by dependencies\n";
eprintf " -noglob | -no-glob : \n";
+ eprintf " -f file : read -I, -Q, -R and filenames from _CoqProject-formatted FILE.";
eprintf " -I dir -as logname : add (non recursively) dir to coq load path under logical name logname\n";
eprintf " -I dir : add (non recursively) dir to ocaml path\n";
eprintf " -R dir -as logname : add and import dir recursively to coq load path under logical name logname\n"; (* deprecate? *)
@@ -462,6 +463,19 @@ let usage () =
let split_period = Str.split (Str.regexp (Str.quote "."))
+let add_q_include path l = add_rec_dir_no_import add_known path (split_period l)
+
+let add_r_include path l = add_rec_dir_import add_known path (split_period l)
+
+let treat_coqproject f =
+ let open CoqProject_file in
+ let iter_sourced f = List.iter (fun {thing} -> f thing) in
+ let project = read_project_file f in
+ iter_sourced (fun { path } -> add_caml_dir path) project.ml_includes;
+ iter_sourced (fun ({ path }, l) -> add_q_include path l) project.q_includes;
+ iter_sourced (fun ({ path }, l) -> add_r_include path l) project.r_includes;
+ iter_sourced (fun f -> treat_file None f) (all_files project)
+
let rec parse = function
| "-c" :: ll -> option_c := true; parse ll
| "-D" :: ll -> option_D := true; parse ll
@@ -469,10 +483,11 @@ let rec parse = function
| "-boot" :: ll -> option_boot := true; parse ll
| "-sort" :: ll -> option_sort := true; parse ll
| ("-noglob" | "-no-glob") :: ll -> option_noglob := true; parse ll
+ | "-f" :: f :: ll -> treat_coqproject f; parse ll
| "-I" :: r :: ll -> add_caml_dir r; parse ll
| "-I" :: [] -> usage ()
- | "-R" :: r :: ln :: ll -> add_rec_dir_import add_known r (split_period ln); parse ll
- | "-Q" :: r :: ln :: ll -> add_rec_dir_no_import add_known r (split_period ln); parse ll
+ | "-R" :: r :: ln :: ll -> add_r_include r ln; parse ll
+ | "-Q" :: r :: ln :: ll -> add_q_include r ln; parse ll
| "-R" :: ([] | [_]) -> usage ()
| "-dumpgraph" :: f :: ll -> option_dump := Some (false, f); parse ll
| "-dumpgraphbox" :: f :: ll -> option_dump := Some (true, f); parse ll
diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml
index 9b7b88b51..2879feba7 100644
--- a/vernac/auto_ind_decl.ml
+++ b/vernac/auto_ind_decl.ml
@@ -323,7 +323,7 @@ let build_beq_scheme mode kn =
raise NoDecidabilityCoInductive;
let fix = mkFix (((Array.make nb_ind 0),i),(names,types,cores)) in
create_input fix),
- Evd.make_evar_universe_context (Global.env ()) None),
+ UState.make (Global.universes ())),
!eff
let beq_scheme_kind = declare_mutual_scheme_object "_beq" build_beq_scheme
@@ -671,7 +671,7 @@ let make_bl_scheme mode mind =
let lnonparrec,lnamesparrec = (* TODO subst *)
context_chop (nparams-nparrec) mib.mind_params_ctxt in
let bl_goal, eff = compute_bl_goal ind lnamesparrec nparrec in
- let ctx = Evd.make_evar_universe_context (Global.env ()) None in
+ let ctx = UState.make (Global.universes ()) in
let side_eff = side_effect_of_mode mode in
let bl_goal = EConstr.of_constr bl_goal in
let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx bl_goal
@@ -795,7 +795,7 @@ let make_lb_scheme mode mind =
let lnonparrec,lnamesparrec =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
let lb_goal, eff = compute_lb_goal ind lnamesparrec nparrec in
- let ctx = Evd.make_evar_universe_context (Global.env ()) None in
+ let ctx = UState.make (Global.universes ()) in
let side_eff = side_effect_of_mode mode in
let lb_goal = EConstr.of_constr lb_goal in
let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx lb_goal
@@ -965,7 +965,7 @@ let make_eq_decidability mode mind =
let nparams = mib.mind_nparams in
let nparrec = mib.mind_nparams_rec in
let u = Univ.Instance.empty in
- let ctx = Evd.make_evar_universe_context (Global.env ()) None in
+ let ctx = UState.make (Global.universes ()) in
let lnonparrec,lnamesparrec =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
let side_eff = side_effect_of_mode mode in
diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml
index 1712634da..6a590758f 100644
--- a/vernac/comAssumption.ml
+++ b/vernac/comAssumption.ml
@@ -160,7 +160,7 @@ let do_assumptions kind nl l =
in
let sigma = solve_remaining_evars all_and_fail_flags env sigma Evd.empty in
(* The universe constraints come from the whole telescope. *)
- let sigma = Evd.nf_constraints sigma in
+ let sigma = Evd.minimize_universes sigma in
let nf_evar c = EConstr.to_constr sigma c in
let uvars, l = List.fold_left_map (fun uvars (coe,t,imps) ->
let t = nf_evar t in
diff --git a/vernac/comDefinition.ml b/vernac/comDefinition.ml
index 01dbe0a0d..b18a60a1f 100644
--- a/vernac/comDefinition.ml
+++ b/vernac/comDefinition.ml
@@ -85,7 +85,7 @@ let interp_definition pl bl poly red_option c ctypopt =
evd, c, imps1@Impargs.lift_implicits (Context.Rel.nhyps ctx) impsty, Some ty
in
(* universe minimization *)
- let evd = Evd.nf_constraints evd in
+ let evd = Evd.minimize_universes evd in
(* Substitute evars and universes, and add parameters.
Note: in program mode some evars may remain. *)
let ctx = List.map (EConstr.to_rel_decl evd) ctx in
diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml
index 41c44b126..27587416b 100644
--- a/vernac/indschemes.ml
+++ b/vernac/indschemes.ml
@@ -381,7 +381,7 @@ let do_mutual_induction_scheme lnamedepindsort =
| None ->
let _, ctx = Global.type_of_global_in_context env0 (IndRef ind) in
let u, ctx = Universes.fresh_instance_from ctx None in
- let evd = Evd.from_ctx (Evd.evar_universe_context_of ctx) in
+ let evd = Evd.from_ctx (UState.of_context_set ctx) in
evd, (ind,u), Some u
| Some ui -> evd, (ind, ui), inst
in
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
index 6447fc350..4f16e1cf6 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -472,7 +472,7 @@ let subst_body expand prg =
let declare_definition prg =
let body, typ = subst_body true prg in
let nf = Universes.nf_evars_and_universes_opt_subst (fun x -> None)
- (Evd.evar_universe_context_subst prg.prg_ctx) in
+ (UState.subst prg.prg_ctx) in
let opaque = prg.prg_opaque in
let fix_exn = Hook.get get_fix_exn () in
let typ = nf typ in
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 5779c07ab..4c9b41b21 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -66,13 +66,13 @@ let show_proof () =
let show_top_evars () =
(* spiwack: new as of Feb. 2010: shows goal evars in addition to non-goal evars. *)
let pfts = Proof_global.give_me_the_proof () in
- let gls,_,_,_,sigma = Proof.proof pfts in
- pr_evars_int sigma 1 (Evd.undefined_map sigma)
+ let gls,_,shelf,givenup,sigma = Proof.proof pfts in
+ pr_evars_int sigma ~shelf ~givenup 1 (Evd.undefined_map sigma)
let show_universes () =
let pfts = Proof_global.give_me_the_proof () in
let gls,_,_,_,sigma = Proof.proof pfts in
- let ctx = Evd.universe_context_set (Evd.nf_constraints sigma) in
+ let ctx = Evd.universe_context_set (Evd.minimize_universes sigma) in
Termops.pr_evar_universe_context (Evd.evar_universe_context sigma) ++ fnl () ++
str "Normalized constraints: " ++ Univ.pr_universe_context_set (Termops.pr_evd_level sigma) ctx