aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/ExtValues.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-04-27 19:23:04 +0200
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-04-27 19:23:04 +0200
commit146bf43966f0d0c7a1587fc4d8dab58958d621fa (patch)
treeb95de34ff87fcbd679056a744e915c3a31e4e25c /mppa_k1c/ExtValues.v
parent28e66eb4b60c485bebb1e217bc8f50bdc2cc6ddb (diff)
downloadcompcert-kvx-146bf43966f0d0c7a1587fc4d8dab58958d621fa.tar.gz
compcert-kvx-146bf43966f0d0c7a1587fc4d8dab58958d621fa.zip
factor expressions into single file
Diffstat (limited to 'mppa_k1c/ExtValues.v')
-rw-r--r--mppa_k1c/ExtValues.v26
1 files changed, 14 insertions, 12 deletions
diff --git a/mppa_k1c/ExtValues.v b/mppa_k1c/ExtValues.v
index d1329583..97e48eb7 100644
--- a/mppa_k1c/ExtValues.v
+++ b/mppa_k1c/ExtValues.v
@@ -2,10 +2,13 @@ Require Import Coqlib.
Require Import Integers.
Require Import Values.
+Definition is_bitfield stop start :=
+ (Z.leb start stop)
+ && (Z.geb start Z.zero)
+ && (Z.ltb stop Int.zwordsize).
+
Definition extfz stop start v :=
- if (Z.leb start stop)
- && (Z.geb start Z.zero)
- && (Z.ltb stop Int.zwordsize)
+ if is_bitfield stop start
then
let stop' := Z.add stop Z.one in
match v with
@@ -17,9 +20,7 @@ Definition extfz stop start v :=
Definition extfs stop start v :=
- if (Z.leb start stop)
- && (Z.geb start Z.zero)
- && (Z.ltb stop Int.zwordsize)
+ if is_bitfield stop start
then
let stop' := Z.add stop Z.one in
match v with
@@ -29,10 +30,13 @@ Definition extfs stop start v :=
end
else Vundef.
+Definition is_bitfieldl stop start :=
+ (Z.leb start stop)
+ && (Z.geb start Z.zero)
+ && (Z.ltb stop Int64.zwordsize).
+
Definition extfzl stop start v :=
- if (Z.leb start stop)
- && (Z.geb start Z.zero)
- && (Z.ltb stop Int64.zwordsize)
+ if is_bitfieldl stop start
then
let stop' := Z.add stop Z.one in
match v with
@@ -44,9 +48,7 @@ Definition extfzl stop start v :=
Definition extfsl stop start v :=
- if (Z.leb start stop)
- && (Z.geb start Z.zero)
- && (Z.ltb stop Int64.zwordsize)
+ if is_bitfieldl stop start
then
let stop' := Z.add stop Z.one in
match v with