aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--plugins/extraction/mlutil.ml2
-rw-r--r--test-suite/bugs/closed/4844.v47
2 files changed, 48 insertions, 1 deletions
diff --git a/plugins/extraction/mlutil.ml b/plugins/extraction/mlutil.ml
index f1bcde2f3..a4c2bcd88 100644
--- a/plugins/extraction/mlutil.ml
+++ b/plugins/extraction/mlutil.ml
@@ -120,7 +120,6 @@ let rec mgu = function
mgu (a, a'); mgu (b, b')
| Tglob (r,l), Tglob (r',l') when Globnames.eq_gr r r' ->
List.iter mgu (List.combine l l')
- | (Tdummy _, _ | _, Tdummy _) when lang() == Haskell -> ()
| Tdummy _, Tdummy _ -> ()
| Tvar i, Tvar j when Int.equal i j -> ()
| Tvar' i, Tvar' j when Int.equal i j -> ()
@@ -1052,6 +1051,7 @@ let rec simpl o = function
| MLmagic(MLcase(typ,e,br)) ->
let br' = Array.map (fun (ids,p,c) -> (ids,p,MLmagic c)) br in
simpl o (MLcase(typ,e,br'))
+ | MLmagic(MLdummy _ as e) when lang () == Haskell -> e
| MLmagic(MLexn _ as e) -> e
| a -> ast_map (simpl o) a
diff --git a/test-suite/bugs/closed/4844.v b/test-suite/bugs/closed/4844.v
new file mode 100644
index 000000000..f140939cc
--- /dev/null
+++ b/test-suite/bugs/closed/4844.v
@@ -0,0 +1,47 @@
+
+(* Bug report 4844 (and 4824):
+ The Haskell extraction was erroneously considering [Any] and
+ [()] as convertible ([Tunknown] an [Tdummy] internally). *)
+
+(* A value with inner logical parts.
+ Its extracted type will be [Sum () ()]. *)
+
+Definition semilogic : True + True := inl I.
+
+(* Higher-order record, whose projection [ST] isn't expressible
+ as an Haskell (or OCaml) type. Hence [ST] is extracted as the
+ unknown type [Any] in Haskell. *)
+
+Record SomeType := { ST : Type }.
+
+Definition SomeTrue := {| ST := True |}.
+
+(* A first version of the issue:
+ [abstrSum] is extracted as [Sum Any Any], so an unsafeCoerce
+ is required to cast [semilogic] into [abstrSum SomeTrue]. *)
+
+Definition abstrSum (t : SomeType) := ((ST t) + (ST t))%type.
+
+Definition semilogic' : abstrSum SomeTrue := semilogic.
+
+(* A deeper version of the issue.
+ In the previous example, the extraction could have reduced
+ [abstrSum SomeTrue] into [True+True], solving the issue.
+ It might do so in future versions. But if we put an inductive
+ in the way, a reduction isn't helpful. *)
+
+Inductive box (t : SomeType) := Box : ST t + ST t -> box t.
+
+Definition boxed_semilogic : box SomeTrue :=
+ Box SomeTrue semilogic.
+
+Require Extraction.
+Extraction Language Haskell.
+Recursive Extraction semilogic' boxed_semilogic.
+(* Warning! To fully check that this bug is still closed,
+ you should run ghc on the extracted code:
+
+Extraction "bug4844.hs" semilogic' boxed_semilogic.
+ghc bug4844.hs
+
+*)