aboutsummaryrefslogtreecommitdiffstats
path: root/backend/ProfilingExploit.v
blob: 2325f582f7e5416d8fdadf949a34525f1cda9bd5 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
(* *************************************************************)
(*                                                             *)
(*             The Compcert verified compiler                  *)
(*                                                             *)
(*           David Monniaux     CNRS, VERIMAG                  *)
(*                                                             *)
(*  Copyright VERIMAG. All rights reserved.                    *)
(*  This file is distributed under the terms of the INRIA      *)
(*  Non-Commercial License Agreement.                          *)
(*                                                             *)
(* *************************************************************)

Require Import Coqlib Maps Errors Integers Floats Lattice Kildall.
Require Import AST Linking.
Require Import Memory Registers Op RTL.

Local Open Scope positive.

Parameter function_id : function -> AST.profiling_id.
Parameter branch_id : AST.profiling_id -> node -> AST.profiling_id.
Parameter condition_oracle : AST.profiling_id -> option bool.

Definition transf_instr (f_id : AST.profiling_id)
           (pc : node) (i : instruction) : instruction :=
  match i with
  | Icond cond args ifso ifnot None =>
    Icond cond args ifso ifnot (condition_oracle (branch_id f_id pc))
  | _ => i
  end.

Definition transf_function (f : function) : function :=
  {| fn_sig := f.(fn_sig);
     fn_params := f.(fn_params);
     fn_stacksize := f.(fn_stacksize);
     fn_code := PTree.map (transf_instr (function_id f)) f.(fn_code);
     fn_entrypoint := f.(fn_entrypoint) |}.

Definition transf_fundef (fd: fundef) : fundef :=
  AST.transf_fundef transf_function fd.

Definition transf_program (p: program) : program :=
  transform_program transf_fundef p.