diff options
author | Matthieu Sozeau <matthieu.sozeau@inria.fr> | 2014-09-25 00:12:26 +0200 |
---|---|---|
committer | Matthieu Sozeau <matthieu.sozeau@inria.fr> | 2014-09-27 21:56:58 +0200 |
commit | 3fe4912b568916676644baeb982a3e10c592d887 (patch) | |
tree | 291c25d55d62c94af8fc3eb5a6d6df1150bc893f /test-suite | |
parent | a95210435f336d89f44052170a7c65563e6e35f2 (diff) |
Keyed unification option, compiling the whole standard library
(but deactivated still).
Set Keyed Unification to activate the option, which changes
subterm selection to _always_ use full conversion _after_ finding a
subterm whose head/key matches the key of the term we're looking for.
This applies to rewrite and higher-order unification in
apply/elim/destruct.
Most proof scripts already abide by these semantics. For those that
don't, it's usually only a matter of using:
Declare Equivalent Keys f g.
This make keyed unification consider f and g to match as keys.
This takes care of most cases of abbreviations: typically Def foo :=
bar and rewriting with a bar-headed lhs in a goal mentioning foo works
once they're set equivalent.
For canonical structures, these hints should be automatically declared.
For non-global-reference headed terms, the key is the constructor name
(Sort, Prod...). Evars and metas are no keys.
INCOMPATIBILITIES:
In FMapFullAVL, a Function definition doesn't go through with keyed
unification on.
Diffstat (limited to 'test-suite')
-rw-r--r-- | test-suite/bugs/opened/1596.v | 10 | ||||
-rw-r--r-- | test-suite/success/keyedrewrite.v | 23 |
2 files changed, 26 insertions, 7 deletions
diff --git a/test-suite/bugs/opened/1596.v b/test-suite/bugs/opened/1596.v index 91cd09910..7c5dc4167 100644 --- a/test-suite/bugs/opened/1596.v +++ b/test-suite/bugs/opened/1596.v @@ -1,10 +1,11 @@ - Require Import Relations. Require Import FSets. Require Import Arith. Require Import Omega. Unset Standard Proposition Elimination Names. +Set Keyed Unification. + Lemma Bool_elim_bool : forall (b:bool),b=true \/ b=false. destruct b;try tauto. Qed. @@ -255,9 +256,6 @@ n). induction m;simpl;intro. elim (Bool_elim_bool (H.mem (MessageSpi.MNam n,n0) h));intros. apply SynInc;apply H.mem_2;trivial. - - Fail rewrite H in H0. (* !! impossible here !! *) -Abort. -(* discriminate H0. - Qed.*) + rewrite H in H0. discriminate. (* !! impossible here !! *) + Qed. End B.
\ No newline at end of file diff --git a/test-suite/success/keyedrewrite.v b/test-suite/success/keyedrewrite.v index 438613b44..c99b16e2b 100644 --- a/test-suite/success/keyedrewrite.v +++ b/test-suite/success/keyedrewrite.v @@ -1 +1,22 @@ -Lemma foo :
\ No newline at end of file +Section foo. +Variable f : nat -> nat. + +Definition g := f. + +Variable lem : g 0 = 0. + +Goal f 0 = 0. +Proof. + Fail rewrite lem. +Abort. + +Declare Equivalent Keys @g @f. +(** Now f and g are considered equivalent heads for subterm selection *) +Goal f 0 = 0. +Proof. + rewrite lem. + reflexivity. +Qed. + +Print Equivalent Keys. +End foo. |