From 51c497b2e5a2b09788f2cf05f414634b037f52bf Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Tue, 23 Apr 2019 15:00:41 +0200 Subject: lib/Coqlib.v: remove defns about multiplication, division, modulus Instead, use definitions and lemmas from the Coq standard library (ZArith, Znumtheory). --- backend/SelectDivproof.v | 28 +++++++++++++++------------- 1 file changed, 15 insertions(+), 13 deletions(-) (limited to 'backend/SelectDivproof.v') diff --git a/backend/SelectDivproof.v b/backend/SelectDivproof.v index 9e24857a..e660677a 100644 --- a/backend/SelectDivproof.v +++ b/backend/SelectDivproof.v @@ -57,13 +57,13 @@ Proof. apply Z.mul_nonneg_nonneg; omega. assert (k * n <= two_p (N + l) - two_p l). apply Z.le_trans with (two_p l * n). - apply Zmult_le_compat_r. omega. omega. + apply Z.mul_le_mono_nonneg_r; omega. replace (N + l) with (l + N) by omega. rewrite two_p_is_exp. replace (two_p l * two_p N - two_p l) with (two_p l * (two_p N - 1)) by ring. - apply Zmult_le_compat_l. omega. exploit (two_p_gt_ZERO l). omega. omega. + apply Z.mul_le_mono_nonneg_l. omega. exploit (two_p_gt_ZERO l). omega. omega. omega. omega. assert (0 <= two_p (N + l) * r). apply Z.mul_nonneg_nonneg. @@ -72,7 +72,7 @@ Proof. assert (two_p (N + l) * r <= two_p (N + l) * d - two_p (N + l)). replace (two_p (N + l) * d - two_p (N + l)) with (two_p (N + l) * (d - 1)) by ring. - apply Zmult_le_compat_l. + apply Z.mul_le_mono_nonneg_l. omega. exploit (two_p_gt_ZERO (N + l)). omega. omega. assert (0 <= m * n - two_p (N + l) * q). @@ -138,13 +138,13 @@ Proof. rewrite H2. assert (k * n <= two_p (N + l)). rewrite Z.add_comm. rewrite two_p_is_exp; try omega. - apply Z.le_trans with (two_p l * n). apply Zmult_le_compat_r. omega. omega. - apply Zmult_le_compat_l. omega. exploit (two_p_gt_ZERO l). omega. omega. + apply Z.le_trans with (two_p l * n). apply Z.mul_le_mono_nonneg_r; omega. + apply Z.mul_le_mono_nonneg_l. omega. exploit (two_p_gt_ZERO l). omega. omega. assert (two_p (N + l) * r <= two_p (N + l) * d - two_p (N + l)). replace (two_p (N + l) * d - two_p (N + l)) with (two_p (N + l) * (d - 1)) by ring. - apply Zmult_le_compat_l. omega. exploit (two_p_gt_ZERO (N + l)). omega. omega. + apply Z.mul_le_mono_nonneg_l. exploit (two_p_gt_ZERO (N + l)). omega. omega. omega. omega. Qed. @@ -246,10 +246,11 @@ Proof. unfold Int.max_signed; omega. apply Zdiv_interval_1. generalize Int.min_signed_neg; omega. apply Int.half_modulus_pos. apply Int.modulus_pos. - split. apply Z.le_trans with (Int.min_signed * m). apply Zmult_le_compat_l_neg. omega. generalize Int.min_signed_neg; omega. - apply Zmult_le_compat_r. unfold n; generalize (Int.signed_range x); tauto. tauto. + split. apply Z.le_trans with (Int.min_signed * m). + apply Z.mul_le_mono_nonpos_l. generalize Int.min_signed_neg; omega. omega. + apply Z.mul_le_mono_nonneg_r. omega. unfold n; generalize (Int.signed_range x); tauto. apply Z.le_lt_trans with (Int.half_modulus * m). - apply Zmult_le_compat_r. generalize (Int.signed_range x); unfold n, Int.max_signed; omega. tauto. + apply Z.mul_le_mono_nonneg_r. tauto. generalize (Int.signed_range x); unfold n, Int.max_signed; omega. apply Zmult_lt_compat_l. generalize Int.half_modulus_pos; omega. tauto. assert (32 < Int.max_unsigned) by (compute; auto). omega. unfold Int.lt; fold n. rewrite Int.signed_zero. destruct (zlt n 0); apply Int.eqm_unsigned_repr. @@ -290,7 +291,7 @@ Proof. apply Int.eqm_sym. eapply Int.eqm_trans. apply Int.eqm_signed_unsigned. apply Int.eqm_unsigned_repr_l. apply Int.eqm_refl2. apply (f_equal (fun x => n * x / Int.modulus)). - rewrite Int.signed_repr_eq. rewrite Zmod_small by assumption. + rewrite Int.signed_repr_eq. rewrite Z.mod_small by assumption. apply zlt_false. assumption. Qed. @@ -400,8 +401,9 @@ Proof. unfold Int64.max_signed; omega. apply Zdiv_interval_1. generalize Int64.min_signed_neg; omega. apply Int64.half_modulus_pos. apply Int64.modulus_pos. - split. apply Z.le_trans with (Int64.min_signed * m). apply Zmult_le_compat_l_neg. omega. generalize Int64.min_signed_neg; omega. - apply Zmult_le_compat_r. unfold n; generalize (Int64.signed_range x); tauto. tauto. + split. apply Z.le_trans with (Int64.min_signed * m). + apply Z.mul_le_mono_nonpos_l. generalize Int64.min_signed_neg; omega. omega. + apply Z.mul_le_mono_nonneg_r. tauto. unfold n; generalize (Int64.signed_range x); tauto. apply Z.le_lt_trans with (Int64.half_modulus * m). apply Zmult_le_compat_r. generalize (Int64.signed_range x); unfold n, Int64.max_signed; omega. tauto. apply Zmult_lt_compat_l. generalize Int64.half_modulus_pos; omega. tauto. @@ -444,7 +446,7 @@ Proof. apply Int64.eqm_sym. eapply Int64.eqm_trans. apply Int64.eqm_signed_unsigned. apply Int64.eqm_unsigned_repr_l. apply Int64.eqm_refl2. apply (f_equal (fun x => n * x / Int64.modulus)). - rewrite Int64.signed_repr_eq. rewrite Zmod_small by assumption. + rewrite Int64.signed_repr_eq. rewrite Z.mod_small by assumption. apply zlt_false. assumption. Qed. -- cgit From 08fd5faf30c18e17caa610076e67cf002a01d8b4 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Wed, 24 Apr 2019 14:02:32 +0200 Subject: Move Z definitions out of Integers and into Zbits The module Integers.Make contained lots of definitions and theorems about Z integers that were independent of the word size. These definitions and theorems are useful outside Integers.Make, but it felt unnatural to fetch them from modules Int or Int64. This commit moves the word-size-independent definitions and theorems to a new module, lib/Zbits.v, and fixes their uses in the code base. --- backend/SelectDivproof.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) (limited to 'backend/SelectDivproof.v') diff --git a/backend/SelectDivproof.v b/backend/SelectDivproof.v index e660677a..f4ff2c86 100644 --- a/backend/SelectDivproof.v +++ b/backend/SelectDivproof.v @@ -12,7 +12,7 @@ (** Correctness of instruction selection for integer division *) -Require Import Zquot Coqlib. +Require Import Zquot Coqlib Zbits. Require Import AST Integers Floats Values Memory Globalenvs Events. Require Import Cminor Op CminorSel. Require Import SelectOp SelectOpproof SplitLong SplitLongproof SelectLong SelectLongproof SelectDiv. @@ -378,7 +378,7 @@ Qed. Remark int64_shr'_div_two_p: forall x y, Int64.shr' x y = Int64.repr (Int64.signed x / two_p (Int.unsigned y)). Proof. - intros; unfold Int64.shr'. rewrite Int64.Zshiftr_div_two_p; auto. generalize (Int.unsigned_range y); omega. + intros; unfold Int64.shr'. rewrite Zshiftr_div_two_p; auto. generalize (Int.unsigned_range y); omega. Qed. Lemma divls_mul_shift_gen: @@ -453,7 +453,7 @@ Qed. Remark int64_shru'_div_two_p: forall x y, Int64.shru' x y = Int64.repr (Int64.unsigned x / two_p (Int.unsigned y)). Proof. - intros; unfold Int64.shru'. rewrite Int64.Zshiftr_div_two_p; auto. generalize (Int.unsigned_range y); omega. + intros; unfold Int64.shru'. rewrite Zshiftr_div_two_p; auto. generalize (Int.unsigned_range y); omega. Qed. Theorem divlu_mul_shift: -- cgit