From 2321c5e73833e7cac019d6a4198489e37129c15e Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 31 Jul 2022 22:33:50 +0100 Subject: Add current changes --- src/hls/Abstr.v | 115 ++------------------------------------------------ src/hls/GiblePargen.v | 2 +- 2 files changed, 5 insertions(+), 112 deletions(-) diff --git a/src/hls/Abstr.v b/src/hls/Abstr.v index bfe49f3..8ca4e46 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. @@ -125,117 +129,6 @@ Definition apred_op := @Predicate.pred_op apred. 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 := *) -- cgit