aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2011-08-08 23:04:24 +0000
committerGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2011-08-08 23:04:24 +0000
commit022f78302514e071ca60c2e8719a1e2ec66eaed6 (patch)
tree7942a37a69f7e591324daff14c20dea10f9c6bcb
parentea09770efde4a7bdd573407a408e54ec2b41b6ad (diff)
New proposition "rewrite Heq in H" for eq_rect (assuming that there is
a smaller risk that "rewrite" clashes with a name used for constr). git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@14390 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r--plugins/romega/ReflOmegaCore.v34
-rw-r--r--theories/Init/Logic.v12
-rw-r--r--theories/Logic/EqdepFacts.v8
-rw-r--r--theories/Vectors/VectorDef.v4
4 files changed, 29 insertions, 29 deletions
diff --git a/plugins/romega/ReflOmegaCore.v b/plugins/romega/ReflOmegaCore.v
index c82abfc85..b4e470470 100644
--- a/plugins/romega/ReflOmegaCore.v
+++ b/plugins/romega/ReflOmegaCore.v
@@ -2038,12 +2038,12 @@ Qed.
(* \subsection{La fonction de normalisation des termes (moteur de réécriture)} *)
-Fixpoint rewrite (s : step) : term -> term :=
+Fixpoint t_rewrite (s : step) : term -> term :=
match s with
- | C_DO_BOTH s1 s2 => apply_both (rewrite s1) (rewrite s2)
- | C_LEFT s => apply_left (rewrite s)
- | C_RIGHT s => apply_right (rewrite s)
- | C_SEQ s1 s2 => fun t : term => rewrite s2 (rewrite s1 t)
+ | C_DO_BOTH s1 s2 => apply_both (t_rewrite s1) (t_rewrite s2)
+ | C_LEFT s => apply_left (t_rewrite s)
+ | C_RIGHT s => apply_right (t_rewrite s)
+ | C_SEQ s1 s2 => fun t : term => t_rewrite s2 (t_rewrite s1 t)
| C_NOP => fun t : term => t
| C_OPP_PLUS => Topp_plus
| C_OPP_OPP => Topp_opp
@@ -2069,7 +2069,7 @@ Fixpoint rewrite (s : step) : term -> term :=
| C_MULT_COMM => Tmult_comm
end.
-Theorem rewrite_stable : forall s : step, term_stable (rewrite s).
+Theorem t_rewrite_stable : forall s : step, term_stable (t_rewrite s).
Proof.
simple induction s; simpl in |- *;
[ intros; apply apply_both_stable; auto
@@ -2453,7 +2453,7 @@ Definition state (m : int) (s : step) (prop1 prop2 : proposition) :=
match prop2 with
| EqTerm b2 b3 =>
if beq Null 0
- then EqTerm (Tint 0) (rewrite s (b1 + (- b3 + b2) * Tint m)%term)
+ then EqTerm (Tint 0) (t_rewrite s (b1 + (- b3 + b2) * Tint m)%term)
else TrueTerm
| _ => TrueTerm
end
@@ -2463,7 +2463,7 @@ Definition state (m : int) (s : step) (prop1 prop2 : proposition) :=
Theorem state_valid : forall (m : int) (s : step), valid2 (state m s).
Proof.
unfold valid2 in |- *; intros m s ep e p1 p2; unfold state in |- *; Simplify;
- simpl in |- *; auto; elim (rewrite_stable s e); simpl in |- *;
+ simpl in |- *; auto; elim (t_rewrite_stable s e); simpl in |- *;
intros H1 H2; elim H1.
now rewrite H2, plus_opp_l, plus_0_l, mult_0_l.
Qed.
@@ -2585,19 +2585,19 @@ Qed.
Definition move_right (s : step) (p : proposition) :=
match p with
- | EqTerm t1 t2 => EqTerm (Tint 0) (rewrite s (t1 + - t2)%term)
- | LeqTerm t1 t2 => LeqTerm (Tint 0) (rewrite s (t2 + - t1)%term)
- | GeqTerm t1 t2 => LeqTerm (Tint 0) (rewrite s (t1 + - t2)%term)
- | LtTerm t1 t2 => LeqTerm (Tint 0) (rewrite s (t2 + Tint (-(1)) + - t1)%term)
- | GtTerm t1 t2 => LeqTerm (Tint 0) (rewrite s (t1 + Tint (-(1)) + - t2)%term)
- | NeqTerm t1 t2 => NeqTerm (Tint 0) (rewrite s (t1 + - t2)%term)
+ | EqTerm t1 t2 => EqTerm (Tint 0) (t_rewrite s (t1 + - t2)%term)
+ | LeqTerm t1 t2 => LeqTerm (Tint 0) (t_rewrite s (t2 + - t1)%term)
+ | GeqTerm t1 t2 => LeqTerm (Tint 0) (t_rewrite s (t1 + - t2)%term)
+ | LtTerm t1 t2 => LeqTerm (Tint 0) (t_rewrite s (t2 + Tint (-(1)) + - t1)%term)
+ | GtTerm t1 t2 => LeqTerm (Tint 0) (t_rewrite s (t1 + Tint (-(1)) + - t2)%term)
+ | NeqTerm t1 t2 => NeqTerm (Tint 0) (t_rewrite s (t1 + - t2)%term)
| p => p
end.
Theorem move_right_valid : forall s : step, valid1 (move_right s).
Proof.
unfold valid1, move_right in |- *; intros s ep e p; Simplify; simpl in |- *;
- elim (rewrite_stable s e); simpl in |- *;
+ elim (t_rewrite_stable s e); simpl in |- *;
[ symmetry in |- *; apply egal_left; assumption
| intro; apply le_left; assumption
| intro; apply le_left; rewrite <- ge_le_iff; assumption
@@ -2950,14 +2950,14 @@ Qed.
Theorem move_right_stable : forall s : step, prop_stable (move_right s).
Proof.
unfold move_right, prop_stable in |- *; intros s ep e p; split;
- [ Simplify; simpl in |- *; elim (rewrite_stable s e); simpl in |- *;
+ [ Simplify; simpl in |- *; elim (t_rewrite_stable s e); simpl in |- *;
[ symmetry in |- *; apply egal_left; assumption
| intro; apply le_left; assumption
| intro; apply le_left; rewrite <- ge_le_iff; assumption
| intro; apply lt_left; rewrite <- gt_lt_iff; assumption
| intro; apply lt_left; assumption
| intro; apply ne_left_2; assumption ]
- | case p; simpl in |- *; intros; auto; generalize H; elim (rewrite_stable s);
+ | case p; simpl in |- *; intros; auto; generalize H; elim (t_rewrite_stable s);
simpl in |- *; intro H1;
[ rewrite (plus_0_r_reverse (interp_term e t0)); rewrite H1;
rewrite plus_permute; rewrite plus_opp_r;
diff --git a/theories/Init/Logic.v b/theories/Init/Logic.v
index 4c4bf6253..00a93efa0 100644
--- a/theories/Init/Logic.v
+++ b/theories/Init/Logic.v
@@ -334,12 +334,12 @@ 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).
+Notation "'rewrite' H 'in' H'" := (eq_rect _ _ H' _ H)
+ (at level 10, H' at level 9).
+Notation "'rewrite' <- H 'in' H'" := (eq_rect_r _ H' H)
+ (at level 10, H' at level 9).
+Notation "'rewrite' -> H 'in' H'" := (eq_rect _ _ H' _ H)
+ (at level 10, H' at level 9, only parsing).
Theorem f_equal2 :
forall (A1 A2 B:Type) (f:A1 -> A2 -> B) (x1 y1:A1)
diff --git a/theories/Logic/EqdepFacts.v b/theories/Logic/EqdepFacts.v
index 2646bb5ae..8fbd79fa4 100644
--- a/theories/Logic/EqdepFacts.v
+++ b/theories/Logic/EqdepFacts.v
@@ -84,7 +84,7 @@ Section Dependent_Equality.
equalities *)
Inductive eq_dep1 (p:U) (x:P p) (q:U) (y:P q) : Prop :=
- eq_dep1_intro : forall h:q = p, x = rew h in y -> eq_dep1 p x q y.
+ eq_dep1_intro : forall h:q = p, x = rewrite h in y -> eq_dep1 p x q y.
Lemma eq_dep1_dep :
forall (p:U) (x:P p) (q:U) (y:P q), eq_dep1 p x q y -> eq_dep p x q y.
@@ -164,7 +164,7 @@ Qed.
Set Implicit Arguments.
-Lemma eq_sigT_sig_eq : forall X P (x1 x2:X) H1 H2, existT P x1 H1 = existT P x2 H2 <-> {H:x1=x2 | rew H in H1 = H2}.
+Lemma eq_sigT_sig_eq : forall X P (x1 x2:X) H1 H2, existT P x1 H1 = existT P x2 H2 <-> {H:x1=x2 | rewrite H in H1 = H2}.
Proof.
intros; split; intro H.
- change x2 with (projT1 (existT P x2 H2)).
@@ -186,7 +186,7 @@ Proof.
Defined.
Lemma eq_sigT_snd :
- forall X P (x1 x2:X) H1 H2 (H:existT P x1 H1 = existT P x2 H2), rew (eq_sigT_fst H) in H1 = H2.
+ forall X P (x1 x2:X) H1 H2 (H:existT P x1 H1 = existT P x2 H2), rewrite (eq_sigT_fst H) in H1 = H2.
Proof.
intros.
unfold eq_sigT_fst.
@@ -206,7 +206,7 @@ Proof.
Defined.
Lemma eq_sig_snd :
- forall X P (x1 x2:X) H1 H2 (H:exist P x1 H1 = exist P x2 H2), rew (eq_sig_fst H) in H1 = H2.
+ forall X P (x1 x2:X) H1 H2 (H:exist P x1 H1 = exist P x2 H2), rewrite (eq_sig_fst H) in H1 = H2.
Proof.
intros.
unfold eq_sig_fst, eq_ind.
diff --git a/theories/Vectors/VectorDef.v b/theories/Vectors/VectorDef.v
index adc2c09a3..9129b94de 100644
--- a/theories/Vectors/VectorDef.v
+++ b/theories/Vectors/VectorDef.v
@@ -202,13 +202,13 @@ 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) :=
- rew <- (plus_tail_plus n p) in (rev_append_tail v w).
+ rewrite <- (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 :=
- rew <- (plus_n_O _) in (rev_append v []).
+ rewrite <- (plus_n_O _) in (rev_append v []).
End BASES.
Local Notation "v [@ p ]" := (nth v p) (at level 1).