aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorYann Herklotz <git@yannherklotz.com>2022-07-31 22:34:36 +0100
committerYann Herklotz <git@yannherklotz.com>2022-07-31 22:34:36 +0100
commitc0ef6778d1279f9733a57cb024bbc7d4f819eff5 (patch)
treedc8af8892995861de89b60e84ae550ffe6777ab1
parent262854b593d5dbcc78794e87360e892cb4aba466 (diff)
parent2321c5e73833e7cac019d6a4198489e37129c15e (diff)
downloadvericert-c0ef6778d1279f9733a57cb024bbc7d4f819eff5.tar.gz
vericert-c0ef6778d1279f9733a57cb024bbc7d4f819eff5.zip
Merge remote-tracking branch 'origin/dev/scheduling' into dev/scheduling
-rw-r--r--src/common/NonEmpty.v110
-rw-r--r--src/hls/Abstr.v115
-rw-r--r--src/hls/GiblePargen.v2
3 files changed, 115 insertions, 112 deletions
diff --git a/src/common/NonEmpty.v b/src/common/NonEmpty.v
new file mode 100644
index 0000000..9cb7650
--- /dev/null
+++ b/src/common/NonEmpty.v
@@ -0,0 +1,110 @@
+(*
+ * Vericert: Verified high-level synthesis.
+ * 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
+ * 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 <https://www.gnu.org/licenses/>.
+ *)
+
+Require Import Coq.Lists.List.
+
+Inductive non_empty (A: Type) :=
+| singleton : A -> non_empty A
+| cons : A -> non_empty A -> non_empty A
+.
+
+Arguments singleton [A].
+Arguments cons [A].
+
+Declare Scope non_empty_scope.
+Delimit Scope non_empty_scope with non_empty.
+
+Module NonEmptyNotation.
+ Infix "::|" := cons (at level 60, right associativity) : non_empty_scope.
+End NonEmptyNotation.
+Import NonEmptyNotation.
+
+#[local] Open Scope non_empty_scope.
+
+Fixpoint map {A B} (f: A -> B) (l: non_empty A): non_empty B :=
+ match l with
+ | singleton a => singleton (f a)
+ | a ::| b => f a ::| map f b
+ end.
+
+Fixpoint to_list {A} (l: non_empty A): list A :=
+ match l with
+ | singleton a => a::nil
+ | a ::| b => a :: to_list b
+ end.
+
+Fixpoint app {A} (l1 l2: non_empty A) :=
+ match l1 with
+ | singleton a => a ::| l2
+ | a ::| b => a ::| app b l2
+ end.
+
+Fixpoint non_empty_prod {A B} (l: non_empty A) (l': non_empty B) :=
+ match l with
+ | singleton a => map (fun x => (a, x)) l'
+ | a ::| b => app (map (fun x => (a, x)) l') (non_empty_prod b l')
+ end.
+
+Fixpoint of_list {A} (l: list A): option (non_empty A) :=
+ match l with
+ | a::b =>
+ match of_list b with
+ | Some b' => Some (a ::| b')
+ | _ => None
+ end
+ | nil => None
+ end.
+
+Fixpoint replace {A} (f: A -> A -> bool) (a b: A) (l: non_empty A) :=
+ match l with
+ | a' ::| l' => if f a a' then b ::| replace f a b l' else a' ::| replace f a b l'
+ | singleton a' => if f a a' then singleton b else singleton a'
+ end.
+
+Inductive In {A: Type} (x: A) : non_empty A -> Prop :=
+| In_cons : forall a b, x = a \/ In x b -> In x (a ::| b)
+| In_single : In x (singleton x).
+
+Ltac inv X := inversion X; clear X; subst.
+
+Lemma in_dec:
+ forall A (a: A) (l: non_empty A),
+ (forall a b: A, {a = b} + {a <> b}) ->
+ {In a l} + {~ In a l}.
+Proof.
+ induction l; intros.
+ { specialize (X a a0). inv X.
+ left. constructor.
+ right. unfold not. intros. apply H. inv H0. auto. }
+ { pose proof X as X2.
+ specialize (X a a0). inv X.
+ left. constructor; tauto.
+ apply IHl in X2. inv X2.
+ left. constructor; tauto.
+ right. unfold not in *; intros. apply H0. inv H1. now inv H3. }
+Qed.
+
+Fixpoint filter {A: Type} (f: A -> bool) (l: non_empty A) : option (non_empty A) :=
+ match l with
+ | singleton a =>
+ if f a then Some (singleton a) else None
+ | a ::| b =>
+ if f a then
+ match filter f b with Some x => Some (a ::| x) | None => Some (singleton a) end
+ else filter f b
+ end.
diff --git a/src/hls/Abstr.v b/src/hls/Abstr.v
index b718c15..20ba4fa 100644
--- a/src/hls/Abstr.v
+++ b/src/hls/Abstr.v
@@ -36,7 +36,11 @@ Require Import vericert.hls.Gible.
Require Import vericert.hls.HashTree.
Require Import vericert.hls.Predicate.
Require Import vericert.common.DecEq.
+Require vericert.common.NonEmpty.
+Module NE := NonEmpty.
+Import NE.NonEmptyNotation.
+#[local] Open Scope non_empty_scope.
#[local] Open Scope positive.
#[local] Open Scope pred_op.
@@ -140,117 +144,6 @@ Variant exit_expression : Type :=
Definition pred_op := @Predicate.pred_op positive.
Definition predicate := positive.
-(* Declare Scope apred_op. *)
-
-(* Notation "A ∧ B" := (Pand A B) (at level 20) : apred_op. *)
-(* Notation "A ∨ B" := (Por A B) (at level 25) : apred_op. *)
-(* Notation "⟂" := (Pfalse) : apred_op. *)
-(* Notation "'T'" := (Ptrue) : apred_op. *)
-
-(* #[local] Open Scope apred_op. *)
-
-(*Inductive pred_expr : Type :=
-| PEsingleton : option pred_op -> expression -> pred_expr
-| PEcons : option pred_op -> expression -> pred_expr -> pred_expr.*)
-
-Module NonEmpty.
-
-Inductive non_empty (A: Type) :=
-| singleton : A -> non_empty A
-| cons : A -> non_empty A -> non_empty A
-.
-
-Arguments singleton [A].
-Arguments cons [A].
-
-Declare Scope non_empty_scope.
-Delimit Scope non_empty_scope with non_empty.
-
-Module NonEmptyNotation.
-Infix "::|" := cons (at level 60, right associativity) : non_empty_scope.
-End NonEmptyNotation.
-Import NonEmptyNotation.
-
-#[local] Open Scope non_empty_scope.
-
-Fixpoint map {A B} (f: A -> B) (l: non_empty A): non_empty B :=
- match l with
- | singleton a => singleton (f a)
- | a ::| b => f a ::| map f b
- end.
-
-Fixpoint to_list {A} (l: non_empty A): list A :=
- match l with
- | singleton a => a::nil
- | a ::| b => a :: to_list b
- end.
-
-Fixpoint app {A} (l1 l2: non_empty A) :=
- match l1 with
- | singleton a => a ::| l2
- | a ::| b => a ::| app b l2
- end.
-
-Fixpoint non_empty_prod {A B} (l: non_empty A) (l': non_empty B) :=
- match l with
- | singleton a => map (fun x => (a, x)) l'
- | a ::| b => app (map (fun x => (a, x)) l') (non_empty_prod b l')
- end.
-
-Fixpoint of_list {A} (l: list A): option (non_empty A) :=
- match l with
- | a::b =>
- match of_list b with
- | Some b' => Some (a ::| b')
- | _ => None
- end
- | nil => None
- end.
-
-Fixpoint replace {A} (f: A -> A -> bool) (a b: A) (l: non_empty A) :=
- match l with
- | a' ::| l' => if f a a' then b ::| replace f a b l' else a' ::| replace f a b l'
- | singleton a' => if f a a' then singleton b else singleton a'
- end.
-
-Inductive In {A: Type} (x: A) : non_empty A -> Prop :=
-| In_cons : forall a b, x = a \/ In x b -> In x (a ::| b)
-| In_single : In x (singleton x).
-
-Lemma in_dec:
- forall A (a: A) (l: non_empty A),
- (forall a b: A, {a = b} + {a <> b}) ->
- {In a l} + {~ In a l}.
-Proof.
- induction l; intros.
- { specialize (X a a0). inv X.
- left. constructor.
- right. unfold not. intros. apply H. inv H0. auto. }
- { pose proof X as X2.
- specialize (X a a0). inv X.
- left. constructor; tauto.
- apply IHl in X2. inv X2.
- left. constructor; tauto.
- right. unfold not in *; intros. apply H0. inv H1. now inv H3. }
-Qed.
-
-Fixpoint filter {A: Type} (f: A -> bool) (l: non_empty A) : option (non_empty A) :=
- match l with
- | singleton a =>
- if f a then Some (singleton a) else None
- | a ::| b =>
- if f a then
- match filter f b with Some x => Some (a ::| x) | None => Some (singleton a) end
- else filter f b
- end.
-
-End NonEmpty.
-
-Module NE := NonEmpty.
-Import NE.NonEmptyNotation.
-
-#[local] Open Scope non_empty_scope.
-
Definition apredicated A := NE.non_empty (apred_op * A).
Definition predicated A := NE.non_empty (pred_op * A).
diff --git a/src/hls/GiblePargen.v b/src/hls/GiblePargen.v
index dae1efc..b7f4446 100644
--- a/src/hls/GiblePargen.v
+++ b/src/hls/GiblePargen.v
@@ -184,7 +184,7 @@ Fixpoint apredicated_to_apred_op (b: bool) (a: apredicated expression): apred_op
match a with
| NE.singleton (p, e) => Pimplies p (Plit (b, e))
| (p, e) ::| r =>
- Pand (Pimplies p (Plit (true, e))) (apredicated_to_apred_op b r)
+ Pand (Pimplies p (Plit (b, e))) (apredicated_to_apred_op b r)
end.
(* Fixpoint get_pred' (f: forest) (ap: pred_op): option apred_op := *)