aboutsummaryrefslogtreecommitdiffstats
path: root/src/hls/HashTree.v
diff options
context:
space:
mode:
authorYann Herklotz <git@yannherklotz.com>2022-07-14 08:46:14 +0100
committerYann Herklotz <git@yannherklotz.com>2022-07-14 08:46:14 +0100
commitaa753acd776638971abb5d9901cc99ef259cb314 (patch)
tree647722ffa3dae5b10e04cdf46b4e96f27699a26d /src/hls/HashTree.v
parent839ae9a65535e25e52207d46e274385e0709a90f (diff)
downloadvericert-aa753acd776638971abb5d9901cc99ef259cb314.tar.gz
vericert-aa753acd776638971abb5d9901cc99ef259cb314.zip
Add work on abstract predicates
Diffstat (limited to 'src/hls/HashTree.v')
-rw-r--r--src/hls/HashTree.v11
1 files changed, 4 insertions, 7 deletions
diff --git a/src/hls/HashTree.v b/src/hls/HashTree.v
index f3c57a8..9f9ec81 100644
--- a/src/hls/HashTree.v
+++ b/src/hls/HashTree.v
@@ -1,6 +1,6 @@
(*
* Vericert: Verified high-level synthesis.
- * Copyright (C) 2021 Yann Herklotz <yann@yannherklotz.com>
+ * Copyright (C) 2021-2022 Yann Herklotz <yann@yannherklotz.com>
*
* This program is free software: you can redistribute it and/or modify
* it under the terms of the GNU General Public License as published by
@@ -16,6 +16,8 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
+Require Import Coq.Structures.Equalities.
+
Require Import compcert.lib.Maps.
Require Import vericert.common.Vericertlib.
@@ -136,12 +138,7 @@ Proof.
eapply IHl. eauto.
Qed.
-Module Type Hashable.
-
- Parameter t: Type.
- Parameter eq_dec: forall (t1 t2: t), {t1 = t2} + {t1 <> t2}.
-
-End Hashable.
+Module Type Hashable := UsualDecidableType.
Module HashTree(H: Hashable).