aboutsummaryrefslogtreecommitdiffstats
path: root/lib/Lattice.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-03-05 16:17:52 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-03-05 16:17:52 +0100
commit84a198eddcc2a8cb825e144c9e97642aa45fed7c (patch)
tree89b0a9b4dba8e075a99f1746b662c1bc0e5cd776 /lib/Lattice.v
parent668912983cd68f5f233bfd3af280f911a8522a84 (diff)
downloadcompcert-kvx-84a198eddcc2a8cb825e144c9e97642aa45fed7c.tar.gz
compcert-kvx-84a198eddcc2a8cb825e144c9e97642aa45fed7c.zip
move lattice stuff where it belongs
Diffstat (limited to 'lib/Lattice.v')
-rw-r--r--lib/Lattice.v115
1 files changed, 115 insertions, 0 deletions
diff --git a/lib/Lattice.v b/lib/Lattice.v
index 85fc03f3..2b81f7af 100644
--- a/lib/Lattice.v
+++ b/lib/Lattice.v
@@ -61,6 +61,121 @@ Module Type SEMILATTICE_WITH_TOP.
End SEMILATTICE_WITH_TOP.
+
+Module Type SEMILATTICE_WITHOUT_BOTTOM.
+
+ Parameter t: Type.
+ Parameter eq: t -> t -> Prop.
+ Axiom eq_refl: forall x, eq x x.
+ Axiom eq_sym: forall x y, eq x y -> eq y x.
+ Axiom eq_trans: forall x y z, eq x y -> eq y z -> eq x z.
+ Parameter beq: t -> t -> bool.
+ Axiom beq_correct: forall x y, beq x y = true -> eq x y.
+ 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 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.
+
+End SEMILATTICE_WITHOUT_BOTTOM.
+
+Module ADD_BOTTOM(L : SEMILATTICE_WITHOUT_BOTTOM).
+ Definition t := option L.t.
+ Definition eq (a b : t) :=
+ match a, b with
+ | None, None => True
+ | Some x, Some y => L.eq x y
+ | Some _, None | None, Some _ => False
+ end.
+
+ Lemma eq_refl: forall x, eq x x.
+ Proof.
+ unfold eq; destruct x; trivial.
+ apply L.eq_refl.
+ Qed.
+
+ Lemma eq_sym: forall x y, eq x y -> eq y x.
+ Proof.
+ unfold eq; destruct x; destruct y; trivial.
+ apply L.eq_sym.
+ Qed.
+
+ Lemma eq_trans: forall x y z, eq x y -> eq y z -> eq x z.
+ Proof.
+ unfold eq; destruct x; destruct y; destruct z; trivial.
+ - apply L.eq_trans.
+ - contradiction.
+ Qed.
+
+ Definition beq (x y : t) :=
+ match x, y with
+ | None, None => true
+ | Some x, Some y => L.beq x y
+ | Some _, None | None, Some _ => false
+ end.
+
+ Lemma beq_correct: forall x y, beq x y = true -> eq x y.
+ Proof.
+ unfold beq, eq.
+ destruct x; destruct y; trivial; try congruence.
+ apply L.beq_correct.
+ Qed.
+
+ Definition ge (x y : t) :=
+ match x, y with
+ | None, Some _ => False
+ | _, None => True
+ | Some a, Some b => L.ge a b
+ end.
+
+ Lemma ge_refl: forall x y, eq x y -> ge x y.
+ Proof.
+ unfold eq, ge.
+ destruct x; destruct y; trivial.
+ apply L.ge_refl.
+ Qed.
+
+ Lemma ge_trans: forall x y z, ge x y -> ge y z -> ge x z.
+ Proof.
+ unfold ge.
+ destruct x; destruct y; destruct z; trivial; try contradiction.
+ apply L.ge_trans.
+ Qed.
+
+ Definition bot: t := None.
+ Lemma ge_bot: forall x, ge x bot.
+ Proof.
+ unfold ge, bot.
+ destruct x; trivial.
+ Qed.
+
+ Definition lub (a b : t) :=
+ match a, b with
+ | None, _ => b
+ | _, None => a
+ | (Some x), (Some y) => Some (L.lub x y)
+ end.
+
+ Lemma ge_lub_left: forall x y, ge (lub x y) x.
+ Proof.
+ unfold ge, lub.
+ destruct x; destruct y; trivial.
+ - apply L.ge_lub_left.
+ - apply L.ge_refl.
+ apply L.eq_refl.
+ Qed.
+
+ Lemma ge_lub_right: forall x y, ge (lub x y) y.
+ Proof.
+ unfold ge, lub.
+ destruct x; destruct y; trivial.
+ - apply L.ge_lub_right.
+ - apply L.ge_refl.
+ apply L.eq_refl.
+ Qed.
+End ADD_BOTTOM.
+
(** * Semi-lattice over maps *)
Set Implicit Arguments.