aboutsummaryrefslogtreecommitdiffstats
path: root/backend/CSE3.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-03-12 16:46:06 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-03-12 16:46:06 +0100
commitc6ebd73465ef895c6ea5e240f9c784463a6a0fe5 (patch)
tree15dec58036a4ddf49954cd25565b50e17bc1b292 /backend/CSE3.v
parent9efb4d87c549087e2ef16103a6993a1f99328348 (diff)
downloadcompcert-kvx-c6ebd73465ef895c6ea5e240f9c784463a6a0fe5.tar.gz
compcert-kvx-c6ebd73465ef895c6ea5e240f9c784463a6a0fe5.zip
removed second analysis phase
Diffstat (limited to 'backend/CSE3.v')
-rw-r--r--backend/CSE3.v33
1 files changed, 10 insertions, 23 deletions
diff --git a/backend/CSE3.v b/backend/CSE3.v
index 2ef16376..e82b7cdb 100644
--- a/backend/CSE3.v
+++ b/backend/CSE3.v
@@ -6,31 +6,21 @@ Require Import RTLtyping.
Local Open Scope error_monad_scope.
-Axiom preanalysis : typing_env -> RTL.function -> analysis_hints.
-
-Definition run f := preanalysis f.
+Axiom preanalysis : typing_env -> RTL.function -> invariants * analysis_hints.
Section REWRITE.
Context {ctx : eq_context}.
Definition find_op_in_fmap fmap pc op args :=
- match fmap with
+ match PMap.get pc fmap with
+ | Some rel => rhs_find (ctx:=ctx) pc (SOp op) args rel
| None => None
- | Some map =>
- match PMap.get pc map with
- | Some rel => rhs_find (ctx:=ctx) pc (SOp op) args rel
- | None => None
- end
end.
Definition find_load_in_fmap fmap pc chunk addr args :=
- match fmap with
+ match PMap.get pc fmap with
+ | Some rel => rhs_find (ctx:=ctx) pc (SLoad chunk addr) args rel
| None => None
- | Some map =>
- match PMap.get pc map with
- | Some rel => rhs_find (ctx:=ctx) pc (SLoad chunk addr) args rel
- | None => None
- end
end.
Definition forward_move_b (rb : RB.t) (x : reg) :=
@@ -39,15 +29,12 @@ Definition forward_move_b (rb : RB.t) (x : reg) :=
| Some rel => forward_move (ctx := ctx) rel x
end.
-Definition subst_arg (fmap : option (PMap.t RB.t)) (pc : node) (x : reg) : reg :=
- match fmap with
- | None => x
- | Some inv => forward_move_b (PMap.get pc inv) x
- end.
+Definition subst_arg (fmap : PMap.t RB.t) (pc : node) (x : reg) : reg :=
+ forward_move_b (PMap.get pc fmap) x.
Definition subst_args fmap pc := List.map (subst_arg fmap pc).
-Definition transf_instr (fmap : option (PMap.t RB.t))
+Definition transf_instr (fmap : PMap.t RB.t)
(pc: node) (instr: instruction) :=
match instr with
| Iop op args dst s =>
@@ -80,8 +67,8 @@ End REWRITE.
Definition transf_function (f: function) : res function :=
do tenv <- type_function f;
- let ctx := context_from_hints (preanalysis tenv f) in
- let invariants := internal_analysis (ctx := ctx) tenv f in
+ let (invariants, hints) := preanalysis tenv f in
+ let ctx := context_from_hints hints in
OK {| fn_sig := f.(fn_sig);
fn_params := f.(fn_params);
fn_stacksize := f.(fn_stacksize);