Skip to content

Commit c88602b

Browse files
committed
Improve value analysis for divu and divs
Some unsigned or signed intervals can be inferred for the result.
1 parent 13f84fc commit c88602b

File tree

2 files changed

+74
-9
lines changed

2 files changed

+74
-9
lines changed

backend/ValueDomain.v

+36-9
Original file line numberDiff line numberDiff line change
@@ -1954,6 +1954,10 @@ Definition divs (v w: aval) :=
19541954
|| Int.eq i1 (Int.repr Int.min_signed) && Int.eq i2 Int.mone
19551955
then if va_strict tt then Vbot else ntop
19561956
else I (Int.divs i1 i2)
1957+
| (I i2 | IU i2), _ =>
1958+
if Int.eq i2 Int.zero
1959+
then if va_strict tt then Vbot else ntop
1960+
else sgn (provenance v) (srange v + 1 - Z.log2 (Z.abs (Int.signed i2)))
19571961
| _, _ => ntop2 v w
19581962
end.
19591963

@@ -1962,24 +1966,47 @@ Lemma divs_sound:
19621966
Proof.
19631967
intros. destruct v; destruct w; try discriminate; simpl in H1.
19641968
destruct orb eqn:E; inv H1.
1965-
inv H; inv H0; auto with va; simpl; rewrite E; constructor.
1969+
rename i0 into j.
1970+
assert (E': Int.eq j Int.zero = false). { apply orb_false_elim in E. tauto. }
1971+
assert (Int.signed j <> 0).
1972+
{ red; intros. rewrite <- (Int.repr_signed j) in E. rewrite H1 in E. discriminate. }
1973+
set (q := srange x + 1 - Z.log2 (Z.abs (Int.signed j))).
1974+
set (q1 := Z.max 0 ((srange x - 1) + 1 - Z.log2 (Z.abs (Int.signed j)))).
1975+
assert (Z.max 1 q - 1 = q1) by lia.
1976+
assert (vmatch (Vint (Int.divs i j)) (sgn (provenance x) q)).
1977+
{ apply srange_sound in H. destruct H as [A B]. apply range_is_sgn in B; auto.
1978+
apply vmatch_sgn'. apply is_sgn_range. lia.
1979+
rewrite ! H2. apply Zdiv_signed_range; auto. lia. }
1980+
unfold divs; inv H; inv H0; auto with va; rewrite ? E, ? E'; auto with va.
19661981
Qed.
19671982

19681983
Definition divu (v w: aval) :=
1969-
match w, v with
1970-
| (I i2 | IU i2), (I i1 | IU i1) =>
1971-
if Int.eq i2 Int.zero
1972-
then if va_strict tt then Vbot else ntop
1973-
else I (Int.divu i1 i2)
1974-
| _, _ => ntop2 v w
1984+
match w with
1985+
| I i2 | IU i2 =>
1986+
if Int.eq i2 Int.zero then
1987+
if va_strict tt then Vbot else ntop
1988+
else
1989+
match v with
1990+
| I i1 | IU i1 => I (Int.divu i1 i2)
1991+
| _ => uns (provenance v) (urange v - Z.log2 (Int.unsigned i2))
1992+
end
1993+
| _ => ntop2 v w
19751994
end.
19761995

19771996
Lemma divu_sound:
19781997
forall v w u x y, vmatch v x -> vmatch w y -> Val.divu v w = Some u -> vmatch u (divu x y).
19791998
Proof.
19801999
intros. destruct v; destruct w; try discriminate; simpl in H1.
1981-
destruct (Int.eq i0 Int.zero) eqn:E; inv H1.
1982-
inv H; inv H0; auto with va; simpl; rewrite E; constructor.
2000+
rename i0 into j. destruct (Int.eq j Int.zero) eqn:E; inv H1.
2001+
assert (Int.unsigned j <> 0).
2002+
{ red; intros. rewrite <- (Int.repr_unsigned j) in E. rewrite H1 in E. discriminate. }
2003+
assert (0 < Int.unsigned j).
2004+
{ pose proof (Int.unsigned_range j). lia. }
2005+
assert (vmatch (Vint (Int.divu i j)) (uns (provenance x) (urange x - Z.log2 (Int.unsigned j)))).
2006+
{ apply urange_sound in H. destruct H as [A B]. apply range_is_uns in B; auto.
2007+
apply vmatch_uns'. apply is_uns_range. lia.
2008+
apply Zdiv_unsigned_range; auto. }
2009+
unfold divu; inv H; inv H0; auto with va; rewrite E; auto with va.
19832010
Qed.
19842011

19852012
Definition mods (v w: aval) :=

lib/Zbits.v

+38
Original file line numberDiff line numberDiff line change
@@ -1147,3 +1147,41 @@ Proof.
11471147
apply two_p_monotone_strict. lia. }
11481148
lia.
11491149
Qed.
1150+
1151+
Lemma Zdiv_unsigned_range: forall n x y,
1152+
0 <= n -> 0 <= x < two_p n -> 0 < y ->
1153+
0 <= x / y < two_p (Z.max 0 (n - Z.log2 y)).
1154+
Proof.
1155+
intros. set (m := Z.log2 y).
1156+
assert (two_p m <= y).
1157+
{ rewrite two_p_correct. apply Z.log2_spec; auto. }
1158+
assert (0 <= m) by (apply Z.log2_nonneg).
1159+
rewrite Zmax_spec. destruct zlt.
1160+
- simpl. rewrite Zdiv_small. lia.
1161+
assert (two_p n <= two_p m) by (apply two_p_monotone; lia).
1162+
lia.
1163+
- split.
1164+
apply Z.div_pos; lia.
1165+
apply Z.div_lt_upper_bound; auto.
1166+
apply Z.lt_le_trans with (two_p m * two_p (n - m)).
1167+
rewrite <- two_p_is_exp by lia. replace (m + (n - m)) with n by lia. lia.
1168+
apply Z.mul_le_mono_nonneg_r; auto.
1169+
assert (two_p (n - m) > 0) by (apply two_p_gt_ZERO; lia). lia.
1170+
Qed.
1171+
1172+
Lemma Zdiv_signed_range: forall n x y,
1173+
0 <= n -> - two_p n <= x < two_p n -> y <> 0 ->
1174+
let q := Z.max 0 (n + 1 - Z.log2 (Z.abs y)) in
1175+
- two_p q <= Z.quot x y < two_p q.
1176+
Proof.
1177+
intros.
1178+
assert (Z.abs x / Z.abs y < two_p q).
1179+
{ apply Zdiv_unsigned_range; auto. lia.
1180+
assert (two_p n < two_p (n + 1)) by (apply two_p_monotone_strict; lia).
1181+
lia.
1182+
lia. }
1183+
assert (Z.abs (Z.quot x y) < two_p q).
1184+
{ rewrite <- Z.quot_abs by lia.
1185+
rewrite Z.quot_div_nonneg by lia. lia. }
1186+
lia.
1187+
Qed.

0 commit comments

Comments
 (0)