From 24406a351e9d64c2953b0b9fc7ef0b3d79db9b85 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Sep 2019 18:51:54 +0200 Subject: fix compiling for aarch64 --- aarch64/SelectOpproof.v | 27 +++++++++++++++++++++++++-- 1 file changed, 25 insertions(+), 2 deletions(-) (limited to 'aarch64/SelectOpproof.v') diff --git a/aarch64/SelectOpproof.v b/aarch64/SelectOpproof.v index b78a5ed8..54c6a9fd 100644 --- a/aarch64/SelectOpproof.v +++ b/aarch64/SelectOpproof.v @@ -16,6 +16,7 @@ Require Import Coqlib Zbits. Require Import AST Integers Floats Values Memory Builtins Globalenvs. Require Import Cminor Op CminorSel. Require Import SelectOp. +Require Import OpHelpers OpHelpersproof. Local Open Scope cminorsel_scope. Local Transparent Archi.ptr64. @@ -74,8 +75,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. @@ -1055,6 +1058,26 @@ Proof. - constructor; auto. Qed. +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. + (** Platform-specific known builtins *) Theorem eval_platform_builtin: -- cgit