aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories/Numbers/Integer
diff options
context:
space:
mode:
authorGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2010-01-08 17:36:28 +0000
committerGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2010-01-08 17:36:28 +0000
commit6477ab0f7ea03a0563ca7ba2731d6aae1d3aa447 (patch)
tree32419bbc5c0cf5b03624a2ede42fa3ac0429b0c7 /theories/Numbers/Integer
parentff01cafe8104f7620aacbfdde5dba738dbadc326 (diff)
Numbers: BigN and BigZ get instantiations of all properties about div and mod
NB: for declaring div and mod as a morphism, even when divisor is zero, I've slightly changed the definition of div_eucl: it now starts by a check of whether the divisor is zero. Not very nice, but this way we can say that BigN.div and BigZ.div _always_ answer like Zdiv.Zdiv. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@12646 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'theories/Numbers/Integer')
-rw-r--r--theories/Numbers/Integer/BigZ/BigZ.v7
-rw-r--r--theories/Numbers/Integer/BigZ/ZMake.v50
-rw-r--r--theories/Numbers/Integer/SpecViaZ/ZSig.v7
-rw-r--r--theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v55
4 files changed, 88 insertions, 31 deletions
diff --git a/theories/Numbers/Integer/BigZ/BigZ.v b/theories/Numbers/Integer/BigZ/BigZ.v
index 1ec7960ae..fc94f693a 100644
--- a/theories/Numbers/Integer/BigZ/BigZ.v
+++ b/theories/Numbers/Integer/BigZ/BigZ.v
@@ -11,10 +11,7 @@
(*i $Id$ i*)
Require Export BigN.
-Require Import ZProperties.
-Require Import ZSig.
-Require Import ZSigZAxioms.
-Require Import ZMake.
+Require Import ZProperties ZDivFloor ZSig ZSigZAxioms ZMake.
Module BigZ <: ZType := ZMake.Make BigN.
@@ -22,6 +19,7 @@ Module BigZ <: ZType := ZMake.Make BigN.
Module Export BigZAxiomsMod := ZSig_ZAxioms BigZ.
Module Export BigZPropMod := ZPropFunct BigZAxiomsMod.
+Module Export BigZDivPropMod := ZDivPropFunct BigZAxiomsMod BigZPropMod.
(** Notations about [BigZ] *)
@@ -71,6 +69,7 @@ Infix "<=" := BigZ.le : bigZ_scope.
Notation "x > y" := (BigZ.lt y x)(only parsing) : bigZ_scope.
Notation "x >= y" := (BigZ.le y x)(only parsing) : bigZ_scope.
Notation "[ i ]" := (BigZ.to_Z i) : bigZ_scope.
+Infix "mod" := modulo (at level 40, no associativity) : bigN_scope.
Local Open Scope bigZ_scope.
diff --git a/theories/Numbers/Integer/BigZ/ZMake.v b/theories/Numbers/Integer/BigZ/ZMake.v
index 827877fc5..0ab509650 100644
--- a/theories/Numbers/Integer/BigZ/ZMake.v
+++ b/theories/Numbers/Integer/BigZ/ZMake.v
@@ -369,17 +369,17 @@ Module Make (N:NType) <: ZType.
end.
- Theorem spec_div_eucl: forall x y,
+ Theorem spec_div_eucl_nz: forall x y,
to_Z y <> 0 ->
let (q,r) := div_eucl x y in
(to_Z q, to_Z r) = Zdiv_eucl (to_Z x) (to_Z y).
unfold div_eucl, to_Z; intros [x | x] [y | y] H.
assert (H1: 0 < N.to_Z y).
generalize (N.spec_pos y); auto with zarith.
- generalize (N.spec_div_eucl x y H1); case N.div_eucl; auto.
+ generalize (N.spec_div_eucl x y); case N.div_eucl; auto.
assert (HH: 0 < N.to_Z y).
generalize (N.spec_pos y); auto with zarith.
- generalize (N.spec_div_eucl x y HH); case N.div_eucl; auto.
+ generalize (N.spec_div_eucl x y); case N.div_eucl; auto.
intros q r; generalize (N.spec_pos x) HH; unfold Zdiv_eucl;
case_eq (N.to_Z x); case_eq (N.to_Z y);
try (intros; apply False_ind; auto with zarith; fail).
@@ -411,7 +411,7 @@ Module Make (N:NType) <: ZType.
intros; apply False_ind; auto with zarith.
assert (HH: 0 < N.to_Z y).
generalize (N.spec_pos y); auto with zarith.
- generalize (N.spec_div_eucl x y HH); case N.div_eucl; auto.
+ generalize (N.spec_div_eucl x y); case N.div_eucl; auto.
intros q r; generalize (N.spec_pos x) HH; unfold Zdiv_eucl;
case_eq (N.to_Z x); case_eq (N.to_Z y);
try (intros; apply False_ind; auto with zarith; fail).
@@ -441,7 +441,7 @@ Module Make (N:NType) <: ZType.
rewrite N.spec_0; generalize (N.spec_pos r); intros; apply False_ind; auto with zarith.
assert (H1: 0 < N.to_Z y).
generalize (N.spec_pos y); auto with zarith.
- generalize (N.spec_div_eucl x y H1); case N.div_eucl; auto.
+ generalize (N.spec_div_eucl x y); case N.div_eucl; auto.
intros q r; generalize (N.spec_pos x) H1; unfold Zdiv_eucl;
case_eq (N.to_Z x); case_eq (N.to_Z y);
try (intros; apply False_ind; auto with zarith; fail).
@@ -455,11 +455,43 @@ Module Make (N:NType) <: ZType.
rewrite <- H2; auto.
Qed.
+ Lemma Zdiv_eucl_0 : forall a, Zdiv_eucl a 0 = (0,0).
+ Proof. destruct a; auto. Qed.
+
+ Theorem spec_div_eucl: forall x y,
+ let (q,r) := div_eucl x y in
+ (to_Z q, to_Z r) = Zdiv_eucl (to_Z x) (to_Z y).
+ Proof.
+ intros. destruct (Z_eq_dec (to_Z y) 0) as [EQ|NEQ];
+ [|apply spec_div_eucl_nz; auto].
+ unfold div_eucl.
+ destruct x; destruct y; simpl in *.
+ generalize (N.spec_div_eucl t0 t1). destruct N.div_eucl; simpl; auto.
+ generalize (N.spec_div_eucl t0 t1). destruct N.div_eucl; simpl; auto.
+ assert (EQ' : N.to_Z t1 = 0) by auto with zarith.
+ rewrite EQ'. simpl. rewrite Zdiv_eucl_0. injection 1; intros.
+ generalize (N.spec_compare N.zero t3); destruct N.compare.
+ simpl. intros. f_equal; auto with zarith.
+ rewrite N.spec_0; intro; exfalso; auto with zarith.
+ rewrite N.spec_0; intro; exfalso; auto with zarith.
+ generalize (N.spec_div_eucl t0 t1). destruct N.div_eucl; simpl; auto.
+ assert (EQ' : N.to_Z t1 = 0) by auto with zarith.
+ rewrite EQ'. simpl. rewrite 2 Zdiv_eucl_0. injection 1; intros.
+ generalize (N.spec_compare N.zero t3); destruct N.compare.
+ simpl. intros. f_equal; auto with zarith.
+ rewrite N.spec_0; intro; exfalso; auto with zarith.
+ rewrite N.spec_0; intro; exfalso; auto with zarith.
+ generalize (N.spec_div_eucl t0 t1). destruct N.div_eucl; simpl; auto.
+ assert (EQ' : N.to_Z t1 = 0) by auto with zarith.
+ rewrite EQ'. simpl. rewrite 2 Zdiv_eucl_0. injection 1; intros.
+ f_equal; auto with zarith.
+ Qed.
+
Definition div x y := fst (div_eucl x y).
Definition spec_div: forall x y,
- to_Z y <> 0 -> to_Z (div x y) = to_Z x / to_Z y.
- intros x y H1; generalize (spec_div_eucl x y H1); unfold div, Zdiv.
+ to_Z (div x y) = to_Z x / to_Z y.
+ intros x y; generalize (spec_div_eucl x y); unfold div, Zdiv.
case div_eucl; case Zdiv_eucl; simpl; auto.
intros q r q11 r1 H; injection H; auto.
Qed.
@@ -467,8 +499,8 @@ Module Make (N:NType) <: ZType.
Definition modulo x y := snd (div_eucl x y).
Theorem spec_modulo:
- forall x y, to_Z y <> 0 -> to_Z (modulo x y) = to_Z x mod to_Z y.
- intros x y H1; generalize (spec_div_eucl x y H1); unfold modulo, Zmod.
+ forall x y, to_Z (modulo x y) = to_Z x mod to_Z y.
+ intros x y; generalize (spec_div_eucl x y); unfold modulo, Zmod.
case div_eucl; case Zdiv_eucl; simpl; auto.
intros q r q11 r1 H; injection H; auto.
Qed.
diff --git a/theories/Numbers/Integer/SpecViaZ/ZSig.v b/theories/Numbers/Integer/SpecViaZ/ZSig.v
index ef3cd5c34..a7c5473aa 100644
--- a/theories/Numbers/Integer/SpecViaZ/ZSig.v
+++ b/theories/Numbers/Integer/SpecViaZ/ZSig.v
@@ -98,17 +98,16 @@ Module Type ZType.
Parameter div_eucl : t -> t -> t * t.
- Parameter spec_div_eucl: forall x y, [y] <> 0 ->
+ Parameter spec_div_eucl: forall x y,
let (q,r) := div_eucl x y in ([q], [r]) = Zdiv_eucl [x] [y].
Parameter div : t -> t -> t.
- Parameter spec_div: forall x y, [y] <> 0 -> [div x y] = [x] / [y].
+ Parameter spec_div: forall x y, [div x y] = [x] / [y].
Parameter modulo : t -> t -> t.
- Parameter spec_modulo: forall x y, [y] <> 0 ->
- [modulo x y] = [x] mod [y].
+ Parameter spec_modulo: forall x y, [modulo x y] = [x] mod [y].
Parameter gcd : t -> t -> t.
diff --git a/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v b/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
index 80166d5bf..cb5a771da 100644
--- a/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
+++ b/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
@@ -8,11 +8,15 @@
(*i $Id$ i*)
-Require Import ZArith ZAxioms ZSgnAbs ZSig.
+Require Import ZArith ZAxioms ZDivFloor ZSig.
-(** * The interface [ZSig.ZType] implies the interface [ZAxiomsSig] *)
+(** * The interface [ZSig.ZType] implies the interface [ZAxiomsSig]
-Module ZSig_ZAxioms (Z:ZType) <: ZAxiomsExtSig.
+ It also provides [sgn], [abs], [div], [mod]
+*)
+
+
+Module ZSig_ZAxioms (Z:ZType) <: ZAxiomsSig <: ZDivSig.
Local Notation "[ x ]" := (Z.to_Z x).
Local Infix "==" := Z.eq (at level 70).
@@ -26,7 +30,7 @@ Local Infix "<" := Z.lt.
Hint Rewrite
Z.spec_0 Z.spec_1 Z.spec_add Z.spec_sub Z.spec_pred Z.spec_succ
- Z.spec_mul Z.spec_opp Z.spec_of_Z : zspec.
+ Z.spec_mul Z.spec_opp Z.spec_of_Z Z.spec_div Z.spec_modulo: zspec.
Ltac zsimpl := unfold Z.eq in *; autorewrite with zspec.
Ltac zcongruence := repeat red; intros; zsimpl; congruence.
@@ -235,32 +239,53 @@ Qed.
Theorem abs_eq : forall n, 0 <= n -> Z.abs n == n.
Proof.
-intros. red. rewrite Z.spec_abs. apply Zabs_eq.
- rewrite <- Z.spec_0, <- spec_le; auto.
+intros n. red. rewrite spec_le, Z.spec_0, Z.spec_abs. apply Zabs_eq.
Qed.
Theorem abs_neq : forall n, n <= 0 -> Z.abs n == -n.
Proof.
-intros. red. rewrite Z.spec_abs, Z.spec_opp. apply Zabs_non_eq.
- rewrite <- Z.spec_0, <- spec_le; auto.
+intros n. red. rewrite spec_le, Z.spec_0, Z.spec_abs, Z.spec_opp.
+ apply Zabs_non_eq.
Qed.
Theorem sgn_null : forall n, n==0 -> Z.sgn n == 0.
Proof.
-intros. red. rewrite Z.spec_sgn, Z.spec_0. rewrite Zsgn_null.
- rewrite <- Z.spec_0; auto.
+intros n. unfold Z.eq. rewrite Z.spec_sgn, Z.spec_0. rewrite Zsgn_null; auto.
Qed.
Theorem sgn_pos : forall n, 0<n -> Z.sgn n == Z.succ 0.
Proof.
-intros. red. rewrite Z.spec_sgn. zsimpl. rewrite Zsgn_pos.
- rewrite <- Z.spec_0, <- spec_lt; auto.
+intros n. red. rewrite spec_lt, Z.spec_sgn. zsimpl. rewrite Zsgn_pos; auto.
Qed.
Theorem sgn_neg : forall n, n<0 -> Z.sgn n == Z.opp (Z.succ 0).
Proof.
-intros. red. rewrite Z.spec_sgn. zsimpl. rewrite Zsgn_neg.
- rewrite <- Z.spec_0, <- spec_lt; auto.
+intros n. red. rewrite spec_lt, Z.spec_sgn. zsimpl. rewrite Zsgn_neg; auto.
+Qed.
+
+Program Instance div_wd : Proper (Z.eq==>Z.eq==>Z.eq) Z.div.
+Program Instance mod_wd : Proper (Z.eq==>Z.eq==>Z.eq) Z.modulo.
+
+Theorem div_mod : forall a b, ~b==0 -> a == b*(Z.div a b) + (Z.modulo a b).
+Proof.
+intros a b. unfold Z.eq; zsimpl. intros.
+apply Z_div_mod_eq_full; auto.
+Qed.
+
+Theorem mod_pos_bound :
+ forall a b, 0 < b -> 0 <= Z.modulo a b /\ Z.modulo a b < b.
+Proof.
+intros a b. rewrite 2 spec_lt, spec_le, Z.spec_0. intros.
+rewrite Z.spec_modulo; auto with zarith.
+apply Z_mod_lt; auto with zarith.
+Qed.
+
+Theorem mod_neg_bound :
+ forall a b, b < 0 -> b < Z.modulo a b /\ Z.modulo a b <= 0.
+Proof.
+intros a b. rewrite 2 spec_lt, spec_le, Z.spec_0. intros.
+rewrite Z.spec_modulo; auto with zarith.
+apply Z_mod_neg; auto with zarith.
Qed.
(** Aliases *)
@@ -280,5 +305,7 @@ Definition min := Z.min.
Definition max := Z.max.
Definition abs := Z.abs.
Definition sgn := Z.sgn.
+Definition div := Z.div.
+Definition modulo := Z.modulo.
End ZSig_ZAxioms.