aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <matthieu.sozeau@inria.fr>2014-08-29 00:50:08 +0200
committerGravatar Matthieu Sozeau <matthieu.sozeau@inria.fr>2014-08-29 00:50:08 +0200
commite368bcd7e16fda4d011ad2c960c647c7da72bcb6 (patch)
tree352297625b07485c667250b691b20a12cf8663ef
parent23f064547758a491bb7cb709797c2b1338a17558 (diff)
Add test-suite file. Compute the name for the record binder in the
eta-expanded version of a projection as before.
-rw-r--r--kernel/entries.mli3
-rw-r--r--kernel/indtypes.ml6
-rw-r--r--kernel/indtypes.mli3
-rw-r--r--test-suite/bugs/closed/3402.v7
4 files changed, 14 insertions, 5 deletions
diff --git a/kernel/entries.mli b/kernel/entries.mli
index 5fa0527ef..0a06eef16 100644
--- a/kernel/entries.mli
+++ b/kernel/entries.mli
@@ -42,7 +42,8 @@ type one_inductive_entry = {
type mutual_inductive_entry = {
mind_entry_record : bool option;
- (** Some true: primitive record, Some false: non-primitive record *)
+ (** Some true: primitive record
+ Some false: non-primitive record *)
mind_entry_finite : bool;
mind_entry_params : (Id.t * local_entry) list;
mind_entry_inds : one_inductive_entry list;
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index 715bdb5da..7ae787d22 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -661,7 +661,7 @@ 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 indsp) nparamargs params
+let compute_projections ((kn, _ as ind), u as indsp) n nparamargs params
mind_consnrealdecls mind_consnrealargs ctx =
let mp, dp, l = repr_mind kn in
let rp = mkApp (mkIndU indsp, rel_vect 0 nparamargs) in
@@ -674,7 +674,7 @@ let compute_projections ((kn, _ as ind), u as indsp) nparamargs params
ci_pp_info = print_info }
in
let len = List.length ctx in
- let x = Name (id_of_string "r") in
+ let x = Name (Id.of_string (Unicode.lowercase_first_char (Id.to_string n))) in
let compat_body ccl i =
(* [ccl] is defined in context [params;x:rp] *)
(* [ccl'] is defined in context [params;x:rp;x:rp] *)
@@ -786,7 +786,7 @@ let build_inductive env p prv ctx env_ar params kn isrecord isfinite inds nmr re
(try
let fields = List.firstn pkt.mind_consnrealdecls.(0) rctx in
let kns, projs =
- compute_projections ((kn, 0), u) nparamargs params
+ compute_projections ((kn, 0), u) pkt.mind_typename nparamargs params
pkt.mind_consnrealdecls pkt.mind_consnrealargs fields
in Some (kns, projs)
with UndefinableExpansion -> Some ([||], [||]))
diff --git a/kernel/indtypes.mli b/kernel/indtypes.mli
index 85e6a6327..0d5b03404 100644
--- a/kernel/indtypes.mli
+++ b/kernel/indtypes.mli
@@ -41,6 +41,7 @@ val check_inductive : env -> mutual_inductive -> mutual_inductive_entry -> mutua
val enforce_indices_matter : unit -> unit
val is_indices_matter : unit -> bool
-val compute_projections : pinductive -> int -> Context.rel_context -> int array -> int array ->
+val compute_projections : pinductive -> Id.t ->
+ int -> Context.rel_context -> int array -> int array ->
Context.rel_context ->
(constant array * projection_body array)
diff --git a/test-suite/bugs/closed/3402.v b/test-suite/bugs/closed/3402.v
new file mode 100644
index 000000000..ed47ec825
--- /dev/null
+++ b/test-suite/bugs/closed/3402.v
@@ -0,0 +1,7 @@
+Set Primitive Projections.
+Record prod A B := pair { fst : A ; snd : B }.
+Goal forall A B (p : prod A B), p = let (x, y) := p in pair A B x y.
+Proof.
+ intros A B p.
+ exact eq_refl.
+Qed. \ No newline at end of file