aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2017-08-21 23:01:04 +0200
committerGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2018-02-20 10:03:06 +0100
commit50970e4043d73d9a4fbd17ffe765745f6d726317 (patch)
tree30af940838c330d2b50a2da6c669667c23dfc7fc
parent15abe33f55b317410223bd48576fa35c81943ff9 (diff)
Using an "as" clause when needed for printing irrefutable patterns.
Example which is now reprinted as parsed: fun '((x,y) as z) => (y,x)=z
-rw-r--r--interp/constrextern.ml6
-rw-r--r--interp/notation_ops.ml6
-rw-r--r--pretyping/glob_ops.ml5
-rw-r--r--pretyping/glob_ops.mli2
-rw-r--r--test-suite/output/Notations3.out8
-rw-r--r--test-suite/output/Notations3.v20
6 files changed, 43 insertions, 4 deletions
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index 013f5713e..67e19d125 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -924,7 +924,8 @@ and sub_extern inctx (_,scopes) = extern inctx (None,scopes)
and factorize_prod scopes vars na bk aty c =
match na, DAst.get c with
| Name id, GCases (LetPatternStyle, None, [(e,(Anonymous,None))],[(_,(_,[p],b))])
- when is_gvar id e && not (occur_glob_constr id b) ->
+ when is_gvar id e ->
+ let p = if occur_glob_constr id b then set_pat_alias id p else p in
let b = extern_typ scopes vars b in
let p = extern_cases_pattern_in_scope scopes vars p in
let binder = CLocalPattern (c.loc,(p,None)) in
@@ -946,7 +947,8 @@ and factorize_prod scopes vars na bk aty c =
and factorize_lambda inctx scopes vars na bk aty c =
match na, DAst.get c with
| Name id, GCases (LetPatternStyle, None, [(e,(Anonymous,None))],[(_,(_,[p],b))])
- when is_gvar id e && not (occur_glob_constr id b) ->
+ when is_gvar id e ->
+ let p = if occur_glob_constr id b then set_pat_alias id p else p in
let b = sub_extern inctx scopes vars b in
let p = extern_cases_pattern_in_scope scopes vars p in
let binder = CLocalPattern (c.loc,(p,None)) in
diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml
index 5fe12e570..c44863791 100644
--- a/interp/notation_ops.ml
+++ b/interp/notation_ops.ml
@@ -1208,11 +1208,13 @@ and match_extended_binders ?loc isprod u alp metas na1 na2 bk t sigma b1 b2 =
match na1, DAst.get b1, na2 with
(* Matching individual binders as part of a recursive pattern *)
| Name p, GCases (LetPatternStyle,None,[(e,_)],[(_,(ids,[cp],b1))]), Name id
- when is_gvar p e && is_bindinglist_meta id metas && not (occur_glob_constr p b1) ->
+ when is_gvar p e && is_bindinglist_meta id metas ->
+ let cp = if occur_glob_constr p b1 then set_pat_alias p cp else cp in
let alp,sigma = bind_bindinglist_env alp sigma id [DAst.make ?loc @@ GLocalPattern ((cp,ids),p,bk,t)] in
match_in u alp metas sigma b1 b2
| Name p, GCases (LetPatternStyle,None,[(e,_)],[(_,(_,[cp],b1))]), Name id
- when is_gvar p e && is_onlybinding_pattern_like_meta id metas && not (occur_glob_constr p b1) ->
+ when is_gvar p e && is_onlybinding_pattern_like_meta id metas ->
+ let cp = if occur_glob_constr p b1 then set_pat_alias p cp else cp in
let alp,sigma = bind_binding_env alp sigma id cp in
match_in u alp metas sigma b1 b2
| _, _, Name id when is_bindinglist_meta id metas && (not isprod || na1 != Anonymous)->
diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml
index b3e6aa059..25817478e 100644
--- a/pretyping/glob_ops.ml
+++ b/pretyping/glob_ops.ml
@@ -24,6 +24,11 @@ let alias_of_pat pat = DAst.with_val (function
| PatCstr(_,_,name) -> name
) pat
+let set_pat_alias id = DAst.map (function
+ | PatVar Anonymous -> PatVar (Name id)
+ | PatCstr (cstr,patl,Anonymous) -> PatCstr (cstr,patl,Name id)
+ | pat -> assert false)
+
let cases_predicate_names tml =
List.flatten (List.map (function
| (tm,(na,None)) -> [na]
diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli
index 7a6d50114..0d9fb1f45 100644
--- a/pretyping/glob_ops.mli
+++ b/pretyping/glob_ops.mli
@@ -15,6 +15,8 @@ val cases_pattern_eq : 'a cases_pattern_g -> 'a cases_pattern_g -> bool
val alias_of_pat : 'a cases_pattern_g -> Name.t
+val set_pat_alias : Id.t -> 'a cases_pattern_g -> 'a cases_pattern_g
+
val cast_type_eq : ('a -> 'a -> bool) ->
'a Misctypes.cast_type -> 'a Misctypes.cast_type -> bool
diff --git a/test-suite/output/Notations3.out b/test-suite/output/Notations3.out
index cb18fa356..0463e5bfb 100644
--- a/test-suite/output/Notations3.out
+++ b/test-suite/output/Notations3.out
@@ -201,3 +201,11 @@ exists2' {{x, y}}, x = 0 & y = 0
exists2 x : nat * nat,
let '{{y, z}} := x in y > z & let '{{y, z}} := x in z > y
: Prop
+fun '({{x, y}} as z) => x + y = 0 /\ z = z
+ : nat * nat -> Prop
+myexists ({{x, y}} as z), x + y = 0 /\ z = z
+ : Prop
+exists '({{x, y}} as z), x + y = 0 /\ z = z
+ : Prop
+∀ '({{x, y}} as z), x + y = 0 /\ z = z
+ : Prop
diff --git a/test-suite/output/Notations3.v b/test-suite/output/Notations3.v
index d768b9ba4..9ec459ed6 100644
--- a/test-suite/output/Notations3.v
+++ b/test-suite/output/Notations3.v
@@ -356,3 +356,23 @@ End D.
pattern where only an ident could be reparsed *)
Check ex2 (fun x => let '(y,z) := x in y>z) (fun x => let '(y,z) := x in z>y).
+
+(* A canonical example of a notation with a non-recursive binder *)
+
+Parameter myex : forall {A}, (A -> Prop) -> Prop.
+Notation "'myexists' x , p" := (myex (fun x => p))
+ (at level 200, x pattern, p at level 200, right associativity).
+
+(* A canonical example of a notation with recursive binders *)
+
+Notation "∀ x .. y , P" := (forall x, .. (forall y, P) ..)
+ (at level 200, x binder, y binder, right associativity) : type_scope.
+
+(* Check that printing 'pat uses an "as" when the variable bound to
+ the pattern is dependent. We check it for the three kinds of
+ notations involving bindings of patterns *)
+
+Check fun '((x,y) as z) => x+y=0/\z=z. (* Primitive fun/forall *)
+Check myexists ((x,y) as z), x+y=0/\z=z. (* Isolated binding pattern *)
+Check exists '((x,y) as z), x+y=0/\z=z. (* Applicative recursive binder *)
+Check ∀ '((x,y) as z), x+y=0/\z=z. (* Other example of recursive binder, now treated as the exists case *)