aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/success/simpl_tuning.v
diff options
context:
space:
mode:
authorGravatar gareuselesinge <gareuselesinge@85f007b7-540e-0410-9357-904b9bb8a0f7>2011-11-21 17:03:42 +0000
committerGravatar gareuselesinge <gareuselesinge@85f007b7-540e-0410-9357-904b9bb8a0f7>2011-11-21 17:03:42 +0000
commitb1bfd9757d33d36b9fc009a97173ea7db2d5196d (patch)
tree0a6f2627ad5490a35679814b849a34378891f238 /test-suite/success/simpl_tuning.v
parent5e62a6a476c925e58e169e43468ed0cee422bb1a (diff)
Configurable simpl tactic
The problem that this patch tries to solve is that sometimes the unfolding behaviour of simpl is too aggressive, frequently exposing match constructs. Moreover one may want some constants to never be unfolded by simpl (like in the Math-Comp library where the nosimpl hack is pervasive). Finally, one may want some constants to be volatile, always unfolded by simple when applied to enough arguments, even if they do not hide a fixpoint or a match. A new vernacular command to be indroduced in a following patch attaches to constants an extra-logical piece of information. It declares wich arguments must be checked to evaluate to a constructor before performing the unfolding of the constant (in the same spirit simpl considers as such the recursive argument of an inner fixpoint). Examples: Arguments minus !x !y. (* x and y must evaluate to a constructor for simpl to unfold minus *) (S x - y) --- simpl ---> (S x - y) (S x - S y) --- simpl ---> (x - y) Definition fcomp A B C (f : B -> C) (g : A -> B) x := f (g x). Arguments fcomp A B C f g x /. Infix "\o" := fcomp (at level 50, left associativity) : nat_scope. (* fcomp hides no fix, but simpl will unfold if applied to 6 arguments *) (fun x => (f \o g) x) = f \o g --- simpl ---> (fun x => f (g x)) = f \o g Arguments minus x y : simpl never. (* simpl will never unfold minus *) (S x - S y) --- simpl ---> (S x - S y) Definition nid (x : nat) := x. Arguments nid / x. (* nid is volatile, unfolded by simpl when applied to 0 or more arguments *) nid --- simpl ---> (fun x => x) Arguments minus x y : simpl nomatch. (* triggers for minus the "new" simpl heuristic I discussed with Hugo: if reducing the fix would leave a match in "head" position then don't unfold a suggestion for a better name (for "nomatch") would be appreciated *) (0 - y) --- simpl ---> 0 (S x - y) --- simpl ---> (S x - y) (S x - S y) --- simpl ---> (x - y) (minus 0) --- simpl ---> (minus 0) (* This last test does not work as one may expect, i.e. (minus 0) --- simpl ---> (fun _ => 0) The point is that simpl is implemented as "strong whd_simpl" and after unfolding minus you have (fun m => match 0 => 0 | S n' => ...n' - m'... end) that is already in whd and exposes a match, that of course "strong" would reduce away but at that stage we don't know, and reducing by hand under the lambda is against whd reduction. So further discussion on that idea is needed. *) git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@14716 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'test-suite/success/simpl_tuning.v')
-rw-r--r--test-suite/success/simpl_tuning.v149
1 files changed, 149 insertions, 0 deletions
diff --git a/test-suite/success/simpl_tuning.v b/test-suite/success/simpl_tuning.v
new file mode 100644
index 000000000..d4191b939
--- /dev/null
+++ b/test-suite/success/simpl_tuning.v
@@ -0,0 +1,149 @@
+(* as it is dynamically inferred by simpl *)
+Arguments minus !n / m.
+
+Lemma foo x y : S (S x) - S y = 0.
+simpl.
+match goal with |- (match y with O => S x | S _ => _ end = 0) => idtac end.
+Abort.
+
+(* we avoid exposing a match *)
+Arguments minus n m : simpl nomatch.
+
+Lemma foo x : minus 0 x = 0.
+simpl.
+match goal with |- (0 = 0) => idtac end.
+Abort.
+
+Lemma foo x y : S (S x) - S y = 0.
+simpl.
+match goal with |- (S x - y = 0) => idtac end.
+Abort.
+
+Lemma foo x y : S (S x) - (S (match y with O => O | S z => S z end)) = 0.
+simpl.
+match goal with |-(S x - (match y with O => _ | S _ => _ end) = 0) => idtac end.
+Abort.
+
+(* we unfold as soon as we have 1 args, but we avoid exposing a match *)
+Arguments minus n / m : simpl nomatch.
+
+Lemma foo : minus 0 = fun x => 0.
+simpl.
+match goal with |- minus 0 = _ => idtac end.
+Abort.
+(* This does not work as one may expect. The point is that simpl is implemented
+ as "strong (whd_simpl_state)" and after unfolding minus you have
+ (fun m => match 0 => 0 | S n => ...) that is already in whd and exposes
+ a match, that of course "strong" would reduce away but at that stage
+ we don't know, and reducing by hand under the lambda is against whd *)
+
+(* extra tuning for the usual heuristic *)
+Arguments minus !n / m : simpl nomatch.
+
+Lemma foo x y : S (S x) - S y = 0.
+simpl.
+match goal with |- (S x - y = 0) => idtac end.
+Abort.
+
+Lemma foo x y : S (S x) - (S (match y with O => O | S z => S z end)) = 0.
+simpl.
+match goal with |-(S x - (match y with O => _ | S _ => _ end) = 0) => idtac end.
+Abort.
+
+(* full control *)
+Arguments minus !n !m /.
+
+Lemma foo x y : S (S x) - S y = 0.
+simpl.
+match goal with |- (S x - y = 0) => idtac end.
+Abort.
+
+Lemma foo x y : S (S x) - (S (match y with O => O | S z => S z end)) = 0.
+simpl.
+match goal with |-(S x - (match y with O => _ | S _ => _ end) = 0) => idtac end.
+Abort.
+
+(* omitting /, that being immediately after the last ! is irrelevant *)
+Arguments minus !n !m.
+
+Lemma foo x y : S (S x) - S y = 0.
+simpl.
+match goal with |- (S x - y = 0) => idtac end.
+Abort.
+
+Lemma foo x y : S (S x) - (S (match y with O => O | S z => S z end)) = 0.
+simpl.
+match goal with |-(S x - (match y with O => _ | S _ => _ end) = 0) => idtac end.
+Abort.
+
+Definition pf (D1 C1 : Type) (f : D1 -> C1) (D2 C2 : Type) (g : D2 -> C2) :=
+ fun x => (f (fst x), g (snd x)).
+
+Delimit Scope foo_scope with F.
+Notation "@@" := nat (only parsing) : foo_scope.
+Notation "@@" := (fun x => x) (only parsing).
+
+Arguments pf {D1%F C1%type} f [D2 C2] g x : simpl never.
+
+Lemma foo x : @pf @@ nat @@ nat nat @@ x = pf @@ @@ x.
+Abort.
+
+Definition fcomp A B C f (g : A -> B) (x : A) : C := f (g x).
+
+(* fcomp is unfolded if applied to 6 args *)
+Arguments fcomp {A B C}%type f g x /.
+
+Notation "f \o g" := (fcomp f g) (at level 50).
+
+Lemma foo (f g h : nat -> nat) x : pf (f \o g) h x = pf f h (g (fst x), snd x).
+simpl.
+match goal with |- (pf (f \o g) h x = _) => idtac end.
+case x; intros x1 x2.
+simpl.
+match goal with |- (pf (f \o g) h _ = pf f h _) => idtac end.
+unfold pf; simpl.
+match goal with |- (f (g x1), h x2) = (f (g x1), h x2) => idtac end.
+Abort.
+
+Definition volatile := fun x : nat => x.
+Arguments volatile /.
+
+Lemma foo : volatile = volatile.
+simpl.
+match goal with |- (fun _ => _) = _ => idtac end.
+Abort.
+
+Set Implicit Arguments.
+
+Section S1.
+
+Variable T1 : Type.
+
+Section S2.
+
+Variable T2 : Type.
+
+Fixpoint f (x : T1) (y : T2) n (v : unit) m {struct n} : nat :=
+ match n, m with
+ | 0,_ => 0
+ | S _, 0 => n
+ | S n', S m' => f x y n' v m' end.
+
+Global Arguments f x y !n !v !m.
+
+Lemma foo x y n m : f x y (S n) tt m = f x y (S n) tt (S m).
+simpl.
+match goal with |- (f _ _ _ _ _ = f _ _ _ _ _) => idtac end.
+Abort.
+
+End S2.
+
+Lemma foo T x y n m : @f T x y (S n) tt m = @f T x y (S n) tt (S m).
+simpl.
+match goal with |- (f _ _ _ _ _ = f _ _ _ _ _) => idtac end.
+Abort.
+
+End S1.
+
+Arguments f : clear implicits and scopes.
+