aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorKazuhiko Sakaguchi2019-08-09 21:19:16 +0200
committerKazuhiko Sakaguchi2019-08-10 20:33:48 +0200
commit970404f2bfa7a26ad64692844441261cebd80082 (patch)
tree68f76ecb0664017f2988b44ba006b73a8960835c
parentb8477fb38842016c226ba9d7be8f60486411a2ee (diff)
[extraction] Fix #7191: Avoid unsound eta-reduction
`Mlutil.simpl` and `Mlutil.atomic_eta_red` did some unsound eta-reductions as follows: (fun x0 ... xn => MLexn x0 ... xn) ->eta MLexn. `MLexn` raises an exception thus is not a value in OCaml. So the above simplification may change the behavior of extracted programs. This patch restricts `atomic_eta_red` to eta-redexes whose core is both atomic and value. Acknowledgement: This work is financially supported by Peano System Inc. on-behalf-of: @peano-system <info@peano-system.jp>
-rw-r--r--plugins/extraction/mlutil.ml4
-rw-r--r--test-suite/output/bug7191.out9
-rw-r--r--test-suite/output/bug7191.v3
3 files changed, 14 insertions, 2 deletions
diff --git a/plugins/extraction/mlutil.ml b/plugins/extraction/mlutil.ml
index 2d5872718f..c57daf0047 100644
--- a/plugins/extraction/mlutil.ml
+++ b/plugins/extraction/mlutil.ml
@@ -779,7 +779,7 @@ let eta_red e =
else e
| _ -> e
-(* Performs an eta-reduction when the core is atomic,
+(* Performs an eta-reduction when the core is atomic and value,
or otherwise returns None *)
let atomic_eta_red e =
@@ -789,7 +789,7 @@ let atomic_eta_red e =
| MLapp (f,a) when test_eta_args_lift 0 n a ->
(match f with
| MLrel k when k>n -> Some (MLrel (k-n))
- | MLglob _ | MLexn _ | MLdummy _ -> Some f
+ | MLglob _ | MLdummy _ -> Some f
| _ -> None)
| _ -> None
diff --git a/test-suite/output/bug7191.out b/test-suite/output/bug7191.out
new file mode 100644
index 0000000000..005455e30c
--- /dev/null
+++ b/test-suite/output/bug7191.out
@@ -0,0 +1,9 @@
+
+type unit0 =
+| Tt
+
+(** val f : unit0 -> unit0 **)
+
+let f _ =
+ assert false (* absurd case *)
+
diff --git a/test-suite/output/bug7191.v b/test-suite/output/bug7191.v
new file mode 100644
index 0000000000..1aa4625b6c
--- /dev/null
+++ b/test-suite/output/bug7191.v
@@ -0,0 +1,3 @@
+Require Extraction.
+Definition f (x : False) : unit -> unit := match x with end.
+Recursive Extraction f.