aboutsummaryrefslogtreecommitdiffstats
path: root/backend/ValueDomain.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@college-de-france.fr>2019-07-03 11:06:07 +0200
committerXavier Leroy <xavierleroy@users.noreply.github.com>2019-07-17 09:17:28 +0200
commite16b34d9fc1aa7854759787fd52ad59c964c2d4b (patch)
treeba63485dfd5706cb9191d8ce4a2cfb71d3604287 /backend/ValueDomain.v
parentd08b080747225160b80c3f04bdfd9cd67550b425 (diff)
downloadcompcert-kvx-e16b34d9fc1aa7854759787fd52ad59c964c2d4b.tar.gz
compcert-kvx-e16b34d9fc1aa7854759787fd52ad59c964c2d4b.zip
Perform constant propagation for known built-in functions
When an external function is a known built-in function and it is applied to compile-time integer or FP constants, we can use the known semantics of the builtin to compute the result at compile-time.
Diffstat (limited to 'backend/ValueDomain.v')
-rw-r--r--backend/ValueDomain.v44
1 files changed, 42 insertions, 2 deletions
diff --git a/backend/ValueDomain.v b/backend/ValueDomain.v
index f6afa836..fd3bd5ae 100644
--- a/backend/ValueDomain.v
+++ b/backend/ValueDomain.v
@@ -13,7 +13,7 @@
Require Import FunInd.
Require Import Zwf Coqlib Maps Zbits Integers Floats Lattice.
Require Import Compopts AST.
-Require Import Values Memory Globalenvs Events.
+Require Import Values Memory Globalenvs Builtins Events.
Require Import Registers RTL.
(** The abstract domains for value analysis *)
@@ -3038,7 +3038,47 @@ Proof with (auto using provenance_monotone with va).
- destruct (zlt n 16)...
Qed.
-(** Abstracting memory blocks *)
+(** Analysis of known builtin functions. All we have is a dynamic semantics
+ as a function [list val -> option val], but we can still perform
+ some constant propagation. *)
+
+Definition val_of_aval (a: aval) : val :=
+ match a with
+ | I n => Vint n
+ | L n => Vlong n
+ | F f => Vfloat f
+ | FS f => Vsingle f
+ | _ => Vundef
+ end.
+
+Definition aval_of_val (v: val) : option aval :=
+ match v with
+ | Vint n => Some (I n)
+ | Vlong n => Some (L n)
+ | Vfloat f => Some (F f)
+ | Vsingle f => Some (FS f)
+ | _ => None
+ end.
+
+Lemma val_of_aval_sound:
+ forall v a, vmatch v a -> Val.lessdef (val_of_aval a) v.
+Proof.
+ destruct 1; simpl; auto.
+Qed.
+
+Corollary list_val_of_aval_sound:
+ forall vl al, list_forall2 vmatch vl al -> Val.lessdef_list (map val_of_aval al) vl.
+Proof.
+ induction 1; simpl; constructor; auto using val_of_aval_sound.
+Qed.
+
+Lemma aval_of_val_sound:
+ forall v a, aval_of_val v = Some a -> vmatch v a.
+Proof.
+ intros v a E; destruct v; simpl in E; inv E; constructor.
+Qed.
+
+(** * Abstracting memory blocks *)
Inductive acontent : Type :=
| ACval (chunk: memory_chunk) (av: aval).