Skip to content

Commit e9f40aa

Browse files
committed
Int.sign_ext_shr_shl: weaker hypothesis
Works also for sign_ext 32. ARM, RISC-V: adapt Asmgenproof1 accordingly
1 parent 4fe221a commit e9f40aa

File tree

3 files changed

+6
-6
lines changed

3 files changed

+6
-6
lines changed

arm/Asmgenproof1.v

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1218,7 +1218,7 @@ Proof.
12181218
split. unfold rs2; Simpl. unfold rs1; Simpl.
12191219
unfold Val.shr, Val.shl; destruct (rs x0); auto.
12201220
change (Int.ltu (Int.repr 24) Int.iwordsize) with true; simpl.
1221-
f_equal. symmetry. apply (Int.sign_ext_shr_shl 8). compute; auto.
1221+
f_equal. symmetry. apply (Int.sign_ext_shr_shl 8). compute; intuition congruence.
12221222
intros. unfold rs2, rs1; Simpl.
12231223
(* Ocast16signed *)
12241224
destruct Archi.thumb2_support.
@@ -1231,7 +1231,7 @@ Proof.
12311231
split. unfold rs2; Simpl. unfold rs1; Simpl.
12321232
unfold Val.shr, Val.shl; destruct (rs x0); auto.
12331233
change (Int.ltu (Int.repr 16) Int.iwordsize) with true; simpl.
1234-
f_equal. symmetry. apply (Int.sign_ext_shr_shl 16). compute; auto.
1234+
f_equal. symmetry. apply (Int.sign_ext_shr_shl 16). compute; intuition congruence.
12351235
intros. unfold rs2, rs1; Simpl.
12361236
(* Oaddimm *)
12371237
generalize (addimm_correct x x0 i k rs m).

lib/Integers.v

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -2766,7 +2766,7 @@ Qed.
27662766

27672767
Corollary sign_ext_shr_shl:
27682768
forall n x,
2769-
0 < n < zwordsize ->
2769+
0 < n <= zwordsize ->
27702770
let y := repr (zwordsize - n) in
27712771
sign_ext n x = shr (shl x y) y.
27722772
Proof.
@@ -2801,7 +2801,7 @@ Qed.
28012801
Lemma sign_ext_range:
28022802
forall n x, 0 < n < zwordsize -> -two_p (n-1) <= signed (sign_ext n x) < two_p (n-1).
28032803
Proof.
2804-
intros. rewrite sign_ext_shr_shl; auto.
2804+
intros. rewrite sign_ext_shr_shl by lia.
28052805
set (X := shl x (repr (zwordsize - n))).
28062806
assert (two_p (n - 1) > 0) by (apply two_p_gt_ZERO; lia).
28072807
assert (unsigned (repr (zwordsize - n)) = zwordsize - n).

riscV/Asmgenproof1.v

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1010,14 +1010,14 @@ Opaque Int.eq.
10101010
split; intros; Simpl.
10111011
assert (A: Int.ltu (Int.repr 24) Int.iwordsize = true) by auto.
10121012
destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A.
1013-
apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity.
1013+
apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. compute; intuition congruence.
10141014
- (* cast16signed *)
10151015
econstructor; split.
10161016
eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto.
10171017
split; intros; Simpl.
10181018
assert (A: Int.ltu (Int.repr 16) Int.iwordsize = true) by auto.
10191019
destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A.
1020-
apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity.
1020+
apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. compute; intuition congruence.
10211021
- (* addimm *)
10221022
exploit (opimm32_correct Paddw Paddiw Val.add); auto. instantiate (1 := x0); eauto with asmgen.
10231023
intros (rs' & A & B & C).

0 commit comments

Comments
 (0)