From 32990af23915c15e5cfd4cfe32c47cafa508f11d Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 8 May 2020 23:34:15 +0100 Subject: Add AssocMap --- src/verilog/AssocMap.v | 60 ++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 60 insertions(+) create mode 100644 src/verilog/AssocMap.v (limited to 'src/verilog/AssocMap.v') diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v new file mode 100644 index 0000000..e550072 --- /dev/null +++ b/src/verilog/AssocMap.v @@ -0,0 +1,60 @@ +(* + * CoqUp: Verified high-level synthesis. + * Copyright (C) 2020 Yann Herklotz + * + * 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 + * the Free Software Foundation, either version 3 of the License, or + * (at your option) any later version. + * + * This program is distributed in the hope that it will be useful, + * but WITHOUT ANY WARRANTY; without even the implied warranty of + * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the + * GNU General Public License for more details. + * + * You should have received a copy of the GNU General Public License + * along with this program. If not, see . + *) + +From coqup Require Import Coquplib Value. +From Coq Require Import FSets.FMapPositive. + +Definition reg := positive. + +Module AssocMap := PositiveMap. + +Module AssocMapExt. + Import AssocMap. + + Section Operations. + + Variable elt : Type. + + Definition merge : t elt -> t elt -> t elt := fold (@add elt). + + Definition find_default (a : elt) (k : reg) (m : t elt) : elt := + match find k m with + | None => a + | Some v => v + end. + + End Operations. + +End AssocMapExt. +Import AssocMapExt. + +Definition assocmap := AssocMap.t value. + +Definition find_assocmap (n : nat) : reg -> assocmap -> value := + find_default value (ZToValue n 0). + +Definition empty_assocmap : assocmap := AssocMap.empty value. + +Definition merge_assocmap : assocmap -> assocmap -> assocmap := merge value. + +Module AssocMapNotation. + Notation "a ! b" := (AssocMap.find b a) (at level 1). + Notation "a # ( b , c )" := (find_assocmap c b a) (at level 1). + Notation "a # b" := (find_assocmap 32 b a) (at level 1). + Notation "a ## b" := (List.map (fun c => find_assocmap 32 c a) b) (at level 1). +End AssocMapNotation. -- cgit