diff options
| author | Hugo Herbelin | 2017-05-20 21:27:34 +0200 |
|---|---|---|
| committer | Hugo Herbelin | 2017-05-31 00:35:46 +0200 |
| commit | 9ee5808746cbcf6e04c08e6a2e798b6cbb34bb06 (patch) | |
| tree | 4304ee584c02cbee626dc1ace75e384faf990e4a | |
| parent | bcc9165aec1a80d563d7060ef127ad022e9ed008 (diff) | |
Fixing a too lax constraint for finding recursive binder part of a notation.
This was preventing to work examples such as:
Notation "[ x ; .. ; y ; z ]" := ((x,((fun u => u), .. (y,(fun u =>u,z)) ..))).
| -rw-r--r-- | interp/notation_ops.ml | 2 | ||||
| -rw-r--r-- | test-suite/output/Notations3.out | 5 | ||||
| -rw-r--r-- | test-suite/output/Notations3.v | 6 |
3 files changed, 12 insertions, 1 deletions
diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 6f91009111..95873ea80d 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -287,7 +287,7 @@ let compare_recursive_parts found f f' (iterator,subc) = | Some _ -> false end | GLambda (Name x,_,t_x,c), GLambda (Name y,_,t_y,term) - | GProd (Name x,_,t_x,c), GProd (Name y,_,t_y,term) -> + | GProd (Name x,_,t_x,c), GProd (Name y,_,t_y,term) when not (Id.equal x y) -> (* We found a binding position where it differs *) begin match !diff with | None -> diff --git a/test-suite/output/Notations3.out b/test-suite/output/Notations3.out index 4d59a92cbf..f4ecfd7362 100644 --- a/test-suite/output/Notations3.out +++ b/test-suite/output/Notations3.out @@ -98,5 +98,10 @@ fun n : nat => foo4 n (fun _ y : nat => ETA z : nat, (fun _ : nat => y = 0)) : nat -> Prop tele (t : Type) '(y, z) (x : t0) := tt : forall t : Type, nat * nat -> t -> fpack +[fun x : nat => x + 0;; fun x : nat => x + 1;; fun x : nat => x + 2] + : (nat -> nat) * + ((nat -> nat) * + ((nat -> nat) * + ((nat -> nat) * ((nat -> nat) * ((nat -> nat) * (nat -> nat)))))) foo5 x nat x : nat -> nat diff --git a/test-suite/output/Notations3.v b/test-suite/output/Notations3.v index 96d831944f..71536c68fb 100644 --- a/test-suite/output/Notations3.v +++ b/test-suite/output/Notations3.v @@ -140,6 +140,12 @@ Notation "'tele' x .. z := b" := Check tele (t:Type) '((y,z):nat*nat) (x:t) := tt. +(* Checking that "fun" in a notation does not mixed up with the + detection of a recursive binder *) + +Notation "[ x ;; .. ;; y ]" := ((x,((fun u => S u), .. (y,(fun u => S u,fun v:nat => v)) ..))). +Check [ fun x => x+0 ;; fun x => x+1 ;; fun x => x+2 ]. + (* Cyprien's part of bug #4765 *) Notation foo5 x T y := (fun x : T => y). |
