aboutsummaryrefslogtreecommitdiffstats
path: root/backend/ForwardMoves.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-01-08 08:50:46 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-01-08 08:50:46 +0100
commitb24140ca3257e9d77df8dcea22172a4b06679243 (patch)
treeaf8c1c058331478d64c464a4e3dd30142ffe2df6 /backend/ForwardMoves.v
parent22c73e96f9dfb4120541ee5e11334d0ba2d38fe8 (diff)
downloadcompcert-kvx-b24140ca3257e9d77df8dcea22172a4b06679243.tar.gz
compcert-kvx-b24140ca3257e9d77df8dcea22172a4b06679243.zip
continue implementing semilattice
Diffstat (limited to 'backend/ForwardMoves.v')
-rw-r--r--backend/ForwardMoves.v72
1 files changed, 65 insertions, 7 deletions
diff --git a/backend/ForwardMoves.v b/backend/ForwardMoves.v
index 420e71f9..2a135768 100644
--- a/backend/ForwardMoves.v
+++ b/backend/ForwardMoves.v
@@ -45,13 +45,71 @@ Proof.
destruct (Pos.eq_dec R1x R2x) in *; congruence.
Qed.
+Definition ge (r1 r2 : t) :=
+ forall x,
+ match PTree.get x r1 with
+ | None => True
+ | Some v => (PTree.get x r2) = Some v
+ end.
+
+Lemma ge_refl: forall r1 r2, eq r1 r2 -> ge r1 r2.
+Proof.
+ unfold eq, ge.
+ intros r1 r2 EQ x.
+ pose proof (EQ x) as EQx.
+ clear EQ.
+ destruct (r1 ! x).
+ - congruence.
+ - trivial.
+Qed.
+
+Lemma ge_trans: forall x y z, ge x y -> ge y z -> ge x z.
+Proof.
+ unfold ge.
+ intros r1 r2 r3 GE12 GE23 x.
+ pose proof (GE12 x) as GE12x; clear GE12.
+ pose proof (GE23 x) as GE23x; clear GE23.
+ destruct (r1 ! x); trivial.
+ destruct (r2 ! x); congruence.
+Qed.
+
+Definition lub (r1 r2 : t) :=
+ PTree.combine
+ (fun ov1 ov2 =>
+ match ov1, ov2 with
+ | (Some v1), (Some v2) =>
+ if Pos.eq_dec v1 v2
+ then ov1
+ else None
+ | None, _
+ | _, None => None
+ end)
+ r1 r2.
+
+Lemma ge_lub_left: forall x y, ge (lub x y) x.
+Proof.
+ unfold ge, lub.
+ intros r1 r2 x.
+ rewrite PTree.gcombine by reflexivity.
+ destruct (_ ! _); trivial.
+ destruct (_ ! _); trivial.
+ destruct (Pos.eq_dec _ _); trivial.
+Qed.
+
+Lemma ge_lub_right: forall x y, ge (lub x y) y.
+Proof.
+ unfold ge, lub.
+ intros r1 r2 x.
+ rewrite PTree.gcombine by reflexivity.
+ destruct (_ ! _); trivial.
+ destruct (_ ! _); trivial.
+ destruct (Pos.eq_dec _ _); trivial.
+ congruence.
+Qed.
+
+End RELATION.
+
(*
- Parameter ge: t -> t -> Prop.
- Axiom ge_refl: forall x y, eq x y -> ge x y.
- Axiom ge_trans: forall x y z, ge x y -> ge y z -> ge x z.
Parameter bot: t.
Axiom ge_bot: forall x, ge x bot.
- Parameter lub: t -> t -> t.
- Axiom ge_lub_left: forall x y, ge (lub x y) x.
- Axiom ge_lub_right: forall x y, ge (lub x y) y.
-*) \ No newline at end of file
+ *)