aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2012-12-18 07:01:07 +0000
committerGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2012-12-18 07:01:07 +0000
commitf61e682857596f0274b956295efd2bfba63bc8c0 (patch)
tree46a3ea60f4da1633cf91ce965ffc3d3df52ee300
parent7f4b6b875daaf1d31ff74a58d8916c2a4889cf3a (diff)
Fixing parsing of specific primitive tokens used as notations for patterns
(e.g. 1 for eq_refl). git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@16094 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r--interp/constrexpr_ops.ml2
-rw-r--r--test-suite/output/Notations.out9
-rw-r--r--test-suite/output/Notations.v14
3 files changed, 24 insertions, 1 deletions
diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml
index 602c2314a..e9dae6421 100644
--- a/interp/constrexpr_ops.ml
+++ b/interp/constrexpr_ops.ml
@@ -345,5 +345,5 @@ let rec raw_cases_pattern_expr_of_glob_constr looked_for = function
RCPatCstr (loc, g,[],[])
| GApp (loc,GRef (_,g),l) ->
looked_for g;
- RCPatCstr (loc, g,[],List.map (raw_cases_pattern_expr_of_glob_constr looked_for) l)
+ RCPatCstr (loc, g,List.map (raw_cases_pattern_expr_of_glob_constr looked_for) l,[])
| _ -> raise Not_found
diff --git a/test-suite/output/Notations.out b/test-suite/output/Notations.out
index 47eafe553..7a6c189c4 100644
--- a/test-suite/output/Notations.out
+++ b/test-suite/output/Notations.out
@@ -132,3 +132,12 @@ Identifier 'foo' now a keyword
: nat
fun _ : nat => 9
: nat -> nat
+Identifier 'ONE' now a keyword
+fun (x : nat) (p : x = x) => match p with
+ | ONE => ONE
+ end = p
+ : forall x : nat, x = x -> Prop
+fun (x : nat) (p : x = x) => match p with
+ | 1 => 1
+ end = p
+ : forall x : nat, x = x -> Prop
diff --git a/test-suite/output/Notations.v b/test-suite/output/Notations.v
index 75d5037a8..bfd5482ab 100644
--- a/test-suite/output/Notations.v
+++ b/test-suite/output/Notations.v
@@ -266,3 +266,17 @@ Notation "'foo' n" := (S n) (at level 50): nat_scope.
Check (foo 9).
Check (fun _ : nat => 9).
+
+(* Checking parsing and printing of numerical and non-numerical notations for eq_refl *)
+
+(* This notation was not correctly printed until Pierre B.'s
+ improvements to the interpretation of patterns *)
+
+Notation "'ONE'" := eq_refl.
+Check fun (x:nat) (p : x=x) => match p with ONE => ONE end = p.
+
+(* This one used to failed at parsing until now *)
+
+Notation "1" := eq_refl.
+Check fun (x:nat) (p : x=x) => match p with 1 => 1 end = p.
+