aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--kernel/reduction.ml12
-rw-r--r--kernel/reduction.mli6
-rw-r--r--printing/printmod.ml10
-rw-r--r--test-suite/output/Inductive.out4
-rw-r--r--test-suite/output/Inductive.v4
5 files changed, 32 insertions, 4 deletions
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index b0f4a1e5f..7a4f25e63 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -938,6 +938,18 @@ let hnf_prod_app env t n =
let hnf_prod_applist env t nl =
List.fold_left (hnf_prod_app env) t nl
+let hnf_prod_applist_assum env n c l =
+ let rec app n subst t l =
+ if Int.equal n 0 then
+ if l == [] then substl subst t
+ else anomaly (Pp.str "Too many arguments.")
+ else match kind (whd_allnolet env t), l with
+ | Prod(_,_,c), arg::l -> app (n-1) (arg::subst) c l
+ | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l
+ | _, [] -> anomaly (Pp.str "Not enough arguments.")
+ | _ -> anomaly (Pp.str "Not enough prod/let's.") in
+ app n [] c l
+
(* Dealing with arities *)
let dest_prod env =
diff --git a/kernel/reduction.mli b/kernel/reduction.mli
index 573e4c8bd..ce019f021 100644
--- a/kernel/reduction.mli
+++ b/kernel/reduction.mli
@@ -103,6 +103,12 @@ val beta_app : constr -> constr -> constr
(** Pseudo-reduction rule Prod(x,A,B) a --> B[x\a] *)
val hnf_prod_applist : env -> types -> constr list -> types
+(** In [hnf_prod_applist_assum n c args], [c] is supposed to (whd-)reduce to
+ the form [∀Γ.t] with [Γ] of length [n] and possibly with let-ins; it
+ returns [t] with the assumptions of [Γ] instantiated by [args] and
+ the local definitions of [Γ] expanded. *)
+val hnf_prod_applist_assum : env -> int -> types -> constr list -> types
+
(** Compatibility alias for Term.lambda_appvect_assum *)
val betazeta_appvect : int -> constr -> constr array -> constr
diff --git a/printing/printmod.ml b/printing/printmod.ml
index 05292b06b..bed7c39fe 100644
--- a/printing/printmod.ml
+++ b/printing/printmod.ml
@@ -93,10 +93,11 @@ let print_one_inductive env sigma mib ((_,i) as ind) =
else Univ.Instance.empty in
let mip = mib.mind_packets.(i) in
let params = Inductive.inductive_paramdecls (mib,u) in
+ let nparamdecls = Context.Rel.length params in
let args = Context.Rel.to_extended_list mkRel 0 params in
- let arity = hnf_prod_applist env (build_ind_type env ((mib,mip),u)) args in
+ let arity = hnf_prod_applist_assum env nparamdecls (build_ind_type env ((mib,mip),u)) args in
let cstrtypes = Inductive.type_of_constructors (ind,u) (mib,mip) in
- let cstrtypes = Array.map (fun c -> hnf_prod_applist env c args) cstrtypes in
+ let cstrtypes = Array.map (fun c -> hnf_prod_applist_assum env nparamdecls c args) cstrtypes in
let envpar = push_rel_context params env in
let inst =
if Declareops.inductive_is_polymorphic mib then
@@ -174,10 +175,11 @@ let print_record env mind mib udecl =
in
let mip = mib.mind_packets.(0) in
let params = Inductive.inductive_paramdecls (mib,u) in
+ let nparamdecls = Context.Rel.length params in
let args = Context.Rel.to_extended_list mkRel 0 params in
- let arity = hnf_prod_applist env (build_ind_type env ((mib,mip),u)) args in
+ let arity = hnf_prod_applist_assum env nparamdecls (build_ind_type env ((mib,mip),u)) args in
let cstrtypes = Inductive.type_of_constructors ((mind,0),u) (mib,mip) in
- let cstrtype = hnf_prod_applist env cstrtypes.(0) args in
+ let cstrtype = hnf_prod_applist_assum env nparamdecls cstrtypes.(0) args in
let fields = get_fields cstrtype in
let envpar = push_rel_context params env in
let bl = Universes.universe_binders_with_opt_names (IndRef (mind,0))
diff --git a/test-suite/output/Inductive.out b/test-suite/output/Inductive.out
index e912003f0..af202ea01 100644
--- a/test-suite/output/Inductive.out
+++ b/test-suite/output/Inductive.out
@@ -1,3 +1,7 @@
The command has indeed failed with message:
Last occurrence of "list'" must have "A" as 1st argument in
"A -> list' A -> list' (A * A)%type".
+Inductive foo (A : Type) (x : A) (y : A := x) : Prop := Foo : foo A x
+
+For foo: Argument scopes are [type_scope _]
+For Foo: Argument scopes are [type_scope _]
diff --git a/test-suite/output/Inductive.v b/test-suite/output/Inductive.v
index 8db8956e3..8ff91268a 100644
--- a/test-suite/output/Inductive.v
+++ b/test-suite/output/Inductive.v
@@ -1,3 +1,7 @@
Fail Inductive list' (A:Set) : Set :=
| nil' : list' A
| cons' : A -> list' A -> list' (A*A).
+
+(* Check printing of let-ins *)
+Inductive foo (A : Type) (x : A) (y := x) := Foo.
+Print foo.