aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--CHANGES1
-rw-r--r--theories/Init/Logic.v7
-rw-r--r--theories/Vectors/VectorDef.v8
3 files changed, 12 insertions, 4 deletions
diff --git a/CHANGES b/CHANGES
index d02e29e4d..530e25096 100644
--- a/CHANGES
+++ b/CHANGES
@@ -86,6 +86,7 @@ Libraries
Vcons become Vector.cons. You can use notations by importing
Vector.VectorNotations.
- Removal of TheoryList. Requiring List instead should work most of the time.
+- New syntax "rew Heq in H" and "rew <- Heq in H" for eq_rect and eq_rect_r.
Internal infrastructure
diff --git a/theories/Init/Logic.v b/theories/Init/Logic.v
index 8551ca97d..82d2e04d7 100644
--- a/theories/Init/Logic.v
+++ b/theories/Init/Logic.v
@@ -331,6 +331,13 @@ Section Logic_lemmas.
Defined.
End Logic_lemmas.
+Notation "'rew' H 'in' H'" := (eq_rect _ _ H' _ H)
+ (at level 0, H' at level 9).
+Notation "'rew' <- H 'in' H'" := (eq_rect_r _ H' H)
+ (at level 0, H' at level 9).
+Notation "'rew' -> H 'in' H'" := (eq_rect _ _ H' _ H)
+ (at level 0, H' at level 9, only parsing).
+
Theorem f_equal2 :
forall (A1 A2 B:Type) (f:A1 -> A2 -> B) (x1 y1:A1)
(x2 y2:A2), x1 = y1 -> x2 = y2 -> f x1 x2 = f y1 y2.
diff --git a/theories/Vectors/VectorDef.v b/theories/Vectors/VectorDef.v
index 777e68b45..adc2c09a3 100644
--- a/theories/Vectors/VectorDef.v
+++ b/theories/Vectors/VectorDef.v
@@ -197,18 +197,18 @@ Fixpoint rev_append_tail {A n p} (v : t A n) (w: t A p)
| a :: v' => rev_append_tail v' (a :: w)
end.
+Import EqdepFacts.
+
(** This one has a better type *)
Definition rev_append {A n p} (v: t A n) (w: t A p)
:t A (n + p) :=
-eq_rect _ (fun n => t A n) (rev_append_tail v w) _
- (eq_sym _ _ _ (plus_tail_plus n p)).
+ rew <- (plus_tail_plus n p) in (rev_append_tail v w).
(** rev [a₁ ; a₂ ; .. ; an] is [an ; a{n-1} ; .. ; a₁]
Caution : There is a lot of rewrite garbage in this definition *)
Definition rev {A n} (v : t A n) : t A n :=
- eq_rect _ (fun n => t A n) (rev_append v [])
- _ (eq_sym _ _ _ (plus_n_O _)).
+ rew <- (plus_n_O _) in (rev_append v []).
End BASES.
Local Notation "v [@ p ]" := (nth v p) (at level 1).