aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--API/API.mli2
-rw-r--r--plugins/extraction/extraction.ml6
-rw-r--r--test-suite/bugs/closed/4709.v18
3 files changed, 25 insertions, 1 deletions
diff --git a/API/API.mli b/API/API.mli
index 3ae8ac64c..a0e77edd1 100644
--- a/API/API.mli
+++ b/API/API.mli
@@ -169,6 +169,8 @@ sig
val map : (Constant.t -> Constant.t) -> t -> t
val constant : t -> Constant.t
val equal : t -> t -> bool
+ val unfolded : t -> bool
+ val unfold : t -> t
end
type evaluable_global_reference =
diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml
index 661842790..7644b49ce 100644
--- a/plugins/extraction/extraction.ml
+++ b/plugins/extraction/extraction.ml
@@ -295,7 +295,11 @@ let rec extract_type env db j c args =
| 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
- | Case _ | Fix _ | CoFix _ | Proj _ -> Tunknown
+ | 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 (Term.mkProj (Projection.unfold p, t)) args
+ | Case _ | Fix _ | CoFix _ -> Tunknown
| _ -> assert false
(*s Auxiliary function dealing with type application.
diff --git a/test-suite/bugs/closed/4709.v b/test-suite/bugs/closed/4709.v
new file mode 100644
index 000000000..a9edcc804
--- /dev/null
+++ b/test-suite/bugs/closed/4709.v
@@ -0,0 +1,18 @@
+
+(** Bug 4709 https://coq.inria.fr/bug/4709
+ Extraction wasn't reducing primitive projections in types. *)
+
+Require Extraction.
+
+Set Primitive Projections.
+
+Record t := Foo { foo : Type }.
+Definition ty := foo (Foo nat).
+
+(* Without proper reduction of primitive projections in
+ [extract_type], the type [ty] was extracted as [Tunknown].
+ Let's check it isn't the case anymore. *)
+
+Parameter check : nat.
+Extract Constant check => "(O:ty)".
+Extraction TestCompile ty check.