aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorherbelin2007-05-10 15:31:46 +0000
committerherbelin2007-05-10 15:31:46 +0000
commitde6e0fdd9f3078b9b0ccf87ba8a7dcc9cc47c9aa (patch)
tree33bc723ec517a6ad737b7ca52a7e71f643ed3158
parentafaa5d6ee5c27329423c152362e0969dc2d9afdd (diff)
Prise en compte réversibilité des notations de la forme "Notation Nil := @nil".
Ajout @ref au niveau constr pour allègement syntaxe. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@9819 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r--interp/constrextern.ml7
-rw-r--r--interp/notation.ml4
-rw-r--r--parsing/g_constr.ml43
-rw-r--r--test-suite/output/Notations.out4
-rw-r--r--test-suite/output/Notations.v10
5 files changed, 23 insertions, 5 deletions
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index 771a0b299d..cf4d2db0e4 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -834,10 +834,13 @@ and extern_symbol (tmp_scope,scopes as allscopes) vars t = function
try
(* Adjusts to the number of arguments expected by the notation *)
let (t,args) = match t,n with
- | RApp (_,f,args), Some n when List.length args > n ->
+ | RApp (_,(RRef _ as f),args), Some n when List.length args >= n ->
let args1, args2 = list_chop n args in
(if n = 0 then f else RApp (dummy_loc,f,args1)), args2
- | _ -> t,[] in
+ | RApp (_,(RRef _ as f),args), None -> f, args
+ | RRef _, Some 0 -> RApp (dummy_loc,t,[]), []
+ | _, None -> t,[]
+ | _ -> raise No_match in
(* Try matching ... *)
let subst = match_aconstr t pat in
(* Try availability of interpretation ... *)
diff --git a/interp/notation.ml b/interp/notation.ml
index c31ef54ec8..aaab6a933f 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -187,10 +187,10 @@ let cases_pattern_key = function
| PatCstr (_,ref,_,_) -> RefKey (ConstructRef ref)
| _ -> Oth
-let aconstr_key = function
+let aconstr_key = function (* Rem: AApp(ARef ref,[]) stands for @ref *)
| AApp (ARef ref,args) -> RefKey ref, Some (List.length args)
| AList (_,_,AApp (ARef ref,args),_,_) -> RefKey ref, Some (List.length args)
- | ARef ref -> RefKey ref, Some 0
+ | ARef ref -> RefKey ref, None
| _ -> Oth, None
let pattern_key = function
diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4
index 821426a3cd..11c9ff48a1 100644
--- a/parsing/g_constr.ml4
+++ b/parsing/g_constr.ml4
@@ -131,7 +131,8 @@ GEXTEND Gram
[ [ c = operconstr LEVEL "200" -> c ] ]
;
constr:
- [ [ c = operconstr LEVEL "8" -> c ] ]
+ [ [ c = operconstr LEVEL "8" -> c
+ | "@"; f=global -> CAppExpl(loc,(None,f),[]) ] ]
;
operconstr:
[ "200" RIGHTA
diff --git a/test-suite/output/Notations.out b/test-suite/output/Notations.out
index be4cd4faea..2066a7ef39 100644
--- a/test-suite/output/Notations.out
+++ b/test-suite/output/Notations.out
@@ -46,3 +46,7 @@ fun x : nat => ifn x is succ n then n else 0
: bool
-4
: Z
+Nil
+ : forall A : Type, list A
+NIL:list nat
+ : list nat
diff --git a/test-suite/output/Notations.v b/test-suite/output/Notations.v
index 3cc0a189d0..6e637aca39 100644
--- a/test-suite/output/Notations.v
+++ b/test-suite/output/Notations.v
@@ -119,3 +119,13 @@ Require Import ZArith.
Open Scope Z_scope.
Notation "- 4" := (-2 + -2).
Check -4.
+
+(**********************************************************************)
+(* Check notations for references with activated or deactivated *)
+(* implicit arguments *)
+
+Notation Nil := @nil.
+Check Nil.
+
+Notation NIL := nil.
+Check NIL : list nat.