From ec0fa1ac249a8eeb0df9700c50a3e6c4f1b540f2 Mon Sep 17 00:00:00 2001 From: James Pollard Date: Tue, 23 Jun 2020 23:00:08 +0100 Subject: Normalise entire expression to avoid overflow issues. --- src/common/Coquplib.v | 10 + src/common/IntegerExtra.v | 235 +++++++++++++++++++ src/translation/HTLgen.v | 11 +- src/translation/HTLgenproof.v | 532 +++++++++++++++++++++--------------------- src/translation/HTLgenspec.v | 7 +- 5 files changed, 523 insertions(+), 272 deletions(-) create mode 100644 src/common/IntegerExtra.v (limited to 'src') diff --git a/src/common/Coquplib.v b/src/common/Coquplib.v index 59b23ae..b4ca906 100644 --- a/src/common/Coquplib.v +++ b/src/common/Coquplib.v @@ -84,6 +84,11 @@ Ltac unfold_constants := | [ H : context[Integers.Ptrofs.max_signed] |- _ ] => replace Integers.Ptrofs.max_signed with 2147483647 in H by reflexivity + | [ _ : _ |- context[Integers.Ptrofs.max_unsigned] ] => + replace Integers.Ptrofs.max_unsigned with 4294967295 by reflexivity + | [ H : context[Integers.Ptrofs.max_unsigned] |- _ ] => + replace Integers.Ptrofs.max_unsigned with 4294967295 in H by reflexivity + | [ _ : _ |- context[Integers.Int.modulus] ] => replace Integers.Int.modulus with 4294967296 by reflexivity | [ H : context[Integers.Int.modulus] |- _ ] => @@ -98,6 +103,11 @@ Ltac unfold_constants := replace Integers.Int.max_signed with 2147483647 by reflexivity | [ H : context[Integers.Int.max_signed] |- _ ] => replace Integers.Int.max_signed with 2147483647 in H by reflexivity + + | [ _ : _ |- context[Integers.Int.max_unsigned] ] => + replace Integers.Int.max_unsigned with 4294967295 by reflexivity + | [ H : context[Integers.Int.max_unsigned] |- _ ] => + replace Integers.Int.max_unsigned with 4294967295 in H by reflexivity end. Ltac simplify := unfold_constants; simpl in *; diff --git a/src/common/IntegerExtra.v b/src/common/IntegerExtra.v new file mode 100644 index 0000000..ad01015 --- /dev/null +++ b/src/common/IntegerExtra.v @@ -0,0 +1,235 @@ +Require Import BinInt. +Require Import Lia. +Require Import ZBinary. + +From bbv Require Import ZLib. +From compcert Require Import Integers Coqlib. + +Require Import Coquplib. + +Local Open Scope Z_scope. + +Module PtrofsExtra. + + Lemma mul_divs : + forall x y, + 0 <= Ptrofs.signed y -> + 0 < Ptrofs.signed x -> + Ptrofs.signed y mod Ptrofs.signed x = 0 -> + (Integers.Ptrofs.mul x (Integers.Ptrofs.divs y x)) = y. + Proof. + intros. + + pose proof (Ptrofs.mods_divs_Euclid y x). + pose proof (Zquot.Zrem_Zmod_zero (Ptrofs.signed y) (Ptrofs.signed x)). + apply <- H3 in H1; try lia; clear H3. + unfold Ptrofs.mods in H2. + rewrite H1 in H2. + replace (Ptrofs.repr 0) with (Ptrofs.zero) in H2 by reflexivity. + rewrite Ptrofs.add_zero in H2. + rewrite Ptrofs.mul_commut. + congruence. + Qed. + + Lemma Z_div_distr_add : + forall x y z, + x mod z = 0 -> + y mod z = 0 -> + z <> 0 -> + x / z + y / z = (x + y) / z. + Proof. + intros. + + assert ((x + y) mod z = 0). + { rewrite <- Z.add_mod_idemp_l; try assumption. + rewrite H. assumption. } + + rewrite <- Z.mul_cancel_r with (p := z); try assumption. + rewrite Z.mul_add_distr_r. + repeat rewrite ZLib.div_mul_undo; lia. + Qed. + + Lemma mul_unsigned : + forall x y, + Ptrofs.mul x y = + Ptrofs.repr (Ptrofs.unsigned x * Ptrofs.unsigned y). + Proof. + intros; unfold Ptrofs.mul. + apply Ptrofs.eqm_samerepr. + apply Ptrofs.eqm_mult; exists 0; lia. + Qed. + + Lemma mul_repr : + forall x y, + Ptrofs.min_signed <= y <= Ptrofs.max_signed -> + Ptrofs.min_signed <= x <= Ptrofs.max_signed -> + Ptrofs.mul (Ptrofs.repr y) (Ptrofs.repr x) = Ptrofs.repr (x * y). + Proof. + intros; unfold Ptrofs.mul. + destruct (Z_ge_lt_dec x 0); destruct (Z_ge_lt_dec y 0). + + - f_equal. + repeat rewrite Ptrofs.unsigned_repr_eq. + repeat rewrite Z.mod_small; simplify; lia. + + - assert (Ptrofs.lt (Ptrofs.repr y) Ptrofs.zero = true). + { + unfold Ptrofs.lt. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.signed_zero. + destruct (zlt y 0); try lia; auto. + } + + rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr y). + rewrite H1. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.unsigned_repr_eq. + rewrite Z.mod_small; simplify; try lia. + rewrite Z.mul_add_distr_r. + apply Ptrofs.eqm_samerepr. + exists x. simplify. lia. + + - assert (Ptrofs.lt (Ptrofs.repr x) Ptrofs.zero = true). + { + unfold Ptrofs.lt. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.signed_zero. + destruct (zlt x 0); try lia; auto. + } + + rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr x). + rewrite H1. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.unsigned_repr_eq. + rewrite Z.mod_small; simplify; try lia. + rewrite Z.mul_add_distr_l. + apply Ptrofs.eqm_samerepr. + exists y. simplify. lia. + + - assert (Ptrofs.lt (Ptrofs.repr x) Ptrofs.zero = true). + { + unfold Ptrofs.lt. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.signed_zero. + destruct (zlt x 0); try lia; auto. + } + assert (Ptrofs.lt (Ptrofs.repr y) Ptrofs.zero = true). + { + unfold Ptrofs.lt. + rewrite Ptrofs.signed_repr; auto. + rewrite Ptrofs.signed_zero. + destruct (zlt y 0); try lia; auto. + } + rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr x). + rewrite Ptrofs.unsigned_signed with (n := Ptrofs.repr y). + rewrite H2. + rewrite H1. + repeat rewrite Ptrofs.signed_repr; auto. + replace ((y + Ptrofs.modulus) * (x + Ptrofs.modulus)) with + (x * y + (x + y + Ptrofs.modulus) * Ptrofs.modulus) by lia. + apply Ptrofs.eqm_samerepr. + exists (x + y + Ptrofs.modulus). lia. + Qed. +End PtrofsExtra. + +Module IntExtra. + Lemma mul_unsigned : + forall x y, + Int.mul x y = + Int.repr (Int.unsigned x * Int.unsigned y). + Proof. + intros; unfold Int.mul. + apply Int.eqm_samerepr. + apply Int.eqm_mult; exists 0; lia. + Qed. + + Lemma mul_repr : + forall x y, + Int.min_signed <= y <= Int.max_signed -> + Int.min_signed <= x <= Int.max_signed -> + Int.mul (Int.repr y) (Int.repr x) = Int.repr (x * y). + Proof. + intros; unfold Int.mul. + destruct (Z_ge_lt_dec x 0); destruct (Z_ge_lt_dec y 0). + + - f_equal. + repeat rewrite Int.unsigned_repr_eq. + repeat rewrite Z.mod_small; simplify; lia. + + - assert (Int.lt (Int.repr y) Int.zero = true). + { + unfold Int.lt. + rewrite Int.signed_repr; auto. + rewrite Int.signed_zero. + destruct (zlt y 0); try lia; auto. + } + + rewrite Int.unsigned_signed with (n := Int.repr y). + rewrite H1. + rewrite Int.signed_repr; auto. + rewrite Int.unsigned_repr_eq. + rewrite Z.mod_small; simplify; try lia. + rewrite Z.mul_add_distr_r. + apply Int.eqm_samerepr. + exists x. simplify. lia. + + - assert (Int.lt (Int.repr x) Int.zero = true). + { + unfold Int.lt. + rewrite Int.signed_repr; auto. + rewrite Int.signed_zero. + destruct (zlt x 0); try lia; auto. + } + + rewrite Int.unsigned_signed with (n := Int.repr x). + rewrite H1. + rewrite Int.signed_repr; auto. + rewrite Int.unsigned_repr_eq. + rewrite Z.mod_small; simplify; try lia. + rewrite Z.mul_add_distr_l. + apply Int.eqm_samerepr. + exists y. simplify. lia. + + - assert (Int.lt (Int.repr x) Int.zero = true). + { + unfold Int.lt. + rewrite Int.signed_repr; auto. + rewrite Int.signed_zero. + destruct (zlt x 0); try lia; auto. + } + assert (Int.lt (Int.repr y) Int.zero = true). + { + unfold Int.lt. + rewrite Int.signed_repr; auto. + rewrite Int.signed_zero. + destruct (zlt y 0); try lia; auto. + } + rewrite Int.unsigned_signed with (n := Int.repr x). + rewrite Int.unsigned_signed with (n := Int.repr y). + rewrite H2. + rewrite H1. + repeat rewrite Int.signed_repr; auto. + replace ((y + Int.modulus) * (x + Int.modulus)) with + (x * y + (x + y + Int.modulus) * Int.modulus) by lia. + apply Int.eqm_samerepr. + exists (x + y + Int.modulus). lia. + Qed. +End IntExtra. + +Lemma mul_of_int : + forall x y, + 0 <= x < Integers.Ptrofs.modulus -> + Integers.Ptrofs.mul (Integers.Ptrofs.repr x) (Integers.Ptrofs.of_int y) = + Integers.Ptrofs.of_int (Integers.Int.mul (Integers.Int.repr x) y). +Proof. + intros. + pose proof (Integers.Ptrofs.agree32_of_int eq_refl y) as A. + pose proof (Integers.Ptrofs.agree32_to_int eq_refl (Integers.Ptrofs.repr x)) as B. + exploit Integers.Ptrofs.agree32_mul; [> reflexivity | exact B | exact A | intro C]. + unfold Integers.Ptrofs.to_int in C. + unfold Integers.Ptrofs.of_int in C. + rewrite Integers.Ptrofs.unsigned_repr_eq in C. + rewrite Z.mod_small in C; auto. + symmetry. + apply Integers.Ptrofs.agree32_of_int_eq; auto. +Qed. diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index cc6a3f8..92e40f5 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -349,17 +349,18 @@ Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) match mem, addr, args with (* TODO: We should be more methodical here; what are the possibilities?*) | Mint32, Op.Aindexed off, r1::nil => if (check_address_parameter off) - then ret (Vvari stack (Vbinop Vadd (boplitz Vdiv r1 4) (Vlit (ZToValue 32 (off / 4))))) + then ret (Vvari stack (Vbinop Vdiv (boplitz Vadd r1 off) (Vlit (ZToValue 32 4)))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") | Mint32, Op.Ascaled scale offset, r1::nil => if (check_address_parameter scale) && (check_address_parameter offset) - then ret (Vvari stack (Vbinop Vadd (boplitz Vmul r1 (scale / 4)) (Vlit (ZToValue 32 (offset / 4))))) + then ret (Vvari stack (Vbinop Vdiv (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset))) (Vlit (ZToValue 32 4)))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) if (check_address_parameter scale) && (check_address_parameter offset) then ret (Vvari stack - (Vbinop Vadd (Vbinop Vadd (boplitz Vdiv r1 4) (Vlit (ZToValue 32 (offset / 4)))) - (boplitz Vmul r2 (scale / 4)))) + (Vbinop Vdiv + (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) + (ZToValue 32 4))) else error (Errors.msg "Veriloggen: translate_arr_access address misaligned") | Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) let a := Integers.Ptrofs.unsigned a in @@ -451,7 +452,7 @@ Definition create_arr (i : option io) (sz : nat) (ln : nat) : mon (reg * nat) := (create_arr_state_incr s sz ln i). Definition stack_correct (sz : Z) : bool := - (0 <=? sz) && (Z.modulo sz 4 =? 0). + (0 <=? sz) && (sz Prop := tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) -> stk_len = Z.to_nat (f.(RTL.fn_stacksize) / 4) -> Z.modulo (f.(RTL.fn_stacksize)) 4 = 0 -> - 0 <= f.(RTL.fn_stacksize) -> + 0 <= f.(RTL.fn_stacksize) < Integers.Ptrofs.modulus -> m = (mkmodule f.(RTL.fn_params) data control @@ -454,7 +454,8 @@ Proof. unfold transf_module in *. unfold stack_correct in *. - destruct (0 <=? RTL.fn_stacksize f) eqn:STACK_BOUND; + destruct (0 <=? RTL.fn_stacksize f) eqn:STACK_BOUND_LOW; + destruct (RTL.fn_stacksize f