aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorGaëtan Gilbert2020-07-02 15:47:15 +0200
committerGaëtan Gilbert2020-07-02 15:49:08 +0200
commit7bc3fb4def33b846b7ac01efe7e42ae6617d46d7 (patch)
tree82155eed74bf2a8cfbb3ae032e956891629c1c85
parenta1835c755f07b2199d122632360a5d236cd9984e (diff)
Fix record pattern interpretation with implicit arguments
We interpret `{|proj=pat|}` as `@C _ pat` instead of `C _ pat` (where the `_` stands for the parameters). Fix #12534
-rw-r--r--interp/constrintern.ml6
-rw-r--r--test-suite/bugs/closed/bug_12534.v9
2 files changed, 12 insertions, 3 deletions
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index ee041ed088..1e64d78f94 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -1659,12 +1659,12 @@ let drop_notations_pattern looked_for genv =
| None -> DAst.make ?loc @@ RCPatAtom None
| Some (n, head, pl) ->
let pl =
- if get_asymmetric_patterns () then pl else
let pars = List.make n (CAst.make ?loc @@ CPatAtom None) in
- List.rev_append pars pl in
+ List.rev_append pars pl
+ in
let (_,argscs) = find_remaining_scopes [] pl head in
let pats = List.map2 (in_pat_sc scopes) argscs pl in
- DAst.make ?loc @@ RCPatCstr(head, [], pats)
+ DAst.make ?loc @@ RCPatCstr(head, pats, [])
end
| CPatCstr (head, None, pl) ->
begin
diff --git a/test-suite/bugs/closed/bug_12534.v b/test-suite/bugs/closed/bug_12534.v
new file mode 100644
index 0000000000..a55515feb6
--- /dev/null
+++ b/test-suite/bugs/closed/bug_12534.v
@@ -0,0 +1,9 @@
+Record C {PROP: Prop} (P : PROP) : Type := { c : unit}.
+Check fun '{|c:=x|} => tt. (* Fine *)
+Arguments Build_C {_ _} _.
+Check fun '(Build_C _) => tt. (* Works. Note: just 1 argument! *)
+Check fun '{|c:=x|} => tt.
+(* Error: The constructor @Build_C (in type @C) expects 1 argument. *)
+
+Set Asymmetric Patterns.
+Check fun '{|c:=x|} => tt.