aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite
diff options
context:
space:
mode:
authorGravatar Arnaud Spiwack <arnaud@spiwack.net>2016-06-25 21:32:19 +0200
committerGravatar Maxime Dénès <mail@maximedenes.fr>2016-06-27 12:36:17 +0200
commit5ee1e3dabdee3073f179a7839b06b0ef9288fd81 (patch)
treee28aed7a6fec21b4e12e538825d0706dd02c886f /test-suite
parentcbbc76d0ac6c9710d341b8bc0362aa9d7599f45d (diff)
Patterns in binders: printing tests
Diffstat (limited to 'test-suite')
-rw-r--r--test-suite/output/PatternsInBinders.out31
-rw-r--r--test-suite/output/PatternsInBinders.v54
2 files changed, 85 insertions, 0 deletions
diff --git a/test-suite/output/PatternsInBinders.out b/test-suite/output/PatternsInBinders.out
new file mode 100644
index 000000000..6a28475d7
--- /dev/null
+++ b/test-suite/output/PatternsInBinders.out
@@ -0,0 +1,31 @@
+swap = fun '(x, y) => (y, x)
+ : A * B -> B * A
+fun '(x, y) => (y, x)
+ : A * B -> B * A
+forall '(x, y), swap (x, y) = (y, x)
+ : Prop
+proj_informative = fun 'exist _ x _ => x : A
+ : {x : A | P x} -> A
+foo = fun 'Bar n b tt p => if b then n + p else n - p
+ : Foo -> nat
+baz =
+fun 'Bar n1 _ tt p1 => fun 'Bar _ _ tt _ => n1 + p1
+ : Foo -> Foo -> nat
+λ '(x, y), (y, x)
+ : A * B → B * A
+∀ '(x, y), swap (x, y) = (y, x)
+ : Prop
+swap =
+fun (A B : Type) (pat : A * B) => let '(x, y) := pat in (y, x)
+ : forall A B : Type, A * B -> B * A
+
+Arguments A, B are implicit and maximally inserted
+Argument scopes are [type_scope type_scope _]
+forall (A B : Type) (pat : A * B), let '(x, y) := pat in swap (x, y) = (y, x)
+ : Prop
+exists pat : A * A, let '(x, y) := pat in swap (x, y) = (y, x)
+ : Prop
+both_z =
+fun pat : nat * nat =>
+let '(n, p) as pat0 := pat return (F pat0) in (Z n, Z p) : F (n, p)
+ : forall pat : nat * nat, F pat
diff --git a/test-suite/output/PatternsInBinders.v b/test-suite/output/PatternsInBinders.v
new file mode 100644
index 000000000..8911909ab
--- /dev/null
+++ b/test-suite/output/PatternsInBinders.v
@@ -0,0 +1,54 @@
+(** The purpose of this file is to test printing of the destructive
+ patterns used in binders ([fun] and [forall]). *)
+
+Parameters (A B : Type) (P:A->Prop).
+
+Definition swap '((x,y) : A*B) := (y,x).
+Print swap.
+
+Check fun '((x,y) : A*B) => (y,x).
+
+Check forall '(x,y), swap (x,y) = (y,x).
+
+Definition proj_informative '(exist _ x _ : { x:A | P x }) : A := x.
+Print proj_informative.
+
+Inductive Foo := Bar : nat -> bool -> unit -> nat -> Foo.
+Definition foo '(Bar n b tt p) :=
+ if b then n+p else n-p.
+Print foo.
+
+Definition baz '(Bar n1 b1 tt p1) '(Bar n2 b2 tt p2) := n1+p1.
+Print baz.
+
+(** Some test involving unicode noations. *)
+Module WithUnicode.
+
+ Require Import Coq.Unicode.Utf8.
+
+ Check λ '((x,y) : A*B), (y,x).
+ Check ∀ '(x,y), swap (x,y) = (y,x).
+
+End WithUnicode.
+
+
+(** * Suboptimal printing *)
+
+(** These tests show examples which expose the [let] introduced by
+ the pattern notation in binders. *)
+
+Module Suboptimal.
+
+Definition swap {A B} '((x,y) : A*B) := (y,x).
+Print swap.
+
+Check forall (A B:Type) '((x,y) : A*B), swap (x,y) = (y,x).
+
+Check exists '((x,y):A*A), swap (x,y) = (y,x).
+
+Inductive Fin (n:nat) := Z : Fin n.
+Definition F '(n,p) : Type := (Fin n * Fin p)%type.
+Definition both_z '(n,p) : F (n,p) := (Z _,Z _).
+Print both_z.
+
+End Suboptimal.