aboutsummaryrefslogtreecommitdiffstats
path: root/backend/CSE3analysis.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-05-06 22:33:02 +0200
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-05-06 22:33:02 +0200
commit6171f6a0880acbf0d007a7715cc37984ac25d851 (patch)
tree6d9fdc4e935a6c616623f3dc1529a0945f8d79d2 /backend/CSE3analysis.v
parentf1f535cad98f3db3e586f0f7a2dbc329fc5bff6f (diff)
downloadcompcert-kvx-6171f6a0880acbf0d007a7715cc37984ac25d851.tar.gz
compcert-kvx-6171f6a0880acbf0d007a7715cc37984ac25d851.zip
-fcse3-glb
Diffstat (limited to 'backend/CSE3analysis.v')
-rw-r--r--backend/CSE3analysis.v24
1 files changed, 17 insertions, 7 deletions
diff --git a/backend/CSE3analysis.v b/backend/CSE3analysis.v
index bd507dec..ade79c28 100644
--- a/backend/CSE3analysis.v
+++ b/backend/CSE3analysis.v
@@ -78,7 +78,10 @@ Module RELATION <: SEMILATTICE_WITHOUT_BOTTOM.
Definition lub x y :=
if Compopts.optim_CSE3_across_merges tt
then PSet.inter x y
- else PSet.empty.
+ else
+ if PSet.eq x y
+ then x
+ else PSet.empty.
Definition glb := PSet.union.
@@ -94,8 +97,10 @@ Module RELATION <: SEMILATTICE_WITHOUT_BOTTOM.
intuition.
- apply PSet.is_subset_spec.
intro.
- rewrite PSet.gempty.
- discriminate.
+ destruct (PSet.eq x y).
+ + auto.
+ + rewrite PSet.gempty.
+ discriminate.
Qed.
Lemma ge_lub_right: forall x y, ge (lub x y) y.
@@ -110,8 +115,10 @@ Module RELATION <: SEMILATTICE_WITHOUT_BOTTOM.
intuition.
- apply PSet.is_subset_spec.
intro.
- rewrite PSet.gempty.
- discriminate.
+ destruct (PSet.eq x y).
+ + subst. auto.
+ + rewrite PSet.gempty.
+ discriminate.
Qed.
Definition top := PSet.empty.
@@ -310,8 +317,11 @@ Section OPERATIONS.
else
let args' := forward_move_l rel args in
match rhs_find op args' rel with
- | Some r => (* FIXME RELATION.glb ( *) move r dst rel (* )
- (oper1 dst op args' rel) *)
+ | Some r =>
+ if Compopts.optim_CSE3_glb tt
+ then RELATION.glb (move r dst rel)
+ (oper1 dst op args' rel)
+ else oper1 dst op args' rel
| None => oper1 dst op args' rel
end.