aboutsummaryrefslogtreecommitdiffstats
path: root/x86
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-03-22 08:12:41 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-03-22 08:12:41 +0100
commit5a425ba34f3f2520de752ea95a4636a9437c312a (patch)
tree91a1ef113e01f6e7a1006e8ef4e84e9d1f6bab04 /x86
parentba2b825a42807102a93cd6df1ce90b41d415569c (diff)
downloadcompcert-kvx-5a425ba34f3f2520de752ea95a4636a9437c312a.tar.gz
compcert-kvx-5a425ba34f3f2520de752ea95a4636a9437c312a.zip
ça recompile sur x86
Diffstat (limited to 'x86')
-rw-r--r--x86/SelectLong.vp2
-rw-r--r--x86/SelectLongproof.v1
-rw-r--r--x86/SelectOp.vp8
-rw-r--r--x86/SelectOpproof.v29
4 files changed, 37 insertions, 3 deletions
diff --git a/x86/SelectLong.vp b/x86/SelectLong.vp
index b213e23f..4f9fb518 100644
--- a/x86/SelectLong.vp
+++ b/x86/SelectLong.vp
@@ -16,7 +16,7 @@ Require Import Coqlib.
Require Import Compopts.
Require Import AST Integers Floats.
Require Import Op CminorSel.
-Require Import SelectOp SplitLong.
+Require Import OpHelpers SelectOp SplitLong.
Local Open Scope cminorsel_scope.
Local Open Scope string_scope.
diff --git a/x86/SelectLongproof.v b/x86/SelectLongproof.v
index 3bef632d..f008f39e 100644
--- a/x86/SelectLongproof.v
+++ b/x86/SelectLongproof.v
@@ -16,6 +16,7 @@ Require Import String Coqlib Maps Integers Floats Errors.
Require Archi.
Require Import AST Values Memory Globalenvs Events.
Require Import Cminor Op CminorSel.
+Require Import OpHelpers OpHelpersproof.
Require Import SelectOp SelectOpproof SplitLong SplitLongproof.
Require Import SelectLong.
diff --git a/x86/SelectOp.vp b/x86/SelectOp.vp
index a1583600..eadda093 100644
--- a/x86/SelectOp.vp
+++ b/x86/SelectOp.vp
@@ -40,6 +40,7 @@ Require Import Coqlib.
Require Import Compopts.
Require Import AST Integers Floats.
Require Import Op CminorSel.
+Require Import OpHelpers.
Local Open Scope cminorsel_scope.
@@ -540,3 +541,10 @@ Nondetfunction builtin_arg (e: expr) :=
| Eload chunk (Ainstack ofs) Enil => BA_loadstack chunk ofs
| _ => BA e
end.
+
+(* floats *)
+Definition divf_base (e1: expr) (e2: expr) :=
+ Eop Odivf (e1 ::: e2 ::: Enil).
+
+Definition divfs_base (e1: expr) (e2: expr) :=
+ Eop Odivfs (e1 ::: e2 ::: Enil).
diff --git a/x86/SelectOpproof.v b/x86/SelectOpproof.v
index fdbadaa8..1eeb5906 100644
--- a/x86/SelectOpproof.v
+++ b/x86/SelectOpproof.v
@@ -23,6 +23,8 @@ Require Import Cminor.
Require Import Op.
Require Import CminorSel.
Require Import SelectOp.
+Require Import OpHelpers.
+Require Import OpHelpersproof.
Local Open Scope cminorsel_scope.
@@ -74,8 +76,10 @@ Ltac TrivialExists :=
(** * Correctness of the smart constructors *)
Section CMCONSTR.
-
-Variable ge: genv.
+Variable prog: program.
+Variable hf: helper_functions.
+Hypothesis HELPERS: helper_functions_declared prog hf.
+Let ge := Genv.globalenv prog.
Variable sp: val.
Variable e: env.
Variable m: mem.
@@ -984,4 +988,25 @@ Proof.
- constructor; auto.
Qed.
+
+(* floating-point division without HELPERS *)
+Theorem eval_divf_base:
+ forall le a b x y,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ exists v, eval_expr ge sp e m le (divf_base a b) v /\ Val.lessdef (Val.divf x y) v.
+Proof.
+ intros; unfold divf_base.
+ TrivialExists.
+Qed.
+
+Theorem eval_divfs_base:
+ forall le a b x y,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ exists v, eval_expr ge sp e m le (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v.
+Proof.
+ intros; unfold divfs_base.
+ TrivialExists.
+Qed.
End CMCONSTR.