aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--CHANGES6
-rw-r--r--checker/check.mllib1
-rw-r--r--checker/cic.mli3
-rw-r--r--checker/environ.ml3
-rw-r--r--checker/environ.mli1
-rw-r--r--checker/subtyping.ml4
-rw-r--r--checker/values.ml9
-rw-r--r--clib/canary.ml28
-rw-r--r--clib/canary.mli27
-rw-r--r--clib/clib.mllib1
-rw-r--r--clib/hashcons.ml40
-rw-r--r--clib/hashcons.mli3
-rw-r--r--dev/ci/README.md17
-rw-r--r--doc/sphinx/proof-engine/ssreflect-proof-language.rst3
-rw-r--r--doc/sphinx/proof-engine/tactics.rst13
-rw-r--r--engine/evd.ml3
-rw-r--r--engine/evd.mli2
-rw-r--r--interp/declare.ml37
-rw-r--r--kernel/cooking.ml2
-rw-r--r--kernel/cooking.mli1
-rw-r--r--kernel/declarations.ml3
-rw-r--r--kernel/declareops.ml3
-rw-r--r--kernel/entries.ml5
-rw-r--r--kernel/environ.ml2
-rw-r--r--kernel/indtypes.ml57
-rw-r--r--kernel/indtypes.mli2
-rw-r--r--kernel/names.ml11
-rw-r--r--kernel/term_typing.ml31
-rw-r--r--library/heads.ml2
-rw-r--r--plugins/extraction/extraction.ml16
-rw-r--r--plugins/ltac/extratactics.ml415
-rw-r--r--pretyping/cases.ml2
-rw-r--r--pretyping/detyping.ml5
-rw-r--r--pretyping/indrec.ml8
-rw-r--r--pretyping/indrec.mli3
-rw-r--r--pretyping/inductiveops.ml124
-rw-r--r--pretyping/inductiveops.mli12
-rw-r--r--tactics/tactics.ml20
-rw-r--r--tactics/tactics.mli5
-rw-r--r--test-suite/bugs/closed/7421.v39
-rw-r--r--test-suite/success/letproj.v2
-rw-r--r--test-suite/success/primitiveproj.v21
-rw-r--r--vernac/indschemes.ml4
-rw-r--r--vernac/indschemes.mli8
44 files changed, 331 insertions, 273 deletions
diff --git a/CHANGES b/CHANGES
index ce8e313b9..c36f59726 100644
--- a/CHANGES
+++ b/CHANGES
@@ -28,6 +28,12 @@ Tactics
- The `simple apply` tactic now respects the `Opaque` flag when called from
Ltac (`auto` still does not respect it).
+- Tactic `constr_eq` now adds universe constraints needed for the
+ identity to the context (it used to ignore them). New tactic
+ `constr_eq_strict` checks that the required constraints already hold
+ without adding new ones. Preexisting tactic `constr_eq_nounivs` can
+ still be used if you really want to ignore universe constraints.
+
Tools
- Coq_makefile lets one override or extend the following variables from
diff --git a/checker/check.mllib b/checker/check.mllib
index f79ba66e3..139fa765b 100644
--- a/checker/check.mllib
+++ b/checker/check.mllib
@@ -3,7 +3,6 @@ Coq_config
Analyze
Hook
Terminal
-Canary
Hashset
Hashcons
CSet
diff --git a/checker/cic.mli b/checker/cic.mli
index 7ec345768..3304b032e 100644
--- a/checker/cic.mli
+++ b/checker/cic.mli
@@ -211,8 +211,6 @@ type projection_body = {
proj_npars : int;
proj_arg : int;
proj_type : constr; (* Type under params *)
- proj_eta : constr * constr; (* Eta-expanded term and type *)
- proj_body : constr; (* For compatibility, the match version *)
}
type constant_def =
@@ -241,7 +239,6 @@ type constant_body = {
const_type : constr;
const_body_code : to_patch_substituted;
const_universes : constant_universes;
- const_proj : bool;
const_inline_code : bool;
const_typing_flags : typing_flags;
}
diff --git a/checker/environ.ml b/checker/environ.ml
index 809150cea..3d5fac806 100644
--- a/checker/environ.ml
+++ b/checker/environ.ml
@@ -166,9 +166,6 @@ let evaluable_constant cst env =
try let _ = constant_value env (cst, Univ.Instance.empty) in true
with Not_found | NotEvaluableConst _ -> false
-let is_projection cst env =
- (lookup_constant cst env).const_proj
-
let lookup_projection p env =
Cmap_env.find (Projection.constant p) env.env_globals.env_projections
diff --git a/checker/environ.mli b/checker/environ.mli
index 4a7597249..acb29d7d2 100644
--- a/checker/environ.mli
+++ b/checker/environ.mli
@@ -58,7 +58,6 @@ exception NotEvaluableConst of const_evaluation_result
val constant_value : env -> Constant.t puniverses -> constr
val evaluable_constant : Constant.t -> env -> bool
-val is_projection : Constant.t -> env -> bool
val lookup_projection : Projection.t -> env -> projection_body
(* Inductives *)
diff --git a/checker/subtyping.ml b/checker/subtyping.ml
index 5c672d04a..f4ae02084 100644
--- a/checker/subtyping.ml
+++ b/checker/subtyping.ml
@@ -130,9 +130,7 @@ let check_inductive env mp1 l info1 mib2 spec2 subst1 subst2=
check (==) (fun x -> x.proj_npars);
check (==) (fun x -> x.proj_arg);
check (eq_constr) (fun x -> x.proj_type);
- check (eq_constr) (fun x -> fst x.proj_eta);
- check (eq_constr) (fun x -> snd x.proj_eta);
- check (eq_constr) (fun x -> x.proj_body); true
+ true
in
let check_inductive_type t1 t2 =
diff --git a/checker/values.ml b/checker/values.ml
index 67032bd1b..31e65729b 100644
--- a/checker/values.ml
+++ b/checker/values.ml
@@ -15,7 +15,7 @@
To ensure this file is up-to-date, 'make' now compares the md5 of cic.mli
with a copy we maintain here:
-MD5 fb80632357e3ffa988c6bba3fa6ade64 checker/cic.mli
+MD5 07651f61f86d91b22ff7056c6a8d86bc checker/cic.mli
*)
@@ -91,7 +91,7 @@ let rec v_mp = Sum("module_path",0,
[|[|v_dp|];
[|v_uid|];
[|v_mp;v_id|]|])
-let v_kn = v_tuple "kernel_name" [|Any;v_mp;v_dp;v_id;Int|]
+let v_kn = v_tuple "kernel_name" [|v_mp;v_dp;v_id;Int|]
let v_cst = v_sum "cst|mind" 0 [|[|v_kn|];[|v_kn;v_kn|]|]
let v_ind = v_tuple "inductive" [|v_cst;Int|]
let v_cons = v_tuple "constructor" [|v_ind;Int|]
@@ -225,9 +225,7 @@ let v_cst_def =
let v_projbody =
v_tuple "projection_body"
- [|v_cst;Int;Int;v_constr;
- v_tuple "proj_eta" [|v_constr;v_constr|];
- v_constr|]
+ [|v_cst;Int;Int;v_constr|]
let v_typing_flags =
v_tuple "typing_flags" [|v_bool; v_bool; v_oracle|]
@@ -241,7 +239,6 @@ let v_cb = v_tuple "constant_body"
Any;
v_const_univs;
v_bool;
- v_bool;
v_typing_flags|]
let v_recarg = v_sum "recarg" 1 (* Norec *)
diff --git a/clib/canary.ml b/clib/canary.ml
deleted file mode 100644
index b8b79ed7f..000000000
--- a/clib/canary.ml
+++ /dev/null
@@ -1,28 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-type t = Obj.t
-
-let obj = Obj.new_block Obj.closure_tag 0
- (** This is an empty closure block. In the current implementation, it is
- sufficient to allow marshalling but forbid equality. Sadly still allows
- hash. *)
- (** FIXME : use custom blocks somehow. *)
-
-module type Obj = sig type t end
-
-module Make(M : Obj) =
-struct
- type canary = t
- type t = (canary * M.t)
-
- let prj (_, x) = x
- let inj x = (obj, x)
-end
diff --git a/clib/canary.mli b/clib/canary.mli
deleted file mode 100644
index d993eabcf..000000000
--- a/clib/canary.mli
+++ /dev/null
@@ -1,27 +0,0 @@
-(************************************************************************)
-(* * The Coq Proof Assistant / The Coq Development Team *)
-(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
-(* <O___,, * (see CREDITS file for the list of authors) *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(* * (see LICENSE file for the text of the license) *)
-(************************************************************************)
-
-type t
-(** Type of canaries. Canaries are used to ensure that an object does not use
- generic operations. *)
-
-val obj : t
-(** Canary. In the current implementation, this object is marshallable,
- forbids generic comparison but still allows generic hashes. *)
-
-module type Obj = sig type t end
-
-module Make(M : Obj) :
-sig
- type t
- val prj : t -> M.t
- val inj : M.t -> t
-end
-(** Adds a canary to any type. *)
diff --git a/clib/clib.mllib b/clib/clib.mllib
index c9b4d72fc..afece4074 100644
--- a/clib/clib.mllib
+++ b/clib/clib.mllib
@@ -1,4 +1,3 @@
-Canary
CObj
CEphemeron
diff --git a/clib/hashcons.ml b/clib/hashcons.ml
index ec73c6d93..39969ebf7 100644
--- a/clib/hashcons.ml
+++ b/clib/hashcons.ml
@@ -10,8 +10,6 @@
(* Hash consing of datastructures *)
-(* The generic hash-consing functions (does not use Obj) *)
-
(* [t] is the type of object to hash-cons
* [u] is the type of hash-cons functions for the sub-structures
* of objects of type t (u usually has the form (t1->t1)*(t2->t2)*...).
@@ -148,41 +146,3 @@ module Hstring = Make(
let len = String.length s in
hash len s 0 0
end)
-
-(* Obj.t *)
-exception NotEq
-
-(* From CAMLLIB/caml/mlvalues.h *)
-let no_scan_tag = 251
-let tuple_p obj = Obj.is_block obj && (Obj.tag obj < no_scan_tag)
-
-let comp_obj o1 o2 =
- if tuple_p o1 && tuple_p o2 then
- let n1 = Obj.size o1 and n2 = Obj.size o2 in
- if n1=n2 then
- try
- for i = 0 to pred n1 do
- if not (Obj.field o1 i == Obj.field o2 i) then raise NotEq
- done; true
- with NotEq -> false
- else false
- else o1=o2
-
-let hash_obj hrec o =
- begin
- if tuple_p o then
- let n = Obj.size o in
- for i = 0 to pred n do
- Obj.set_field o i (hrec (Obj.field o i))
- done
- end;
- o
-
-module Hobj = Make(
- struct
- type t = Obj.t
- type u = (Obj.t -> Obj.t) * unit
- let hashcons (hrec,_) = hash_obj hrec
- let eq = comp_obj
- let hash = Hashtbl.hash
- end)
diff --git a/clib/hashcons.mli b/clib/hashcons.mli
index 3e396ff23..223dd2a4d 100644
--- a/clib/hashcons.mli
+++ b/clib/hashcons.mli
@@ -87,6 +87,3 @@ module Hstring : (S with type t = string and type u = unit)
module Hlist (D:HashedType) :
(S with type t = D.t list and type u = (D.t list -> D.t list)*(D.t->D.t))
(** Hashconsing of lists. *)
-
-module Hobj : (S with type t = Obj.t and type u = (Obj.t -> Obj.t) * unit)
-(** Hashconsing of OCaml values. *)
diff --git a/dev/ci/README.md b/dev/ci/README.md
index 665b3768a..08364c897 100644
--- a/dev/ci/README.md
+++ b/dev/ci/README.md
@@ -47,16 +47,13 @@ CI.
### Add your development by submitting a pull request
-Add a new `ci-mydev.sh` script to [`dev/ci`](.) (have a look at
-[`ci-coq-dpdgraph.sh`](ci-coq-dpdgraph.sh) or
-[`ci-fiat-parsers.sh`](ci-fiat-parsers.sh) for simple examples);
-set the corresponding variables in
-[`ci-basic-overlay.sh`](ci-basic-overlay.sh); add the corresponding
-target to [`Makefile.ci`](../../Makefile.ci); add new jobs to
-[`.gitlab-ci.yml`](../../.gitlab-ci.yml),
-[`.circleci/config.yml`](../../.circleci/config.yml) and
-[`.travis.yml`](../../.travis.yml) so that this new target is run. **Do not
-hesitate to submit an incomplete pull request if you need help to finish it.**
+Add a new `ci-mydev.sh` script to [`dev/ci`](.); set the corresponding
+variables in [`ci-basic-overlay.sh`](ci-basic-overlay.sh); add the
+corresponding target to [`Makefile.ci`](../../Makefile.ci) and a new job to
+[`.gitlab-ci.yml`](../../.gitlab-ci.yml) so that this new target is run.
+Have a look at [#7656](https://github.com/coq/coq/pull/7656/files) for an
+example. **Do not hesitate to submit an incomplete pull request if you need
+help to finish it.**
You may also be interested in having your development tested in our
performance benchmark. Currently this is done by providing an OPAM package
diff --git a/doc/sphinx/proof-engine/ssreflect-proof-language.rst b/doc/sphinx/proof-engine/ssreflect-proof-language.rst
index 3b2009657..1c554acdb 100644
--- a/doc/sphinx/proof-engine/ssreflect-proof-language.rst
+++ b/doc/sphinx/proof-engine/ssreflect-proof-language.rst
@@ -3657,7 +3657,8 @@ selective rewriting, blocking on the fly the reduction in the term ``t``.
.. coqtop:: reset
- From Coq Require Import ssreflect ssrfun ssrbool List.
+ From Coq Require Import ssreflect ssrfun ssrbool.
+ From Coq Require Import List.
Set Implicit Arguments.
Unset Strict Implicit.
Unset Printing Implicit Defensive.
diff --git a/doc/sphinx/proof-engine/tactics.rst b/doc/sphinx/proof-engine/tactics.rst
index 29c2f8b81..d0a0d568e 100644
--- a/doc/sphinx/proof-engine/tactics.rst
+++ b/doc/sphinx/proof-engine/tactics.rst
@@ -3949,9 +3949,20 @@ succeeds, and results in an error otherwise.
:name: constr_eq
This tactic checks whether its arguments are equal modulo alpha
- conversion and casts.
+ conversion, casts and universe constraints. It may unify universes.
.. exn:: Not equal.
+.. exn:: Not equal (due to universes).
+
+.. tacn:: constr_eq_strict @term @term
+ :name: constr_eq_strict
+
+ This tactic checks whether its arguments are equal modulo alpha
+ conversion, casts and universe constraints. It does not add new
+ constraints.
+
+.. exn:: Not equal.
+.. exn:: Not equal (due to universes).
.. tacn:: unify @term @term
:name: unify
diff --git a/engine/evd.ml b/engine/evd.ml
index 0c9c3a29b..f56f9662d 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -894,6 +894,9 @@ let check_eq evd s s' =
let check_leq evd s s' =
UGraph.check_leq (UState.ugraph evd.universes) s s'
+let check_constraints evd csts =
+ UGraph.check_constraints csts (UState.ugraph evd.universes)
+
let fix_undefined_variables evd =
{ evd with universes = UState.fix_undefined_variables evd.universes }
diff --git a/engine/evd.mli b/engine/evd.mli
index c40e925d8..405fcc403 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -552,6 +552,8 @@ val set_eq_instances : ?flex:bool ->
val check_eq : evar_map -> Univ.Universe.t -> Univ.Universe.t -> bool
val check_leq : evar_map -> Univ.Universe.t -> Univ.Universe.t -> bool
+val check_constraints : evar_map -> Univ.Constraint.t -> bool
+
val evar_universe_context : evar_map -> UState.t
val universe_context_set : evar_map -> Univ.ContextSet.t
val universe_subst : evar_map -> UnivSubst.universe_opt_subst
diff --git a/interp/declare.ml b/interp/declare.ml
index bc2d2068a..aa737239b 100644
--- a/interp/declare.ml
+++ b/interp/declare.ml
@@ -382,17 +382,34 @@ let inInductive : inductive_obj -> obj =
discharge_function = discharge_inductive;
rebuild_function = infer_inductive_subtyping }
-let declare_projections mind =
- let spec,_ = Inductive.lookup_mind_specif (Global.env ()) (mind,0) in
+let declare_projections univs mind =
+ let env = Global.env () in
+ let spec,_ = Inductive.lookup_mind_specif env (mind,0) in
match spec.mind_record with
- | Some (Some (_, kns, pjs)) ->
- Array.iteri (fun i kn ->
+ | Some (Some (_, kns, _)) ->
+ let projs = Inductiveops.compute_projections env (mind, 0) in
+ Array.iter2 (fun kn (term, types) ->
let id = Label.to_id (Constant.label kn) in
- let entry = {proj_entry_ind = mind; proj_entry_arg = i} in
- let kn' = declare_constant id (ProjectionEntry entry,
- IsDefinition StructureComponent)
- in
- assert(Constant.equal kn kn')) kns; true,true
+ let univs = match univs with
+ | Monomorphic_ind_entry _ ->
+ (** Global constraints already defined through the inductive *)
+ Monomorphic_const_entry Univ.ContextSet.empty
+ | Polymorphic_ind_entry ctx ->
+ Polymorphic_const_entry ctx
+ | Cumulative_ind_entry ctx ->
+ Polymorphic_const_entry (Univ.CumulativityInfo.univ_context ctx)
+ in
+ let term, types = match univs with
+ | Monomorphic_const_entry _ -> term, types
+ | Polymorphic_const_entry ctx ->
+ let u = Univ.UContext.instance ctx in
+ Vars.subst_instance_constr u term, Vars.subst_instance_constr u types
+ in
+ let entry = definition_entry ~types ~univs term in
+ let kn' = declare_constant id (DefinitionEntry entry, IsDefinition StructureComponent) in
+ assert (Constant.equal kn kn')
+ ) kns projs;
+ true, true
| Some None -> true,false
| None -> false,false
@@ -403,7 +420,7 @@ let declare_mind mie =
| [] -> anomaly (Pp.str "cannot declare an empty list of inductives.") in
let (sp,kn as oname) = add_leaf id (inInductive ([],mie)) in
let mind = Global.mind_of_delta_kn kn in
- let isrecord,isprim = declare_projections mind in
+ let isrecord,isprim = declare_projections mie.mind_entry_universes mind in
declare_mib_implicits mind;
declare_inductive_argument_scopes mind mie;
oname, isprim
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index 5783453e6..68057b389 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -156,7 +156,6 @@ type inline = bool
type result = {
cook_body : constant_def;
cook_type : types;
- cook_proj : bool;
cook_universes : constant_universes;
cook_inline : inline;
cook_context : Context.Named.t option;
@@ -230,7 +229,6 @@ let cook_constant ~hcons env { from = cb; info } =
{
cook_body = body;
cook_type = typ;
- cook_proj = cb.const_proj;
cook_universes = univs;
cook_inline = cb.const_inline_code;
cook_context = Some const_hyps;
diff --git a/kernel/cooking.mli b/kernel/cooking.mli
index 0d907f3de..76c79335f 100644
--- a/kernel/cooking.mli
+++ b/kernel/cooking.mli
@@ -21,7 +21,6 @@ type inline = bool
type result = {
cook_body : constant_def;
cook_type : types;
- cook_proj : bool;
cook_universes : constant_universes;
cook_inline : inline;
cook_context : Context.Named.t option;
diff --git a/kernel/declarations.ml b/kernel/declarations.ml
index 7bd70c050..7bd7d6c9c 100644
--- a/kernel/declarations.ml
+++ b/kernel/declarations.ml
@@ -54,8 +54,6 @@ type projection_body = {
proj_npars : int;
proj_arg : int; (** Projection index, starting from 0 *)
proj_type : types; (* Type under params *)
- proj_eta : constr * types; (* Eta-expanded term and type *)
- proj_body : constr; (* For compatibility with VMs only, the match version *)
}
(* Global declarations (i.e. constants) can be either: *)
@@ -87,7 +85,6 @@ type constant_body = {
const_type : types;
const_body_code : Cemitcodes.to_patch_substituted option;
const_universes : constant_universes;
- const_proj : bool;
const_inline_code : bool;
const_typing_flags : typing_flags; (** The typing options which
were used for
diff --git a/kernel/declareops.ml b/kernel/declareops.ml
index 75c0e5b4c..1b73096f7 100644
--- a/kernel/declareops.ml
+++ b/kernel/declareops.ml
@@ -86,7 +86,7 @@ let subst_const_def sub def = match def with
let subst_const_proj sub pb =
{ pb with proj_ind = subst_mind sub pb.proj_ind;
proj_type = subst_mps sub pb.proj_type;
- proj_body = subst_const_type sub pb.proj_body }
+ }
let subst_const_body sub cb =
assert (List.is_empty cb.const_hyps); (* we're outside sections *)
@@ -100,7 +100,6 @@ let subst_const_body sub cb =
{ const_hyps = [];
const_body = body';
const_type = type';
- const_proj = cb.const_proj;
const_body_code =
Option.map (Cemitcodes.subst_to_patch_subst sub) cb.const_body_code;
const_universes = cb.const_universes;
diff --git a/kernel/entries.ml b/kernel/entries.ml
index 94da00c7e..3c555f8c7 100644
--- a/kernel/entries.ml
+++ b/kernel/entries.ml
@@ -95,14 +95,9 @@ type inline = int option (* inlining level, None for no inlining *)
type parameter_entry =
Context.Named.t option * types in_constant_universes_entry * inline
-type projection_entry = {
- proj_entry_ind : MutInd.t;
- proj_entry_arg : int }
-
type 'a constant_entry =
| DefinitionEntry of 'a definition_entry
| ParameterEntry of parameter_entry
- | ProjectionEntry of projection_entry
(** {6 Modules } *)
diff --git a/kernel/environ.ml b/kernel/environ.ml
index fb89576dd..2d6c9117b 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -490,7 +490,7 @@ let lookup_projection cst env =
Cmap_env.find (Projection.constant cst) env.env_globals.env_projections
let is_projection cst env =
- (lookup_constant cst env).const_proj
+ Cmap_env.mem cst env.env_globals.env_projections
(* Mutual Inductives *)
let polymorphic_ind (mind,i) env =
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index 439acd15b..14f2a3d8f 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -797,16 +797,13 @@ exception UndefinableExpansion
build an expansion function.
The term built is expecting to be substituted first by
a substitution of the form [params, x : ind params] *)
-let compute_projections ((kn, _ as ind), u as indu) n x nparamargs params
+let compute_projections ((kn, _ as ind), u) nparamargs params
mind_consnrealdecls mind_consnrealargs paramslet ctx =
let mp, dp, l = MutInd.repr3 kn in
(** We build a substitution smashing the lets in the record parameters so
that typechecking projections requires just a substitution and not
matching with a parameter context. *)
- let indty, paramsletsubst =
- (* [ty] = [Ind inst] is typed in context [params] *)
- let inst = Context.Rel.to_extended_vect mkRel 0 paramslet in
- let ty = mkApp (mkIndU indu, inst) in
+ let paramsletsubst =
(* [Ind inst] is typed in context [params-wo-let] *)
let inst' = rel_list 0 nparamargs in
(* {params-wo-let |- subst:params] *)
@@ -814,48 +811,21 @@ let compute_projections ((kn, _ as ind), u as indu) n x nparamargs params
(* {params-wo-let, x:Ind inst' |- subst':(params,x:Ind inst)] *)
let subst = (* For the record parameter: *)
mkRel 1 :: List.map (lift 1) subst in
- ty, subst
+ subst
in
- let ci =
- let print_info =
- { ind_tags = []; cstr_tags = [|Context.Rel.to_tags ctx|]; style = LetStyle } in
- { ci_ind = ind;
- ci_npar = nparamargs;
- ci_cstr_ndecls = mind_consnrealdecls;
- ci_cstr_nargs = mind_consnrealargs;
- ci_pp_info = print_info }
- in
- let len = List.length ctx in
- let x = Name x in
- let compat_body ccl i =
- (* [ccl] is defined in context [params;x:indty] *)
- (* [ccl'] is defined in context [params;x:indty;x:indty] *)
- let ccl' = liftn 1 2 ccl in
- let p = mkLambda (x, lift 1 indty, ccl') in
- let branch = it_mkLambda_or_LetIn (mkRel (len - i)) ctx in
- let body = mkCase (ci, p, mkRel 1, [|lift 1 branch|]) in
- it_mkLambda_or_LetIn (mkLambda (x,indty,body)) params
- in
- let projections decl (i, j, kns, pbs, subst, letsubst) =
+ let projections decl (i, j, kns, pbs, letsubst) =
match decl with
| LocalDef (na,c,t) ->
(* From [params, field1,..,fieldj |- c(params,field1,..,fieldj)]
to [params, x:I, field1,..,fieldj |- c(params,field1,..,fieldj)] *)
let c = liftn 1 j c in
(* From [params, x:I, field1,..,fieldj |- c(params,field1,..,fieldj)]
- to [params, x:I |- c(params,proj1 x,..,projj x)] *)
- let c1 = substl subst c in
- (* From [params, x:I |- subst:field1,..,fieldj]
- to [params, x:I |- subst:field1,..,fieldj+1] where [subst]
- is represented with instance of field1 last *)
- let subst = c1 :: subst in
- (* From [params, x:I, field1,..,fieldj |- c(params,field1,..,fieldj)]
to [params-wo-let, x:I |- c(params,proj1 x,..,projj x)] *)
let c2 = substl letsubst c in
(* From [params-wo-let, x:I |- subst:(params, x:I, field1,..,fieldj)]
to [params-wo-let, x:I |- subst:(params, x:I, field1,..,fieldj+1)] *)
let letsubst = c2 :: letsubst in
- (i, j+1, kns, pbs, subst, letsubst)
+ (i, j+1, kns, pbs, letsubst)
| LocalAssum (na,t) ->
match na with
| Name id ->
@@ -868,21 +838,14 @@ let compute_projections ((kn, _ as ind), u as indu) n x nparamargs params
let projty = substl letsubst t in
(* from [params, x:I, field1,..,fieldj |- t(field1,..,fieldj)]
to [params, x:I |- t(proj1 x,..,projj x)] *)
- let ty = substl subst t in
- let term = mkProj (Projection.make kn true, mkRel 1) in
let fterm = mkProj (Projection.make kn false, mkRel 1) in
- let compat = compat_body ty (j - 1) in
- let etab = it_mkLambda_or_LetIn (mkLambda (x, indty, term)) params in
- let etat = it_mkProd_or_LetIn (mkProd (x, indty, ty)) params in
let body = { proj_ind = fst ind; proj_npars = nparamargs;
- proj_arg = i; proj_type = projty; proj_eta = etab, etat;
- proj_body = compat } in
- (i + 1, j + 1, kn :: kns, body :: pbs,
- fterm :: subst, fterm :: letsubst)
+ proj_arg = i; proj_type = projty; } in
+ (i + 1, j + 1, kn :: kns, body :: pbs, fterm :: letsubst)
| Anonymous -> raise UndefinableExpansion
in
- let (_, _, kns, pbs, subst, letsubst) =
- List.fold_right projections ctx (0, 1, [], [], [], paramsletsubst)
+ let (_, _, kns, pbs, letsubst) =
+ List.fold_right projections ctx (0, 1, [], [], paramsletsubst)
in
Array.of_list (List.rev kns),
Array.of_list (List.rev pbs)
@@ -987,7 +950,7 @@ let build_inductive env prv iu env_ar paramsctxt kn isrecord isfinite inds nmr r
(try
let fields, paramslet = List.chop pkt.mind_consnrealdecls.(0) rctx in
let kns, projs =
- compute_projections indsp pkt.mind_typename rid nparamargs paramsctxt
+ compute_projections indsp nparamargs paramsctxt
pkt.mind_consnrealdecls pkt.mind_consnrealargs paramslet fields
in Some (Some (rid, kns, projs))
with UndefinableExpansion -> Some None)
diff --git a/kernel/indtypes.mli b/kernel/indtypes.mli
index 5a38172c2..45228e35e 100644
--- a/kernel/indtypes.mli
+++ b/kernel/indtypes.mli
@@ -43,7 +43,7 @@ val check_inductive : env -> MutInd.t -> mutual_inductive_entry -> mutual_induct
val enforce_indices_matter : unit -> unit
val is_indices_matter : unit -> bool
-val compute_projections : pinductive -> Id.t -> Id.t ->
+val compute_projections : pinductive ->
int -> Context.Rel.t -> int array -> int array ->
Context.Rel.t -> Context.Rel.t ->
(Constant.t array * projection_body array)
diff --git a/kernel/names.ml b/kernel/names.ml
index 597061278..1d2a7c4ce 100644
--- a/kernel/names.ml
+++ b/kernel/names.ml
@@ -17,7 +17,7 @@
the module system, Aug 2002 *)
(* Abstraction over the type of constant for module inlining by Claudio
Sacerdoti, Nov 2004 *)
-(* Miscellaneous features or improvements by Hugo Herbelin,
+(* Miscellaneous features or improvements by Hugo Herbelin,
Élie Soubiran, ... *)
open Pp
@@ -364,7 +364,6 @@ module MPmap = CMap.Make(ModPath)
module KerName = struct
type t = {
- canary : Canary.t;
modpath : ModPath.t;
dirpath : DirPath.t;
knlabel : Label.t;
@@ -372,16 +371,14 @@ module KerName = struct
(** Lazily computed hash. If unset, it is set to negative values. *)
}
- let canary = Canary.obj
-
type kernel_name = t
let make modpath dirpath knlabel =
- { modpath; dirpath; knlabel; refhash = -1; canary; }
+ { modpath; dirpath; knlabel; refhash = -1; }
let repr kn = (kn.modpath, kn.dirpath, kn.knlabel)
let make2 modpath knlabel =
- { modpath; dirpath = DirPath.empty; knlabel; refhash = -1; canary; }
+ { modpath; dirpath = DirPath.empty; knlabel; refhash = -1; }
let modpath kn = kn.modpath
let label kn = kn.knlabel
@@ -437,7 +434,7 @@ module KerName = struct
* (string -> string)
let hashcons (hmod,hdir,hstr) kn =
let { modpath = mp; dirpath = dp; knlabel = l; refhash; } = kn in
- { modpath = hmod mp; dirpath = hdir dp; knlabel = hstr l; refhash; canary; }
+ { modpath = hmod mp; dirpath = hdir dp; knlabel = hstr l; refhash; }
let eq kn1 kn2 =
kn1.modpath == kn2.modpath && kn1.dirpath == kn2.dirpath &&
kn1.knlabel == kn2.knlabel
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index db1109e75..37bf679c5 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -250,7 +250,6 @@ let infer_declaration (type a) ~(trust : a trust) env (dcl : a constant_entry) =
{
Cooking.cook_body = Undef nl;
cook_type = t;
- cook_proj = false;
cook_universes = univs;
cook_inline = false;
cook_context = ctx;
@@ -291,7 +290,6 @@ let infer_declaration (type a) ~(trust : a trust) env (dcl : a constant_entry) =
{
Cooking.cook_body = def;
cook_type = typ;
- cook_proj = false;
cook_universes = Monomorphic_const univs;
cook_inline = c.const_entry_inline_code;
cook_context = c.const_entry_secctx;
@@ -343,39 +341,11 @@ let infer_declaration (type a) ~(trust : a trust) env (dcl : a constant_entry) =
{
Cooking.cook_body = def;
cook_type = typ;
- cook_proj = false;
cook_universes = univs;
cook_inline = c.const_entry_inline_code;
cook_context = c.const_entry_secctx;
}
- | ProjectionEntry {proj_entry_ind = ind; proj_entry_arg = i} ->
- let mib, _ = Inductive.lookup_mind_specif env (ind,0) in
- let kn, pb =
- match mib.mind_record with
- | Some (Some (id, kns, pbs)) ->
- if i < Array.length pbs then
- kns.(i), pbs.(i)
- else assert false
- | _ -> assert false
- in
- let univs =
- match mib.mind_universes with
- | Monomorphic_ind ctx -> Monomorphic_const ctx
- | Polymorphic_ind auctx -> Polymorphic_const auctx
- | Cumulative_ind acumi ->
- Polymorphic_const (Univ.ACumulativityInfo.univ_context acumi)
- in
- let term, typ = pb.proj_eta in
- {
- Cooking.cook_body = Def (Mod_subst.from_val (Constr.hcons term));
- cook_type = typ;
- cook_proj = true;
- cook_universes = univs;
- cook_inline = false;
- cook_context = None;
- }
-
let record_aux env s_ty s_bo =
let in_ty = keep_hyps env s_ty in
let v =
@@ -464,7 +434,6 @@ let build_constant_declaration kn env result =
{ const_hyps = hyps;
const_body = def;
const_type = typ;
- const_proj = result.cook_proj;
const_body_code = tps;
const_universes = univs;
const_inline_code = result.cook_inline;
diff --git a/library/heads.ml b/library/heads.ml
index 3d5f6a6ff..d9d650ac0 100644
--- a/library/heads.ml
+++ b/library/heads.ml
@@ -129,7 +129,7 @@ let compute_head = function
let cb = Environ.lookup_constant cst env in
let is_Def = function Declarations.Def _ -> true | _ -> false in
let body =
- if not cb.Declarations.const_proj && is_Def cb.Declarations.const_body
+ if not (Environ.is_projection cst env) && is_Def cb.Declarations.const_body
then Global.body_of_constant cst else None
in
(match body with
diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml
index 5aee70194..3a61c7747 100644
--- a/plugins/extraction/extraction.ml
+++ b/plugins/extraction/extraction.ml
@@ -1065,11 +1065,15 @@ let extract_constant env kn cb =
(match cb.const_body with
| Undef _ -> warn_info (); mk_typ_ax ()
| Def c ->
- (match cb.const_proj with
+ (match Environ.is_projection kn env with
| false -> mk_typ (get_body c)
| true ->
let pb = lookup_projection (Projection.make kn false) env in
- mk_typ (EConstr.of_constr pb.proj_body))
+ (** FIXME: handle mutual records *)
+ let ind = (pb.Declarations.proj_ind, 0) in
+ let bodies = Inductiveops.legacy_match_projection env ind in
+ let body = bodies.(pb.Declarations.proj_arg) in
+ mk_typ (EConstr.of_constr body))
| OpaqueDef c ->
add_opaque r;
if access_opaque () then mk_typ (get_opaque env c)
@@ -1078,11 +1082,15 @@ let extract_constant env kn cb =
(match cb.const_body with
| Undef _ -> warn_info (); mk_ax ()
| Def c ->
- (match cb.const_proj with
+ (match Environ.is_projection kn env with
| false -> mk_def (get_body c)
| true ->
let pb = lookup_projection (Projection.make kn false) env in
- mk_def (EConstr.of_constr pb.proj_body))
+ (** FIXME: handle mutual records *)
+ let ind = (pb.Declarations.proj_ind, 0) in
+ let bodies = Inductiveops.legacy_match_projection env ind in
+ let body = bodies.(pb.Declarations.proj_arg) in
+ mk_def (EConstr.of_constr body))
| OpaqueDef c ->
add_opaque r;
if access_opaque () then mk_def (get_opaque env c)
diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4
index f2899ab63..660e29ca8 100644
--- a/plugins/ltac/extratactics.ml4
+++ b/plugins/ltac/extratactics.ml4
@@ -793,17 +793,12 @@ END
(* ********************************************************************* *)
-let eq_constr x y =
- Proofview.Goal.enter begin fun gl ->
- let env = Tacmach.New.pf_env gl in
- let evd = Tacmach.New.project gl in
- match EConstr.eq_constr_universes env evd x y with
- | Some _ -> Proofview.tclUNIT ()
- | None -> Tacticals.New.tclFAIL 0 (str "Not equal")
- end
-
TACTIC EXTEND constr_eq
-| [ "constr_eq" constr(x) constr(y) ] -> [ eq_constr x y ]
+| [ "constr_eq" constr(x) constr(y) ] -> [ Tactics.constr_eq ~strict:false x y ]
+END
+
+TACTIC EXTEND constr_eq_strict
+| [ "constr_eq_strict" constr(x) constr(y) ] -> [ Tactics.constr_eq ~strict:true x y ]
END
TACTIC EXTEND constr_eq_nounivs
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 0dd3c5944..93ca9dc5e 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -1425,7 +1425,7 @@ and match_current pb (initial,tomatch) =
let ci = make_case_info pb.env (fst mind) pb.casestyle in
let pred = nf_betaiota pb.env !(pb.evdref) pred in
let case =
- make_case_or_project pb.env !(pb.evdref) indf ci pred current brvals
+ make_case_or_project pb.env !(pb.evdref) indf ci pred current brvals
in
let _ = Evarutil.evd_comb1 (Typing.type_of pb.env) pb.evdref pred in
Typing.check_allowed_sort pb.env !(pb.evdref) mind current pred;
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 019ec3d27..fe49d64c7 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -690,7 +690,10 @@ and detype_r d flags avoid env sigma t =
let c' =
try
let pb = Environ.lookup_projection p (snd env) in
- let body = pb.Declarations.proj_body in
+ (** FIXME: handle mutual records *)
+ let ind = (pb.Declarations.proj_ind, 0) in
+ let bodies = Inductiveops.legacy_match_projection (snd env) ind in
+ let body = bodies.(pb.Declarations.proj_arg) in
let ty = Retyping.get_type_of (snd env) sigma c in
let ((ind,u), args) = Inductiveops.find_mrectype (snd env) sigma ty in
let body' = strip_lam_assum body in
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index 27b029aad..4ab932723 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -304,7 +304,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
process_constr env 0 f (List.rev cstr.cs_args, recargs)
(* Main function *)
-let mis_make_indrec env sigma listdepkind mib u =
+let mis_make_indrec env sigma ?(force_mutual=false) listdepkind mib u =
let nparams = mib.mind_nparams in
let nparrec = mib.mind_nparams_rec in
let evdref = ref sigma in
@@ -469,7 +469,7 @@ let mis_make_indrec env sigma listdepkind mib u =
(* Body on make_one_rec *)
let ((indi,u),mibi,mipi,dep,kind) = List.nth listdepkind p in
- if (mis_is_recursive_subset
+ if force_mutual || (mis_is_recursive_subset
(List.map (fun ((indi,u),_,_,_,_) -> snd indi) listdepkind)
mipi.mind_recargs)
then
@@ -558,7 +558,7 @@ let check_arities env listdepkind =
[] listdepkind
in true
-let build_mutual_induction_scheme env sigma = function
+let build_mutual_induction_scheme env sigma ?(force_mutual=false) = function
| ((mind,u),dep,s)::lrecspec ->
let (mib,mip) = lookup_mind_specif env mind in
if dep && not (Inductiveops.has_dependent_elim mib) then
@@ -577,7 +577,7 @@ let build_mutual_induction_scheme env sigma = function
lrecspec)
in
let _ = check_arities env listdepkind in
- mis_make_indrec env sigma listdepkind mib u
+ mis_make_indrec env sigma ~force_mutual listdepkind mib u
| _ -> anomaly (Pp.str "build_induction_scheme expects a non empty list of inductive types.")
let build_induction_scheme env sigma pind dep kind =
diff --git a/pretyping/indrec.mli b/pretyping/indrec.mli
index d87a19d28..de9d3a0ab 100644
--- a/pretyping/indrec.mli
+++ b/pretyping/indrec.mli
@@ -47,7 +47,8 @@ val build_induction_scheme : env -> evar_map -> pinductive ->
(** Builds mutual (recursive) induction schemes *)
val build_mutual_induction_scheme :
- env -> evar_map -> (pinductive * dep_flag * Sorts.family) list -> evar_map * constr list
+ env -> evar_map -> ?force_mutual:bool ->
+ (pinductive * dep_flag * Sorts.family) list -> evar_map * constr list
(** Scheme combinators *)
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index b1ab2d2b7..1003f86c5 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -356,8 +356,8 @@ let make_case_or_project env sigma indf ci pred c branches =
| None -> (mkCase (ci, pred, c, branches))
| Some ps ->
assert(Array.length branches == 1);
+ let na, ty, t = destLambda sigma pred in
let () =
- let _, _, t = destLambda sigma pred in
let (ind, _), _ = dest_ind_family indf in
let mib, _ = Inductive.lookup_mind_specif env ind in
if (* dependent *) not (Vars.noccurn sigma 1 t) &&
@@ -368,16 +368,18 @@ let make_case_or_project env sigma indf ci pred c branches =
in
let branch = branches.(0) in
let ctx, br = decompose_lam_n_assum sigma (Array.length ps) branch in
- let n, subst =
+ let n, len, ctx =
List.fold_right
- (fun decl (i, subst) ->
+ (fun decl (i, j, ctx) ->
match decl with
- | LocalAssum (na, t) ->
- let t = mkProj (Projection.make ps.(i) true, c) in
- (i + 1, t :: subst)
- | LocalDef (na, b, t) -> (i, Vars.substl subst b :: subst))
- ctx (0, [])
- in Vars.substl subst br
+ | LocalAssum (na, ty) ->
+ let t = mkProj (Projection.make ps.(i) true, mkRel j) in
+ (i + 1, j + 1, LocalDef (na, t, Vars.liftn 1 j ty) :: ctx)
+ | LocalDef (na, b, ty) ->
+ (i, j + 1, LocalDef (na, Vars.liftn 1 j b, Vars.liftn 1 j ty) :: ctx))
+ ctx (0, 1, [])
+ in
+ mkLetIn (na, c, ty, it_mkLambda_or_LetIn (Vars.liftn 1 (Array.length ps + 1) br) ctx)
(* substitution in a signature *)
@@ -454,6 +456,110 @@ let build_branch_type env sigma dep p cs =
(**************************************************)
+(** From a rel context describing the constructor arguments,
+ build an expansion function.
+ The term built is expecting to be substituted first by
+ a substitution of the form [params, x : ind params] *)
+let compute_projections env (kn, _ as ind) =
+ let open Term in
+ let mib = Environ.lookup_mind kn env in
+ let indu = match mib.mind_universes with
+ | Monomorphic_ind _ -> mkInd ind
+ | Polymorphic_ind ctx -> mkIndU (ind, make_abstract_instance ctx)
+ | Cumulative_ind ctx ->
+ mkIndU (ind, make_abstract_instance (ACumulativityInfo.univ_context ctx))
+ in
+ let x = match mib.mind_record with
+ | None | Some None ->
+ anomaly Pp.(str "Trying to build primitive projections for a non-primitive record")
+ | Some (Some (id, _, _)) -> Name id
+ in
+ (** FIXME: handle mutual records *)
+ let pkt = mib.mind_packets.(0) in
+ let { mind_consnrealargs; mind_consnrealdecls } = pkt in
+ let { mind_nparams = nparamargs; mind_params_ctxt = params } = mib in
+ let rctx, _ = decompose_prod_assum (subst1 indu pkt.mind_nf_lc.(0)) in
+ let ctx, paramslet = List.chop pkt.mind_consnrealdecls.(0) rctx in
+ let mp, dp, l = MutInd.repr3 kn in
+ (** We build a substitution smashing the lets in the record parameters so
+ that typechecking projections requires just a substitution and not
+ matching with a parameter context. *)
+ let indty =
+ (* [ty] = [Ind inst] is typed in context [params] *)
+ let inst = Context.Rel.to_extended_vect mkRel 0 paramslet in
+ let ty = mkApp (indu, inst) in
+ (* [Ind inst] is typed in context [params-wo-let] *)
+ ty
+ in
+ let ci =
+ let print_info =
+ { ind_tags = []; cstr_tags = [|Context.Rel.to_tags ctx|]; style = LetStyle } in
+ { ci_ind = ind;
+ ci_npar = nparamargs;
+ ci_cstr_ndecls = mind_consnrealdecls;
+ ci_cstr_nargs = mind_consnrealargs;
+ ci_pp_info = print_info }
+ in
+ let len = List.length ctx in
+ let compat_body ccl i =
+ (* [ccl] is defined in context [params;x:indty] *)
+ (* [ccl'] is defined in context [params;x:indty;x:indty] *)
+ let ccl' = liftn 1 2 ccl in
+ let p = mkLambda (x, lift 1 indty, ccl') in
+ let branch = it_mkLambda_or_LetIn (mkRel (len - i)) ctx in
+ let body = mkCase (ci, p, mkRel 1, [|lift 1 branch|]) in
+ it_mkLambda_or_LetIn (mkLambda (x,indty,body)) params
+ in
+ let projections decl (j, pbs, subst) =
+ match decl with
+ | LocalDef (na,c,t) ->
+ (* From [params, field1,..,fieldj |- c(params,field1,..,fieldj)]
+ to [params, x:I, field1,..,fieldj |- c(params,field1,..,fieldj)] *)
+ let c = liftn 1 j c in
+ (* From [params, x:I, field1,..,fieldj |- c(params,field1,..,fieldj)]
+ to [params, x:I |- c(params,proj1 x,..,projj x)] *)
+ let c1 = substl subst c in
+ (* From [params, x:I |- subst:field1,..,fieldj]
+ to [params, x:I |- subst:field1,..,fieldj+1] where [subst]
+ is represented with instance of field1 last *)
+ let subst = c1 :: subst in
+ (j+1, pbs, subst)
+ | LocalAssum (na,t) ->
+ match na with
+ | Name id ->
+ let kn = Constant.make1 (KerName.make mp dp (Label.of_id id)) in
+ (* from [params, field1,..,fieldj |- t(params,field1,..,fieldj)]
+ to [params, x:I, field1,..,fieldj |- t(params,field1,..,fieldj] *)
+ let t = liftn 1 j t in
+ (* from [params, x:I, field1,..,fieldj |- t(params,field1,..,fieldj)]
+ to [params-wo-let, x:I |- t(params,proj1 x,..,projj x)] *)
+ (* from [params, x:I, field1,..,fieldj |- t(field1,..,fieldj)]
+ to [params, x:I |- t(proj1 x,..,projj x)] *)
+ let ty = substl subst t in
+ let term = mkProj (Projection.make kn true, mkRel 1) in
+ let fterm = mkProj (Projection.make kn false, mkRel 1) in
+ let compat = compat_body ty (j - 1) in
+ let etab = it_mkLambda_or_LetIn (mkLambda (x, indty, term)) params in
+ let etat = it_mkProd_or_LetIn (mkProd (x, indty, ty)) params in
+ let body = (etab, etat, compat) in
+ (j + 1, body :: pbs, fterm :: subst)
+ | Anonymous ->
+ anomaly Pp.(str "Trying to build primitive projections for a non-primitive record")
+ in
+ let (_, pbs, subst) =
+ List.fold_right projections ctx (1, [], [])
+ in
+ Array.rev_of_list pbs
+
+let legacy_match_projection env ind =
+ Array.map pi3 (compute_projections env ind)
+
+let compute_projections ind mib =
+ let ans = compute_projections ind mib in
+ Array.map (fun (prj, ty, _) -> (prj, ty)) ans
+
+(**************************************************)
+
let extract_mrectype sigma t =
let open EConstr in
let (t, l) = decompose_app sigma t in
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index b0d714b03..aa53f7e67 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -194,6 +194,18 @@ val make_case_or_project :
val make_default_case_info : env -> case_style -> inductive -> case_info
i*)
+val compute_projections : Environ.env -> inductive -> (constr * types) array
+(** Given a primitive record type, for every field computes the eta-expanded
+ projection and its type. *)
+
+val legacy_match_projection : Environ.env -> inductive -> constr array
+(** Given a record type, computes the legacy match-based projection of the
+ projections.
+
+ BEWARE: such terms are ill-typed, and should thus only be used in upper
+ layers. The kernel will probably badly fail if presented with one of
+ those. *)
+
(********************)
val type_of_inductive_knowing_conclusion :
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 770e31fea..c430edf2e 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -5034,6 +5034,26 @@ let tclABSTRACT ?(opaque=true) name_op tac =
else name_op_to_name name_op (DefinitionBody Definition) "_subterm" in
abstract_subproof ~opaque s gk tac
+let constr_eq ~strict x y =
+ let fail = Tacticals.New.tclFAIL 0 (str "Not equal") in
+ let fail_universes = Tacticals.New.tclFAIL 0 (str "Not equal (due to universes)") in
+ Proofview.Goal.enter begin fun gl ->
+ let env = Tacmach.New.pf_env gl in
+ let evd = Tacmach.New.project gl in
+ match EConstr.eq_constr_universes env evd x y with
+ | Some csts ->
+ let csts = UnivProblem.to_constraints ~force_weak:false (Evd.universes evd) csts in
+ if strict then
+ if Evd.check_constraints evd csts then Proofview.tclUNIT ()
+ else fail_universes
+ else
+ (match Evd.add_constraints evd csts with
+ | evd -> Proofview.Unsafe.tclEVARS evd
+ | exception Univ.UniverseInconsistency _ ->
+ fail_universes)
+ | None -> fail
+ end
+
let unify ?(state=full_transparent_state) x y =
Proofview.Goal.enter begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index 8d4302450..57f20d2ff 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -409,6 +409,11 @@ val generalize_dep : ?with_let:bool (** Don't lose let bindings *) -> constr -
(** {6 Other tactics. } *)
+(** Syntactic equality up to universes. With [strict] the universe
+ constraints must be already true to succeed, without [strict] they
+ are added to the evar map. *)
+val constr_eq : strict:bool -> constr -> constr -> unit Proofview.tactic
+
val unify : ?state:Names.transparent_state -> constr -> constr -> unit Proofview.tactic
val cache_term_by_tactic_then : opaque:bool -> ?goal_type:(constr option) -> Id.t -> Decl_kinds.goal_kind -> unit Proofview.tactic -> (constr -> constr list -> unit Proofview.tactic) -> unit Proofview.tactic
diff --git a/test-suite/bugs/closed/7421.v b/test-suite/bugs/closed/7421.v
new file mode 100644
index 000000000..afcdd35fc
--- /dev/null
+++ b/test-suite/bugs/closed/7421.v
@@ -0,0 +1,39 @@
+
+
+Universe i j.
+
+Goal False.
+Proof.
+ Check Type@{i} : Type@{j}.
+ Fail constr_eq_strict Type@{i} Type@{j}.
+ assert_succeeds constr_eq Type@{i} Type@{j}. (* <- i=j is forgotten after assert_succeeds *)
+ Fail constr_eq_strict Type@{i} Type@{j}.
+
+ constr_eq Type@{i} Type@{j}. (* <- i=j is retained *)
+ constr_eq_strict Type@{i} Type@{j}.
+ Fail Check Type@{i} : Type@{j}.
+
+ Fail constr_eq Prop Set.
+ Fail constr_eq Prop Type.
+
+ Fail constr_eq_strict Type Type.
+ constr_eq Type Type.
+
+ constr_eq_strict Set Set.
+ constr_eq Set Set.
+ constr_eq Prop Prop.
+
+ let x := constr:(Type) in constr_eq_strict x x.
+ let x := constr:(Type) in constr_eq x x.
+
+ Fail lazymatch type of prod with
+ | ?A -> ?B -> _ => constr_eq_strict A B
+ end.
+ lazymatch type of prod with
+ | ?A -> ?B -> _ => constr_eq A B
+ end.
+ lazymatch type of prod with
+ | ?A -> ?B -> ?C => constr_eq A C
+ end.
+
+Abort.
diff --git a/test-suite/success/letproj.v b/test-suite/success/letproj.v
index de2857b43..2f0d8bf8c 100644
--- a/test-suite/success/letproj.v
+++ b/test-suite/success/letproj.v
@@ -7,3 +7,5 @@ Definition test (A : Type) (f : Foo A) :=
Scheme foo_case := Case for Foo Sort Type.
+Definition test' (A : Type) (f : Foo A) :=
+ let 'Build_Foo _ x y := f in x.
diff --git a/test-suite/success/primitiveproj.v b/test-suite/success/primitiveproj.v
index 31a1608c4..7ca2767a5 100644
--- a/test-suite/success/primitiveproj.v
+++ b/test-suite/success/primitiveproj.v
@@ -199,3 +199,24 @@ split.
reflexivity.
Qed.
*)
+
+(* Primitive projection match compilation *)
+Require Import List.
+Set Primitive Projections.
+
+Record prod (A B : Type) := pair { fst : A ; snd : B }.
+Arguments pair {_ _} _ _.
+
+Fixpoint split_at {A} (l : list A) (n : nat) : prod (list A) (list A) :=
+ match n with
+ | 0 => pair nil l
+ | S n =>
+ match l with
+ | nil => pair nil nil
+ | x :: l => let 'pair l1 l2 := split_at l n in pair (x :: l1) l2
+ end
+ end.
+
+Time Eval vm_compute in split_at (repeat 0 20) 10. (* Takes 0s *)
+Time Eval vm_compute in split_at (repeat 0 40) 20. (* Takes 0.001s *)
+Timeout 1 Time Eval vm_compute in split_at (repeat 0 60) 30. (* Used to take 60s, now takes 0.001s *)
diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml
index 2deca1e06..e86e10877 100644
--- a/vernac/indschemes.ml
+++ b/vernac/indschemes.ml
@@ -370,7 +370,7 @@ requested
| InductionScheme (x,y,z) -> names "_ind" "_rec" x y z
| EqualityScheme x -> l1,((None,smart_global_inductive x)::l2)
-let do_mutual_induction_scheme lnamedepindsort =
+let do_mutual_induction_scheme ?(force_mutual=false) lnamedepindsort =
let lrecnames = List.map (fun ({CAst.v},_,_,_) -> v) lnamedepindsort
and env0 = Global.env() in
let sigma, lrecspec, _ =
@@ -388,7 +388,7 @@ let do_mutual_induction_scheme lnamedepindsort =
(evd, (indu,dep,sort) :: l, inst))
lnamedepindsort (Evd.from_env env0,[],None)
in
- let sigma, listdecl = Indrec.build_mutual_induction_scheme env0 sigma lrecspec in
+ let sigma, listdecl = Indrec.build_mutual_induction_scheme env0 sigma ~force_mutual lrecspec in
let declare decl fi lrecref =
let decltype = Retyping.get_type_of env0 sigma (EConstr.of_constr decl) in
let decltype = EConstr.to_constr sigma decltype in
diff --git a/vernac/indschemes.mli b/vernac/indschemes.mli
index 261c3d813..ebfc76de9 100644
--- a/vernac/indschemes.mli
+++ b/vernac/indschemes.mli
@@ -29,9 +29,13 @@ val declare_congr_scheme : inductive -> unit
val declare_rewriting_schemes : inductive -> unit
-(** Mutual Minimality/Induction scheme *)
+(** Mutual Minimality/Induction scheme.
+ [force_mutual] forces the construction of eliminators having the same predicates and
+ methods even if some of the inductives are not recursive.
+ By default it is [false] and some of the eliminators are defined as simple case analysis.
+ *)
-val do_mutual_induction_scheme :
+val do_mutual_induction_scheme : ?force_mutual:bool ->
(lident * bool * inductive * Sorts.family) list -> unit
(** Main calls to interpret the Scheme command *)