Skip to content

Commit eab6181

Browse files
authored
Merge pull request #746 from patrick-nicodemus/rewrite
Created rewrite tactic
2 parents 153ebf9 + c3df3b1 commit eab6181

File tree

3 files changed

+91
-0
lines changed

3 files changed

+91
-0
lines changed

apps/eltac/tests/test_rewrite.v

+44
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,44 @@
1+
From elpi.apps Require Import eltac.rewrite.
2+
3+
Goal (forall x : nat, 1 + x = x + 1) ->
4+
forall y, 2 * ((y+y) + 1) = ((y + y)+1) * 2.
5+
Proof.
6+
intro H.
7+
intro x.
8+
eltac.rewrite H.
9+
eltac.rewrite PeanoNat.Nat.mul_comm.
10+
exact eq_refl.
11+
Defined.
12+
13+
Section Example_rewrite.
14+
Variable A : Type.
15+
Variable B : A -> Type.
16+
Variable C : forall (a : A) (b : B a), Type.
17+
Variable add : forall {a : A} {b : B a}, C a b -> C a b -> C a b.
18+
Variable sym : forall {a : A} {b : B a} (c c' : C a b), add c c' = add c' c.
19+
20+
Goal forall (a : A) (b : B a) (x y : C a b),
21+
add x y = add y x /\ add x y = add y x.
22+
Proof.
23+
intros a b x y.
24+
eltac.rewrite @sym. (* @sym is a gref *)
25+
(** [add y x = add y x /\ add y x = add y x] *)
26+
easy.
27+
Defined.
28+
29+
Goal forall (a : A) (b : B a) (x y : C a b),
30+
add x y = add y x /\ add x y = add y x.
31+
Proof.
32+
intros a b x y.
33+
eltac.rewrite sym. (* because of implicit arguments, this is sym _ _, which is a term *)
34+
easy.
35+
Defined.
36+
37+
Goal forall n, 2 * n = n * 2.
38+
Proof.
39+
intro n.
40+
Fail eltac.rewrite PeanoNat.Nat.add_comm.
41+
eltac.rewrite PeanoNat.Nat.add_comm "strong".
42+
Abort.
43+
44+
End Example_rewrite.

apps/eltac/theories/rewrite.v

+46
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,46 @@
1+
From elpi Require Export elpi.
2+
3+
Elpi Tactic rewrite.
4+
Elpi Accumulate lp:{{
5+
% Second argument is a type of the form forall x1 x2 x3... P = Q.
6+
% First argument is a term of that type.
7+
% This tactic finds a subterm of the goal that Q unifies with
8+
% and rewrites all instances of that subterm from right to left.
9+
pred rewrite i:list argument, i:term, i:term, o:goal, o:list sealed-goal.
10+
11+
% The copy predicate used below is discussed in the tutorial here:
12+
% https://lpcic.github.io/coq-elpi/tutorial_coq_elpi_tactic.html#let-s-code-set-in-elpi
13+
14+
rewrite Strong Eqpf {{@eq lp:S lp:P lp:Q }} (goal _ _ GoalType _ _ as G) GL :-
15+
% First, introduce a rule that causes "copy" to act as a function
16+
% sending a type T to the same type, but with all
17+
% subterms of T unifiable with Q to be replaced with a fresh constant x.
18+
pi x\
19+
(pi J\ copy J x :- Strong = [str "strong"| _], coq.unify-leq Q J ok) =>
20+
(pi J\ copy J x :- [] = Strong, Q = J) =>
21+
% Apply this copy function to the goal type.
22+
(copy GoalType (A x),
23+
% If the subterm Q did indeed appear in the goal,
24+
% then pattern match on the given equality assumption P = Q,
25+
% causing Q to be replaced with P everywhere.
26+
if (occurs x (A x))
27+
(refine (match
28+
Eqpf {{ fun a (e : @eq lp:S lp:P a) => lp:(A a) }}
29+
% We only want to create one hole,
30+
% the one corresponding to the
31+
% contents of the (single) branch of the match.
32+
[Hole_])
33+
G GL
34+
)
35+
(coq.ltac.fail _ "Couldn't unify")).
36+
37+
solve (goal _ _ _ _ [trm Eq | Strong] as G) GL :-
38+
coq.typecheck Eq Ty ok,
39+
coq.saturate Ty Eq Eq',
40+
coq.typecheck Eq' Ty' ok,
41+
rewrite Strong Eq' Ty' G GL.
42+
}}.
43+
44+
Tactic Notation "eltac.rewrite" ident(T) := elpi rewrite ltac_term:(T).
45+
Tactic Notation "eltac.rewrite" uconstr(T) := elpi rewrite ltac_term:(T).
46+
Tactic Notation "eltac.rewrite" uconstr(T) string(s) := elpi rewrite ltac_term:(T) ltac_string:(s).

apps/eltac/theories/tactics.v

+1
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,6 @@
11
From elpi.apps.eltac Require Export
22
intro
3+
rewrite
34
constructor
45
assumption
56
discriminate

0 commit comments

Comments
 (0)