From e9f40aaca38ba81f3e9e5c0a5e03de9fa074d838 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Thu, 10 Jun 2021 09:52:47 +0200 Subject: Int.sign_ext_shr_shl: weaker hypothesis Works also for sign_ext 32. ARM, RISC-V: adapt Asmgenproof1 accordingly --- arm/Asmgenproof1.v | 4 ++-- lib/Integers.v | 4 ++-- riscV/Asmgenproof1.v | 4 ++-- 3 files changed, 6 insertions(+), 6 deletions(-) diff --git a/arm/Asmgenproof1.v b/arm/Asmgenproof1.v index b94964a0..6f0482dc 100644 --- a/arm/Asmgenproof1.v +++ b/arm/Asmgenproof1.v @@ -1218,7 +1218,7 @@ Proof. split. unfold rs2; Simpl. unfold rs1; Simpl. unfold Val.shr, Val.shl; destruct (rs x0); auto. change (Int.ltu (Int.repr 24) Int.iwordsize) with true; simpl. - f_equal. symmetry. apply (Int.sign_ext_shr_shl 8). compute; auto. + f_equal. symmetry. apply (Int.sign_ext_shr_shl 8). compute; intuition congruence. intros. unfold rs2, rs1; Simpl. (* Ocast16signed *) destruct Archi.thumb2_support. @@ -1231,7 +1231,7 @@ Proof. split. unfold rs2; Simpl. unfold rs1; Simpl. unfold Val.shr, Val.shl; destruct (rs x0); auto. change (Int.ltu (Int.repr 16) Int.iwordsize) with true; simpl. - f_equal. symmetry. apply (Int.sign_ext_shr_shl 16). compute; auto. + f_equal. symmetry. apply (Int.sign_ext_shr_shl 16). compute; intuition congruence. intros. unfold rs2, rs1; Simpl. (* Oaddimm *) generalize (addimm_correct x x0 i k rs m). diff --git a/lib/Integers.v b/lib/Integers.v index b38f8564..63dc2251 100644 --- a/lib/Integers.v +++ b/lib/Integers.v @@ -2766,7 +2766,7 @@ Qed. Corollary sign_ext_shr_shl: forall n x, - 0 < n < zwordsize -> + 0 < n <= zwordsize -> let y := repr (zwordsize - n) in sign_ext n x = shr (shl x y) y. Proof. @@ -2801,7 +2801,7 @@ Qed. Lemma sign_ext_range: forall n x, 0 < n < zwordsize -> -two_p (n-1) <= signed (sign_ext n x) < two_p (n-1). Proof. - intros. rewrite sign_ext_shr_shl; auto. + intros. rewrite sign_ext_shr_shl by lia. set (X := shl x (repr (zwordsize - n))). assert (two_p (n - 1) > 0) by (apply two_p_gt_ZERO; lia). assert (unsigned (repr (zwordsize - n)) = zwordsize - n). diff --git a/riscV/Asmgenproof1.v b/riscV/Asmgenproof1.v index 8195ce44..af53754e 100644 --- a/riscV/Asmgenproof1.v +++ b/riscV/Asmgenproof1.v @@ -1010,14 +1010,14 @@ Opaque Int.eq. split; intros; Simpl. assert (A: Int.ltu (Int.repr 24) Int.iwordsize = true) by auto. destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A. - apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. + apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. compute; intuition congruence. - (* cast16signed *) econstructor; split. eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto. split; intros; Simpl. assert (A: Int.ltu (Int.repr 16) Int.iwordsize = true) by auto. destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A. - apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. + apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. compute; intuition congruence. - (* addimm *) exploit (opimm32_correct Paddw Paddiw Val.add); auto. instantiate (1 := x0); eauto with asmgen. intros (rs' & A & B & C). -- cgit