summaryrefslogtreecommitdiff
path: root/theories/Numbers/Cyclic/Abstract
diff options
context:
space:
mode:
Diffstat (limited to 'theories/Numbers/Cyclic/Abstract')
-rw-r--r--theories/Numbers/Cyclic/Abstract/CyclicAxioms.v159
-rw-r--r--theories/Numbers/Cyclic/Abstract/NZCyclic.v173
2 files changed, 211 insertions, 121 deletions
diff --git a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
index 528d78c3..51df2fa3 100644
--- a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
+++ b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
@@ -8,12 +8,12 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(* $Id: CyclicAxioms.v 11012 2008-05-28 16:34:43Z letouzey $ *)
+(* $Id$ *)
(** * Signature and specification of a bounded integer structure *)
-(** This file specifies how to represent [Z/nZ] when [n=2^d],
- [d] being the number of digits of these bounded integers. *)
+(** This file specifies how to represent [Z/nZ] when [n=2^d],
+ [d] being the number of digits of these bounded integers. *)
Set Implicit Arguments.
@@ -22,7 +22,7 @@ Require Import Znumtheory.
Require Import BigNumPrelude.
Require Import DoubleType.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
(** First, a description via an operator record and a spec record. *)
@@ -33,7 +33,7 @@ Section Z_nZ_Op.
Record znz_op := mk_znz_op {
(* Conversion functions with Z *)
- znz_digits : positive;
+ znz_digits : positive;
znz_zdigits: znz;
znz_to_Z : znz -> Z;
znz_of_pos : positive -> N * znz; (* Euclidean division by [2^digits] *)
@@ -78,12 +78,12 @@ Section Z_nZ_Op.
znz_div : znz -> znz -> znz * znz;
znz_mod_gt : znz -> znz -> znz; (* specialized version of [znz_mod] *)
- znz_mod : znz -> znz -> znz;
+ znz_mod : znz -> znz -> znz;
znz_gcd_gt : znz -> znz -> znz; (* specialized version of [znz_gcd] *)
- znz_gcd : znz -> znz -> znz;
+ znz_gcd : znz -> znz -> znz;
(* [znz_add_mul_div p i j] is a combination of the [(digits-p)]
- low bits of [i] above the [p] high bits of [j]:
+ low bits of [i] above the [p] high bits of [j]:
[znz_add_mul_div p i j = i*2^p+j/2^(digits-p)] *)
znz_add_mul_div : znz -> znz -> znz -> znz;
(* [znz_pos_mod p i] is [i mod 2^p] *)
@@ -135,7 +135,7 @@ Section Z_nZ_Spec.
Let w_mul_c := w_op.(znz_mul_c).
Let w_mul := w_op.(znz_mul).
Let w_square_c := w_op.(znz_square_c).
-
+
Let w_div21 := w_op.(znz_div21).
Let w_div_gt := w_op.(znz_div_gt).
Let w_div := w_op.(znz_div).
@@ -229,25 +229,25 @@ Section Z_nZ_Spec.
spec_div : forall a b, 0 < [|b|] ->
let (q,r) := w_div a b in
[|a|] = [|q|] * [|b|] + [|r|] /\
- 0 <= [|r|] < [|b|];
-
+ 0 <= [|r|] < [|b|];
+
spec_mod_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] ->
[|w_mod_gt a b|] = [|a|] mod [|b|];
spec_mod : forall a b, 0 < [|b|] ->
[|w_mod a b|] = [|a|] mod [|b|];
-
+
spec_gcd_gt : forall a b, [|a|] > [|b|] ->
Zis_gcd [|a|] [|b|] [|w_gcd_gt a b|];
spec_gcd : forall a b, Zis_gcd [|a|] [|b|] [|w_gcd a b|];
-
+
(* shift operations *)
spec_head00: forall x, [|x|] = 0 -> [|w_head0 x|] = Zpos w_digits;
spec_head0 : forall x, 0 < [|x|] ->
- wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB;
+ wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB;
spec_tail00: forall x, [|x|] = 0 -> [|w_tail0 x|] = Zpos w_digits;
- spec_tail0 : forall x, 0 < [|x|] ->
- exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|w_tail0 x|]) ;
+ spec_tail0 : forall x, 0 < [|x|] ->
+ exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|w_tail0 x|]) ;
spec_add_mul_div : forall x y p,
[|p|] <= Zpos w_digits ->
[| w_add_mul_div p x y |] =
@@ -272,23 +272,23 @@ End Z_nZ_Spec.
(** Generic construction of double words *)
Section WW.
-
+
Variable w : Type.
Variable w_op : znz_op w.
Variable op_spec : znz_spec w_op.
-
+
Let wB := base w_op.(znz_digits).
Let w_to_Z := w_op.(znz_to_Z).
Let w_eq0 := w_op.(znz_eq0).
Let w_0 := w_op.(znz_0).
- Definition znz_W0 h :=
+ Definition znz_W0 h :=
if w_eq0 h then W0 else WW h w_0.
- Definition znz_0W l :=
+ Definition znz_0W l :=
if w_eq0 l then W0 else WW w_0 l.
- Definition znz_WW h l :=
+ Definition znz_WW h l :=
if w_eq0 h then znz_0W l else WW h l.
Lemma spec_W0 : forall h,
@@ -300,7 +300,7 @@ Section WW.
unfold w_0; rewrite op_spec.(spec_0); auto with zarith.
Qed.
- Lemma spec_0W : forall l,
+ Lemma spec_0W : forall l,
zn2z_to_Z wB w_to_Z (znz_0W l) = w_to_Z l.
Proof.
unfold zn2z_to_Z, znz_0W, w_to_Z; simpl; intros.
@@ -309,7 +309,7 @@ Section WW.
unfold w_0; rewrite op_spec.(spec_0); auto with zarith.
Qed.
- Lemma spec_WW : forall h l,
+ Lemma spec_WW : forall h l,
zn2z_to_Z wB w_to_Z (znz_WW h l) = (w_to_Z h)*wB + w_to_Z l.
Proof.
unfold znz_WW, w_to_Z; simpl; intros.
@@ -324,7 +324,7 @@ End WW.
(** Injecting [Z] numbers into a cyclic structure *)
Section znz_of_pos.
-
+
Variable w : Type.
Variable w_op : znz_op w.
Variable op_spec : znz_spec w_op.
@@ -349,7 +349,7 @@ Section znz_of_pos.
apply Zle_trans with X; auto with zarith
end.
match goal with |- ?X <= _ =>
- pattern X at 1; rewrite <- (Zmult_1_l);
+ pattern X at 1; rewrite <- (Zmult_1_l);
apply Zmult_le_compat_r; auto with zarith
end.
case p1; simpl; intros; red; simpl; intros; discriminate.
@@ -373,3 +373,112 @@ Module Type CyclicType.
Parameter w_op : znz_op w.
Parameter w_spec : znz_spec w_op.
End CyclicType.
+
+
+(** A Cyclic structure can be seen as a ring *)
+
+Module CyclicRing (Import Cyclic : CyclicType).
+
+Definition t := w.
+
+Local Notation "[| x |]" := (w_op.(znz_to_Z) x) (at level 0, x at level 99).
+
+Definition eq (n m : t) := [| n |] = [| m |].
+Definition zero : t := w_op.(znz_0).
+Definition one := w_op.(znz_1).
+Definition add := w_op.(znz_add).
+Definition sub := w_op.(znz_sub).
+Definition mul := w_op.(znz_mul).
+Definition opp := w_op.(znz_opp).
+
+Local Infix "==" := eq (at level 70).
+Local Notation "0" := zero.
+Local Notation "1" := one.
+Local Infix "+" := add.
+Local Infix "-" := sub.
+Local Infix "*" := mul.
+Local Notation "!!" := (base (znz_digits w_op)).
+
+Hint Rewrite
+ w_spec.(spec_0) w_spec.(spec_1)
+ w_spec.(spec_add) w_spec.(spec_mul) w_spec.(spec_opp) w_spec.(spec_sub)
+ : cyclic.
+
+Ltac zify :=
+ unfold eq, zero, one, add, sub, mul, opp in *; autorewrite with cyclic.
+
+Lemma add_0_l : forall x, 0 + x == x.
+Proof.
+intros. zify. rewrite Zplus_0_l.
+apply Zmod_small. apply w_spec.(spec_to_Z).
+Qed.
+
+Lemma add_comm : forall x y, x + y == y + x.
+Proof.
+intros. zify. now rewrite Zplus_comm.
+Qed.
+
+Lemma add_assoc : forall x y z, x + (y + z) == x + y + z.
+Proof.
+intros. zify. now rewrite Zplus_mod_idemp_r, Zplus_mod_idemp_l, Zplus_assoc.
+Qed.
+
+Lemma mul_1_l : forall x, 1 * x == x.
+Proof.
+intros. zify. rewrite Zmult_1_l.
+apply Zmod_small. apply w_spec.(spec_to_Z).
+Qed.
+
+Lemma mul_comm : forall x y, x * y == y * x.
+Proof.
+intros. zify. now rewrite Zmult_comm.
+Qed.
+
+Lemma mul_assoc : forall x y z, x * (y * z) == x * y * z.
+Proof.
+intros. zify. now rewrite Zmult_mod_idemp_r, Zmult_mod_idemp_l, Zmult_assoc.
+Qed.
+
+Lemma mul_add_distr_r : forall x y z, (x+y)*z == x*z + y*z.
+Proof.
+intros. zify. now rewrite <- Zplus_mod, Zmult_mod_idemp_l, Zmult_plus_distr_l.
+Qed.
+
+Lemma add_opp_r : forall x y, x + opp y == x-y.
+Proof.
+intros. zify. rewrite <- Zminus_mod_idemp_r. unfold Zminus.
+destruct (Z_eq_dec ([|y|] mod !!) 0) as [EQ|NEQ].
+rewrite Z_mod_zero_opp_full, EQ, 2 Zplus_0_r; auto.
+rewrite Z_mod_nz_opp_full by auto.
+rewrite <- Zplus_mod_idemp_r, <- Zminus_mod_idemp_l.
+rewrite Z_mod_same_full. simpl. now rewrite Zplus_mod_idemp_r.
+Qed.
+
+Lemma add_opp_diag_r : forall x, x + opp x == 0.
+Proof.
+intros. red. rewrite add_opp_r. zify. now rewrite Zminus_diag, Zmod_0_l.
+Qed.
+
+Lemma CyclicRing : ring_theory 0 1 add mul sub opp eq.
+Proof.
+constructor.
+exact add_0_l. exact add_comm. exact add_assoc.
+exact mul_1_l. exact mul_comm. exact mul_assoc.
+exact mul_add_distr_r.
+symmetry. apply add_opp_r.
+exact add_opp_diag_r.
+Qed.
+
+Definition eqb x y :=
+ match w_op.(znz_compare) x y with Eq => true | _ => false end.
+
+Lemma eqb_eq : forall x y, eqb x y = true <-> x == y.
+Proof.
+ intros. unfold eqb, eq. generalize (w_spec.(spec_compare) x y).
+ destruct (w_op.(znz_compare) x y); intuition; try discriminate.
+Qed.
+
+Lemma eqb_correct : forall x y, eqb x y = true -> x==y.
+Proof. now apply eqb_eq. Qed.
+
+End CyclicRing. \ No newline at end of file
diff --git a/theories/Numbers/Cyclic/Abstract/NZCyclic.v b/theories/Numbers/Cyclic/Abstract/NZCyclic.v
index fb3f0cef..517e48ad 100644
--- a/theories/Numbers/Cyclic/Abstract/NZCyclic.v
+++ b/theories/Numbers/Cyclic/Abstract/NZCyclic.v
@@ -8,7 +8,7 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NZCyclic.v 11238 2008-07-19 09:34:03Z herbelin $ i*)
+(*i $Id$ i*)
Require Export NZAxioms.
Require Import BigNumPrelude.
@@ -17,89 +17,79 @@ Require Import CyclicAxioms.
(** * From [CyclicType] to [NZAxiomsSig] *)
-(** A [Z/nZ] representation given by a module type [CyclicType]
- implements [NZAxiomsSig], e.g. the common properties between
- N and Z with no ordering. Notice that the [n] in [Z/nZ] is
+(** A [Z/nZ] representation given by a module type [CyclicType]
+ implements [NZAxiomsSig], e.g. the common properties between
+ N and Z with no ordering. Notice that the [n] in [Z/nZ] is
a power of 2.
*)
Module NZCyclicAxiomsMod (Import Cyclic : CyclicType) <: NZAxiomsSig.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
-Definition NZ := w.
+Definition t := w.
-Definition NZ_to_Z : NZ -> Z := znz_to_Z w_op.
-Definition Z_to_NZ : Z -> NZ := znz_of_Z w_op.
-Notation Local wB := (base w_op.(znz_digits)).
+Definition NZ_to_Z : t -> Z := znz_to_Z w_op.
+Definition Z_to_NZ : Z -> t := znz_of_Z w_op.
+Local Notation wB := (base w_op.(znz_digits)).
-Notation Local "[| x |]" := (w_op.(znz_to_Z) x) (at level 0, x at level 99).
+Local Notation "[| x |]" := (w_op.(znz_to_Z) x) (at level 0, x at level 99).
-Definition NZeq (n m : NZ) := [| n |] = [| m |].
-Definition NZ0 := w_op.(znz_0).
-Definition NZsucc := w_op.(znz_succ).
-Definition NZpred := w_op.(znz_pred).
-Definition NZadd := w_op.(znz_add).
-Definition NZsub := w_op.(znz_sub).
-Definition NZmul := w_op.(znz_mul).
+Definition eq (n m : t) := [| n |] = [| m |].
+Definition zero := w_op.(znz_0).
+Definition succ := w_op.(znz_succ).
+Definition pred := w_op.(znz_pred).
+Definition add := w_op.(znz_add).
+Definition sub := w_op.(znz_sub).
+Definition mul := w_op.(znz_mul).
-Theorem NZeq_equiv : equiv NZ NZeq.
-Proof.
-unfold equiv, reflexive, symmetric, transitive, NZeq; repeat split; intros; auto.
-now transitivity [| y |].
-Qed.
+Local Infix "==" := eq (at level 70).
+Local Notation "0" := zero.
+Local Notation S := succ.
+Local Notation P := pred.
+Local Infix "+" := add.
+Local Infix "-" := sub.
+Local Infix "*" := mul.
-Add Relation NZ NZeq
- reflexivity proved by (proj1 NZeq_equiv)
- symmetry proved by (proj2 (proj2 NZeq_equiv))
- transitivity proved by (proj1 (proj2 NZeq_equiv))
-as NZeq_rel.
+Hint Rewrite w_spec.(spec_0) w_spec.(spec_succ) w_spec.(spec_pred)
+ w_spec.(spec_add) w_spec.(spec_mul) w_spec.(spec_sub) : w.
+Ltac wsimpl :=
+ unfold eq, zero, succ, pred, add, sub, mul; autorewrite with w.
+Ltac wcongruence := repeat red; intros; wsimpl; congruence.
-Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd.
+Instance eq_equiv : Equivalence eq.
Proof.
-unfold NZeq; intros n m H. do 2 rewrite w_spec.(spec_succ). now rewrite H.
+unfold eq. firstorder.
Qed.
-Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd.
+Instance succ_wd : Proper (eq ==> eq) succ.
Proof.
-unfold NZeq; intros n m H. do 2 rewrite w_spec.(spec_pred). now rewrite H.
+wcongruence.
Qed.
-Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd.
+Instance pred_wd : Proper (eq ==> eq) pred.
Proof.
-unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_add).
-now rewrite H1, H2.
+wcongruence.
Qed.
-Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd.
+Instance add_wd : Proper (eq ==> eq ==> eq) add.
Proof.
-unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_sub).
-now rewrite H1, H2.
+wcongruence.
Qed.
-Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd.
+Instance sub_wd : Proper (eq ==> eq ==> eq) sub.
Proof.
-unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_mul).
-now rewrite H1, H2.
+wcongruence.
Qed.
-Delimit Scope IntScope with Int.
-Bind Scope IntScope with NZ.
-Open Local Scope IntScope.
-Notation "x == y" := (NZeq x y) (at level 70) : IntScope.
-Notation "x ~= y" := (~ NZeq x y) (at level 70) : IntScope.
-Notation "0" := NZ0 : IntScope.
-Notation S x := (NZsucc x).
-Notation P x := (NZpred x).
-(*Notation "1" := (S 0) : IntScope.*)
-Notation "x + y" := (NZadd x y) : IntScope.
-Notation "x - y" := (NZsub x y) : IntScope.
-Notation "x * y" := (NZmul x y) : IntScope.
+Instance mul_wd : Proper (eq ==> eq ==> eq) mul.
+Proof.
+wcongruence.
+Qed.
Theorem gt_wB_1 : 1 < wB.
Proof.
-unfold base.
-apply Zpower_gt_1; unfold Zlt; auto with zarith.
+unfold base. apply Zpower_gt_1; unfold Zlt; auto with zarith.
Qed.
Theorem gt_wB_0 : 0 < wB.
@@ -107,7 +97,7 @@ Proof.
pose proof gt_wB_1; auto with zarith.
Qed.
-Lemma NZsucc_mod_wB : forall n : Z, (n + 1) mod wB = ((n mod wB) + 1) mod wB.
+Lemma succ_mod_wB : forall n : Z, (n + 1) mod wB = ((n mod wB) + 1) mod wB.
Proof.
intro n.
pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zplus_mod.
@@ -115,7 +105,7 @@ reflexivity.
now rewrite Zmod_small; [ | split; [auto with zarith | apply gt_wB_1]].
Qed.
-Lemma NZpred_mod_wB : forall n : Z, (n - 1) mod wB = ((n mod wB) - 1) mod wB.
+Lemma pred_mod_wB : forall n : Z, (n - 1) mod wB = ((n mod wB) - 1) mod wB.
Proof.
intro n.
pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zminus_mod.
@@ -123,34 +113,32 @@ reflexivity.
now rewrite Zmod_small; [ | split; [auto with zarith | apply gt_wB_1]].
Qed.
-Lemma NZ_to_Z_mod : forall n : NZ, [| n |] mod wB = [| n |].
+Lemma NZ_to_Z_mod : forall n, [| n |] mod wB = [| n |].
Proof.
intro n; rewrite Zmod_small. reflexivity. apply w_spec.(spec_to_Z).
Qed.
-Theorem NZpred_succ : forall n : NZ, P (S n) == n.
+Theorem pred_succ : forall n, P (S n) == n.
Proof.
-intro n; unfold NZsucc, NZpred, NZeq. rewrite w_spec.(spec_pred), w_spec.(spec_succ).
-rewrite <- NZpred_mod_wB.
+intro n. wsimpl.
+rewrite <- pred_mod_wB.
replace ([| n |] + 1 - 1)%Z with [| n |] by auto with zarith. apply NZ_to_Z_mod.
Qed.
-Lemma Z_to_NZ_0 : Z_to_NZ 0%Z == 0%Int.
+Lemma Z_to_NZ_0 : Z_to_NZ 0%Z == 0.
Proof.
-unfold NZeq, NZ_to_Z, Z_to_NZ. rewrite znz_of_Z_correct.
-symmetry; apply w_spec.(spec_0).
+unfold NZ_to_Z, Z_to_NZ. wsimpl.
+rewrite znz_of_Z_correct; auto.
exact w_spec. split; [auto with zarith |apply gt_wB_0].
Qed.
Section Induction.
-Variable A : NZ -> Prop.
-Hypothesis A_wd : predicate_wd NZeq A.
+Variable A : t -> Prop.
+Hypothesis A_wd : Proper (eq ==> iff) A.
Hypothesis A0 : A 0.
-Hypothesis AS : forall n : NZ, A n <-> A (S n). (* Below, we use only -> direction *)
-
-Add Morphism A with signature NZeq ==> iff as A_morph.
-Proof. apply A_wd. Qed.
+Hypothesis AS : forall n, A n <-> A (S n).
+ (* Below, we use only -> direction *)
Let B (n : Z) := A (Z_to_NZ n).
@@ -163,8 +151,8 @@ Lemma BS : forall n : Z, 0 <= n -> n < wB - 1 -> B n -> B (n + 1).
Proof.
intros n H1 H2 H3.
unfold B in *. apply -> AS in H3.
-setoid_replace (Z_to_NZ (n + 1)) with (S (Z_to_NZ n)) using relation NZeq. assumption.
-unfold NZeq. rewrite w_spec.(spec_succ).
+setoid_replace (Z_to_NZ (n + 1)) with (S (Z_to_NZ n)). assumption.
+wsimpl.
unfold NZ_to_Z, Z_to_NZ.
do 2 (rewrite znz_of_Z_correct; [ | exact w_spec | auto with zarith]).
symmetry; apply Zmod_small; auto with zarith.
@@ -177,11 +165,11 @@ apply Zbounded_induction with wB.
apply B0. apply BS. assumption. assumption.
Qed.
-Theorem NZinduction : forall n : NZ, A n.
+Theorem bi_induction : forall n, A n.
Proof.
-intro n. setoid_replace n with (Z_to_NZ (NZ_to_Z n)) using relation NZeq.
+intro n. setoid_replace n with (Z_to_NZ (NZ_to_Z n)).
apply B_holds. apply w_spec.(spec_to_Z).
-unfold NZeq, NZ_to_Z, Z_to_NZ; rewrite znz_of_Z_correct.
+unfold eq, NZ_to_Z, Z_to_NZ; rewrite znz_of_Z_correct.
reflexivity.
exact w_spec.
apply w_spec.(spec_to_Z).
@@ -189,47 +177,40 @@ Qed.
End Induction.
-Theorem NZadd_0_l : forall n : NZ, 0 + n == n.
+Theorem add_0_l : forall n, 0 + n == n.
Proof.
-intro n; unfold NZadd, NZ0, NZeq. rewrite w_spec.(spec_add). rewrite w_spec.(spec_0).
+intro n. wsimpl.
rewrite Zplus_0_l. rewrite Zmod_small; [reflexivity | apply w_spec.(spec_to_Z)].
Qed.
-Theorem NZadd_succ_l : forall n m : NZ, (S n) + m == S (n + m).
+Theorem add_succ_l : forall n m, (S n) + m == S (n + m).
Proof.
-intros n m; unfold NZadd, NZsucc, NZeq. rewrite w_spec.(spec_add).
-do 2 rewrite w_spec.(spec_succ). rewrite w_spec.(spec_add).
-rewrite NZsucc_mod_wB. repeat rewrite Zplus_mod_idemp_l; try apply gt_wB_0.
+intros n m. wsimpl.
+rewrite succ_mod_wB. repeat rewrite Zplus_mod_idemp_l; try apply gt_wB_0.
rewrite <- (Zplus_assoc ([| n |] mod wB) 1 [| m |]). rewrite Zplus_mod_idemp_l.
rewrite (Zplus_comm 1 [| m |]); now rewrite Zplus_assoc.
Qed.
-Theorem NZsub_0_r : forall n : NZ, n - 0 == n.
+Theorem sub_0_r : forall n, n - 0 == n.
Proof.
-intro n; unfold NZsub, NZ0, NZeq. rewrite w_spec.(spec_sub).
-rewrite w_spec.(spec_0). rewrite Zminus_0_r. apply NZ_to_Z_mod.
+intro n. wsimpl. rewrite Zminus_0_r. apply NZ_to_Z_mod.
Qed.
-Theorem NZsub_succ_r : forall n m : NZ, n - (S m) == P (n - m).
+Theorem sub_succ_r : forall n m, n - (S m) == P (n - m).
Proof.
-intros n m; unfold NZsub, NZsucc, NZpred, NZeq.
-rewrite w_spec.(spec_pred). do 2 rewrite w_spec.(spec_sub).
-rewrite w_spec.(spec_succ). rewrite Zminus_mod_idemp_r.
-rewrite Zminus_mod_idemp_l.
-now replace ([|n|] - ([|m|] + 1))%Z with ([|n|] - [|m|] - 1)%Z by auto with zarith.
+intros n m. wsimpl. rewrite Zminus_mod_idemp_r, Zminus_mod_idemp_l.
+now replace ([|n|] - ([|m|] + 1))%Z with ([|n|] - [|m|] - 1)%Z
+ by auto with zarith.
Qed.
-Theorem NZmul_0_l : forall n : NZ, 0 * n == 0.
+Theorem mul_0_l : forall n, 0 * n == 0.
Proof.
-intro n; unfold NZmul, NZ0, NZ, NZeq. rewrite w_spec.(spec_mul).
-rewrite w_spec.(spec_0). now rewrite Zmult_0_l.
+intro n. wsimpl. now rewrite Zmult_0_l.
Qed.
-Theorem NZmul_succ_l : forall n m : NZ, (S n) * m == n * m + m.
+Theorem mul_succ_l : forall n m, (S n) * m == n * m + m.
Proof.
-intros n m; unfold NZmul, NZsucc, NZadd, NZeq. rewrite w_spec.(spec_mul).
-rewrite w_spec.(spec_add), w_spec.(spec_mul), w_spec.(spec_succ).
-rewrite Zplus_mod_idemp_l, Zmult_mod_idemp_l.
+intros n m. wsimpl. rewrite Zplus_mod_idemp_l, Zmult_mod_idemp_l.
now rewrite Zmult_plus_distr_l, Zmult_1_l.
Qed.