File: bug_18342.v

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: forky, sid, trixie
  • size: 44,116 kB
  • sloc: ml: 234,160; sh: 4,301; python: 3,270; ansic: 2,644; makefile: 882; lisp: 172; javascript: 63; xml: 24; sed: 2
file content (40 lines) | stat: -rw-r--r-- 1,240 bytes parent folder | download | duplicates (2)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
Declare Custom Entry stlc.

Module Bug18342.
Definition A (T : True) := 0.
Notation "?" := A.

Notation "<{ e }>" := e (e custom stlc at level 99).
Notation "x" := x (in custom stlc at level 0, x constr at level 0).
Notation "'SS' x" := (S x) (in custom stlc at level 89, x custom stlc at level 99).

Check <{SS (? I)}>.
End Bug18342.

Declare Custom Entry qmark.

Module Bug18342VariantWithExplicitCoercions.
Definition A (T : True) := 0.

Notation "?" := A (in custom qmark).
Notation "{{ x }}" := x (x custom qmark).

Notation "<{ e }>" := e (e custom stlc at level 99).
Notation "[[ x ]]" := x (in custom stlc at level 0, x constr).
Notation "'SS' x" := (S x) (in custom stlc at level 89, x custom stlc at level 99).

Check <{SS [[{{?}} I]]}>.
End Bug18342VariantWithExplicitCoercions.

Module Bug18342VariantPattern.
Inductive I := C : nat -> I | D : I -> I.

Notation "?" := C (in custom qmark).
Notation "{{ x }}" := x (x custom qmark).

Notation "<{ e }>" := e (e custom stlc at level 99).
Notation "[[ x ]]" := x (in custom stlc at level 0, x constr).
Notation "'DD' x" := (D x) (in custom stlc at level 89, x custom stlc at level 99).

Check fun x => match x with <{DD [[{{?}} y]]}> => y | _ => 0 end.
End Bug18342VariantPattern.