aboutsummaryrefslogtreecommitdiffstats
path: root/lib/Integers.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-01-14 15:35:20 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-01-14 15:35:20 +0100
commit93dc602bebb6293283981eac072852a5fcd3f51c (patch)
tree389cb67bc2b4a95f8eb545446ccc1df6aa63c6d3 /lib/Integers.v
parent9c6fac6cd52b824aaefac66089bf5c71e27845be (diff)
downloadcompcert-kvx-93dc602bebb6293283981eac072852a5fcd3f51c.tar.gz
compcert-kvx-93dc602bebb6293283981eac072852a5fcd3f51c.zip
shrx'1_shr'
Diffstat (limited to 'lib/Integers.v')
-rw-r--r--lib/Integers.v128
1 files changed, 127 insertions, 1 deletions
diff --git a/lib/Integers.v b/lib/Integers.v
index 09053cab..277f78e3 100644
--- a/lib/Integers.v
+++ b/lib/Integers.v
@@ -4,7 +4,7 @@
(* *)
(* Xavier Leroy, INRIA Paris-Rocquencourt *)
(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
+(* Copyright Institut National de Recherstestche en Informatique et en *)
(* Automatique. All rights reserved. This file is distributed *)
(* under the terms of the GNU General Public License as published by *)
(* the Free Software Foundation, either version 2 of the License, or *)
@@ -1189,6 +1189,34 @@ Proof.
rewrite <- half_modulus_modulus. apply unsigned_range.
Qed.
+Local Transparent repr.
+Lemma sign_bit_of_signed: forall x,
+ (testbit x (zwordsize - 1)) = lt x zero.
+Proof.
+ intro.
+ rewrite sign_bit_of_unsigned.
+ unfold lt.
+ unfold signed, unsigned.
+ simpl.
+ pose proof half_modulus_pos as HMOD.
+ destruct (zlt 0 half_modulus) as [HMOD' | HMOD'].
+ 2: omega.
+ clear HMOD'.
+ destruct (zlt (intval x) half_modulus) as [ LOW | HIGH].
+ {
+ destruct x as [ix RANGE].
+ simpl in *.
+ destruct (zlt ix 0). omega.
+ reflexivity.
+ }
+ destruct (zlt _ _) as [LOW' | HIGH']; trivial.
+ destruct x as [ix RANGE].
+ simpl in *.
+ rewrite half_modulus_modulus in *.
+ omega.
+Qed.
+Local Opaque repr.
+
Lemma bits_signed:
forall x i, 0 <= i ->
Z.testbit (signed x) i = testbit x (if zlt i zwordsize then i else zwordsize - 1).
@@ -3639,6 +3667,104 @@ Proof.
unfold ltu. apply zlt_true. change (unsigned z < 63). rewrite A; omega.
Qed.
+Lemma shr'63:
+ forall x, (shr' x (Int.repr 63)) = if lt x zero then mone else zero.
+Proof.
+ intro.
+ unfold shr', mone, zero.
+ rewrite Int.unsigned_repr by (change Int.max_unsigned with 4294967295; omega).
+ apply same_bits_eq.
+ intros i BIT.
+ rewrite testbit_repr by assumption.
+ rewrite Z.shiftr_spec by omega.
+ rewrite bits_signed by omega.
+ simpl.
+ change zwordsize with 64 in *.
+ destruct (zlt _ _) as [LT | GE].
+ {
+ replace i with 0 in * by omega.
+ change (0 + 63) with (zwordsize - 1).
+ rewrite sign_bit_of_signed.
+ destruct (lt x _).
+ all: rewrite testbit_repr by (change zwordsize with 64 in *; omega).
+ all: simpl; reflexivity.
+ }
+ change (64 - 1) with (zwordsize - 1).
+ rewrite sign_bit_of_signed.
+ destruct (lt x _).
+ all: rewrite testbit_repr by (change zwordsize with 64 in *; omega).
+ { symmetry.
+ apply Ztestbit_m1.
+ tauto.
+ }
+ symmetry.
+ apply Ztestbit_0.
+Qed.
+
+Lemma shru'63:
+ forall x, (shru' x (Int.repr 63)) = if lt x zero then one else zero.
+Proof.
+ intro.
+ unfold shru'.
+ rewrite Int.unsigned_repr by (change Int.max_unsigned with 4294967295; omega).
+ apply same_bits_eq.
+ intros i BIT.
+ rewrite testbit_repr by assumption.
+ rewrite Z.shiftr_spec by omega.
+ unfold lt.
+ rewrite signed_zero.
+ unfold one, zero.
+ destruct (zlt _ 0) as [LT | GE].
+ {
+ rewrite testbit_repr by assumption.
+ destruct (zeq i 0) as [IZERO | INONZERO].
+ { subst i.
+ change (Z.testbit (unsigned x) (0 + 63)) with (testbit x (zwordsize - 1)).
+ rewrite sign_bit_of_signed.
+ unfold lt.
+ rewrite signed_zero.
+ destruct (zlt _ _); try omega.
+ reflexivity.
+ }
+ change (Z.testbit (unsigned x) (i + 63)) with (testbit x (i+63)).
+ rewrite bits_above by (change zwordsize with 64; omega).
+ rewrite Ztestbit_1.
+ destruct (zeq i 0); trivial.
+ subst i.
+ omega.
+ }
+ destruct (zeq i 0) as [IZERO | INONZERO].
+ { subst i.
+ change (Z.testbit (unsigned x) (0 + 63)) with (testbit x (zwordsize - 1)).
+ rewrite sign_bit_of_signed.
+ unfold lt.
+ rewrite signed_zero.
+ rewrite bits_zero.
+ destruct (zlt _ _); try omega.
+ reflexivity.
+ }
+ change (Z.testbit (unsigned x) (i + 63)) with (testbit x (i + 63)).
+ rewrite bits_zero.
+ apply bits_above.
+ change zwordsize with 64.
+ omega.
+Qed.
+
+Theorem shrx'1_shr':
+ forall x,
+ Int.ltu Int.one (Int.repr (zwordsize - 1)) = true ->
+ shrx' x (Int.repr 1) = shr' (add x (shru' x (Int.repr (Int64.zwordsize - 1)))) (Int.repr 1).
+Proof.
+ intros.
+ rewrite shrx'_shr_2 by reflexivity.
+ change (Int.sub (Int.repr 64) (Int.repr 1)) with (Int.repr 63).
+ f_equal. f_equal.
+ rewrite shr'63.
+ rewrite shru'63.
+ rewrite shru'63.
+ destruct (lt x zero); reflexivity.
+Qed.
+
Remark int_ltu_2_inv:
forall y z,
Int.ltu y iwordsize' = true ->