From 888c41d2bf95bb84fee28a8737515c9ff66aa94e Mon Sep 17 00:00:00 2001 From: letouzey Date: Thu, 14 Oct 2010 11:37:33 +0000 Subject: Numbers: new functions pow, even, odd + many reorganisations - Simplification of functor names, e.g. ZFooProp instead of ZFooPropFunct - The axiomatisations of the different fonctions are now in {N,Z}Axioms.v apart for Z division (three separate flavours in there own files). Content of {N,Z}AxiomsSig is extended, old version is {N,Z}AxiomsMiniSig. - In NAxioms, the recursion field isn't that useful, since we axiomatize other functions and not define them (apart in the toy NDefOps.v). We leave recursion there, but in a separate NAxiomsFullSig. - On Z, the pow function is specified to behave as Zpower : a^(-1)=0 - In BigN/BigZ, (power:t->N->t) is now pow_N, while pow is t->t->t These pow could be more clever (we convert 2nd arg to N and use pow_N). Default "^" is now (pow:t->t->t). BigN/BigZ ring is adapted accordingly - In BigN, is_even is now even, its spec is changed to use Zeven_bool. We add an odd. In BigZ, we add even and odd. - In ZBinary (implem of ZAxioms by ZArith), we create an efficient Zpow to implement pow. This Zpow should replace the current linear Zpower someday. - In NPeano (implem of NAxioms by Arith), we create pow, even, odd functions, and we modify the div and mod functions for them to be linear, structural, tail-recursive. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@13546 85f007b7-540e-0410-9357-904b9bb8a0f7 --- theories/Numbers/Natural/SpecViaZ/NSig.v | 18 +++--- theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v | 76 +++++++++++++++++++++++-- 2 files changed, 82 insertions(+), 12 deletions(-) (limited to 'theories/Numbers/Natural/SpecViaZ') diff --git a/theories/Numbers/Natural/SpecViaZ/NSig.v b/theories/Numbers/Natural/SpecViaZ/NSig.v index ba2864760..3ff2ded62 100644 --- a/theories/Numbers/Natural/SpecViaZ/NSig.v +++ b/theories/Numbers/Natural/SpecViaZ/NSig.v @@ -46,8 +46,9 @@ Module Type NType. Parameter sub : t -> t -> t. Parameter mul : t -> t -> t. Parameter square : t -> t. - Parameter power_pos : t -> positive -> t. - Parameter power : t -> N -> t. + Parameter pow_pos : t -> positive -> t. + Parameter pow_N : t -> N -> t. + Parameter pow : t -> t -> t. Parameter sqrt : t -> t. Parameter log2 : t -> t. Parameter div_eucl : t -> t -> t * t. @@ -56,7 +57,8 @@ Module Type NType. Parameter gcd : t -> t -> t. Parameter shiftr : t -> t -> t. Parameter shiftl : t -> t -> t. - Parameter is_even : t -> bool. + Parameter even : t -> bool. + Parameter odd : t -> bool. Parameter spec_compare: forall x y, compare x y = Zcompare [x] [y]. Parameter spec_eq_bool: forall x y, eq_bool x y = Zeq_bool [x] [y]. @@ -70,8 +72,9 @@ Module Type NType. Parameter spec_sub: forall x y, [sub x y] = Zmax 0 ([x] - [y]). Parameter spec_mul: forall x y, [mul x y] = [x] * [y]. Parameter spec_square: forall x, [square x] = [x] * [x]. - Parameter spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n. - Parameter spec_power: forall x n, [power x n] = [x] ^ Z_of_N n. + Parameter spec_pow_pos: forall x n, [pow_pos x n] = [x] ^ Zpos n. + Parameter spec_pow_N: forall x n, [pow_N x n] = [x] ^ Z_of_N n. + Parameter spec_pow: forall x n, [pow x n] = [x] ^ [n]. Parameter spec_sqrt: forall x, [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2. Parameter spec_log2_0: forall x, [x] = 0 -> [log2 x] = 0. Parameter spec_log2: forall x, [x]<>0 -> 2^[log2 x] <= [x] < 2^([log2 x]+1). @@ -82,8 +85,8 @@ Module Type NType. Parameter spec_gcd: forall a b, [gcd a b] = Zgcd [a] [b]. Parameter spec_shiftr: forall p x, [shiftr p x] = [x] / 2^[p]. Parameter spec_shiftl: forall p x, [shiftl p x] = [x] * 2^[p]. - Parameter spec_is_even: forall x, - if is_even x then [x] mod 2 = 0 else [x] mod 2 = 1. + Parameter spec_even: forall x, even x = Zeven_bool [x]. + Parameter spec_odd: forall x, odd x = Zodd_bool [x]. End NType. @@ -94,6 +97,7 @@ Module Type NType_Notation (Import N:NType). Infix "+" := add. Infix "-" := sub. Infix "*" := mul. + Infix "^" := pow. Infix "<=" := le. Infix "<" := lt. End NType_Notation. diff --git a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v index 90dd16249..568ebeae8 100644 --- a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v +++ b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v @@ -13,13 +13,13 @@ Require Import ZArith Nnat NAxioms NDiv NSig. Module NTypeIsNAxioms (Import N : NType'). Hint Rewrite - spec_0 spec_succ spec_add spec_mul spec_pred spec_sub + spec_0 spec_1 spec_succ spec_add spec_mul spec_pred spec_sub spec_div spec_modulo spec_gcd spec_compare spec_eq_bool - spec_max spec_min spec_power_pos spec_power + spec_max spec_min spec_pow_pos spec_pow_N spec_pow spec_even spec_odd : nsimpl. Ltac nsimpl := autorewrite with nsimpl. -Ltac ncongruence := unfold eq; repeat red; intros; nsimpl; congruence. -Ltac zify := unfold eq, lt, le in *; nsimpl. +Ltac ncongruence := unfold eq, to_N; repeat red; intros; nsimpl; congruence. +Ltac zify := unfold eq, lt, le, to_N in *; nsimpl. Local Obligation Tactic := ncongruence. @@ -177,6 +177,70 @@ Proof. zify. auto. Qed. +(** Power *) + +Program Instance pow_wd : Proper (eq==>eq==>eq) pow. + +Local Notation "1" := (succ 0). +Local Notation "2" := (succ 1). + +Lemma pow_0_r : forall a, a^0 == 1. +Proof. + intros. now zify. +Qed. + +Lemma pow_succ_r : forall a b, 0<=b -> a^(succ b) == a * a^b. +Proof. + intros a b. zify. intro Hb. + rewrite Zpower_exp; auto with zarith. + simpl. unfold Zpower_pos; simpl. ring. +Qed. + +Lemma pow_pow_N : forall a b, a^b == pow_N a (to_N b). +Proof. + intros. zify. f_equal. + now rewrite Z_of_N_abs, Zabs_eq by apply spec_pos. +Qed. + +Lemma pow_N_pow : forall a b, pow_N a b == a^(of_N b). +Proof. + intros. zify. f_equal. symmetry. apply spec_of_N. +Qed. + +Lemma pow_pos_N : forall a p, pow_pos a p == pow_N a (Npos p). +Proof. + intros. now zify. +Qed. + +(** Even / Odd *) + +Definition Even n := exists m, n == 2*m. +Definition Odd n := exists m, n == 2*m+1. + +Lemma even_spec : forall n, even n = true <-> Even n. +Proof. + intros n. unfold Even. zify. + rewrite Zeven_bool_iff, Zeven_ex_iff. + split; intros (m,Hm). + exists (of_N (Zabs_N m)). + zify. rewrite spec_of_N, Z_of_N_abs, Zabs_eq; auto. + generalize (spec_pos n); auto with zarith. + exists [m]. revert Hm. now zify. +Qed. + +Lemma odd_spec : forall n, odd n = true <-> Odd n. +Proof. + intros n. unfold Odd. zify. + rewrite Zodd_bool_iff, Zodd_ex_iff. + split; intros (m,Hm). + exists (of_N (Zabs_N m)). + zify. rewrite spec_of_N, Z_of_N_abs, Zabs_eq; auto. + generalize (spec_pos n); auto with zarith. + exists [m]. revert Hm. now zify. +Qed. + +(** Div / Mod *) + Program Instance div_wd : Proper (eq==>eq==>eq) div. Program Instance mod_wd : Proper (eq==>eq==>eq) modulo. @@ -192,6 +256,8 @@ destruct (Z_mod_lt [a] [b]); auto. generalize (spec_pos b); auto with zarith. Qed. +(** Recursion *) + Definition recursion (A : Type) (a : A) (f : N.t -> A -> A) (n : N.t) := Nrect (fun _ => A) a (fun n a => f (N.of_N n) a) (N.to_N n). Implicit Arguments recursion [A]. @@ -250,5 +316,5 @@ Qed. End NTypeIsNAxioms. Module NType_NAxioms (N : NType) - <: NAxiomsSig <: NDivSig <: HasCompare N <: HasEqBool N <: HasMinMax N + <: NAxiomsSig <: HasCompare N <: HasEqBool N <: HasMinMax N := N <+ NTypeIsNAxioms. -- cgit v1.2.3