diff options
author | Jasper Hugunin <jasperh@cs.washington.edu> | 2018-01-10 17:13:44 +0900 |
---|---|---|
committer | Jasper Hugunin <jasperh@cs.washington.edu> | 2018-01-17 14:54:14 +0900 |
commit | 92bc1c642c2bf9001b86c71bb26f3e73011d02a7 (patch) | |
tree | c5404b63b6acdf924603d1a72460c0c62d0b6893 | |
parent | 58d209fe36e37b6c0ee4acd702dac333388b1b88 (diff) |
Add a test that `prod_applist_assum` reduces the right number of let-ins
-rw-r--r-- | engine/termops.mli | 5 | ||||
-rw-r--r-- | test-suite/success/dtauto-let-deps.v | 24 |
2 files changed, 28 insertions, 1 deletions
diff --git a/engine/termops.mli b/engine/termops.mli index 1f4c85054..a3559a693 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -165,7 +165,10 @@ val prod_applist : Evd.evar_map -> constr -> constr list -> constr (** In [prod_applist_assum n c args], [c] is supposed to have the form [∀Γ.c] with [Γ] of length [m] and possibly with let-ins; it returns [c] with the assumptions of [Γ] instantiated by [args] and - the local definitions of [Γ] expanded. *) + the local definitions of [Γ] expanded. + Note that [n] counts both let-ins and prods, while the length of [args] + only counts prods. In other words, varying [n] changes how many + trailing let-ins are expanded. *) val prod_applist_assum : Evd.evar_map -> int -> constr -> constr list -> constr (** Remove recursively the casts around a term i.e. diff --git a/test-suite/success/dtauto-let-deps.v b/test-suite/success/dtauto-let-deps.v new file mode 100644 index 000000000..094b2f8b3 --- /dev/null +++ b/test-suite/success/dtauto-let-deps.v @@ -0,0 +1,24 @@ +(* +This test is sensitive to changes in which let-ins are expanded when checking +for dependencies in constructors. +If the (x := X) is not reduced, Foo1 won't be recognized as a conjunction, +and if the (y := X) is reduced, Foo2 will be recognized as a conjunction. + +This tests the behavior of engine/termops.ml : prod_applist_assum, +which is currently specified to reduce exactly the parameters. + +If dtauto is changed to reduce lets in constructors before checking dependency, +this test will need to be changed. +*) + +Context (P Q : Type). +Inductive Foo1 (X : Type) (x := X) := foo1 : let y := X in P -> Q -> Foo1 x. +Inductive Foo2 (X : Type) (x := X) := foo2 : let y := X in P -> Q -> Foo2 y. + +Goal P -> Q -> Foo1 nat. +solve [dtauto]. +Qed. + +Goal P -> Q -> Foo2 nat. +Fail solve [dtauto]. +Abort. |