From 178a7c4781c96857fe0a33c777da83e769516152 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 11 Jul 2021 15:59:21 +0200 Subject: Remove unnecessary files and proofs --- src/hls/RTLBlockInstr.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) (limited to 'src/hls/RTLBlockInstr.v') diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 86f8eba..5e123a3 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -211,9 +211,9 @@ Fixpoint trans_pred (bound: nat) (p: pred_op) : - apply orb_true_intro. apply satFormula_mult2 in H. inv H. apply i in H0. auto. apply i0 in H0. auto. -Admitted. +Abort. -Definition sat_pred (bound: nat) (p: pred_op) : +(*Definition sat_pred (bound: nat) (p: pred_op) : option ({al : alist | sat_predicate p (interp_alist al) = true} + {forall a : asgn, sat_predicate p a = false}). refine @@ -243,7 +243,7 @@ Definition sat_pred_temp (bound: nat) (p: pred_op) := match trans_pred_temp bound p with | Some fm => boundedSatSimple bound fm | None => None - end. + end.*) Inductive instr : Type := | RBnop : instr -- cgit From 4f074002b6c2b626a3f41528e9b3bdf62b82e2bc Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 12 Aug 2021 02:03:55 +0200 Subject: Fix proofs for SAT --- src/hls/RTLBlockInstr.v | 33 +++++++++++++++++++++++++++------ 1 file changed, 27 insertions(+), 6 deletions(-) (limited to 'src/hls/RTLBlockInstr.v') diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 5e123a3..5d9d578 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -194,12 +194,33 @@ Fixpoint trans_pred (bound: nat) (p: pred_op) : | _, _ => None end | Pnot (Pvar p') => Some (exist _ (((false, p') :: nil) :: nil) _) - | _ => None + | Pnot (Pnot p') => + match trans_pred n p' with + | Some (exist p1' _) => Some (exist _ p1' _) + | None => None + end + | Pnot (Pand p1 p2) => + match trans_pred n (Por (Pnot p1) (Pnot p2)) with + | Some (exist p1' _) => Some (exist _ p1' _) + | None => None + end + | Pnot (Por p1 p2) => + match trans_pred n (Pand (Pnot p1) (Pnot p2)) with + | Some (exist p1' _) => Some (exist _ p1' _) + | None => None + end end end); split; intros; simpl in *; auto. - inv H. inv H0; auto. - - admit. - - admit. + - split; auto. destruct (a p') eqn:?; crush. + - inv H. inv H0. unfold satLit in H. simplify. rewrite H. auto. + crush. + - rewrite negb_involutive in H. apply i in H. auto. + - rewrite negb_involutive. apply i; auto. + - rewrite negb_andb in H. apply i. auto. + - rewrite negb_andb. apply i. auto. + - rewrite negb_orb in H. apply i. auto. + - rewrite negb_orb. apply i. auto. - apply satFormula_concat. apply andb_prop in H. inv H. apply i in H0. auto. apply andb_prop in H. inv H. apply i0 in H1. auto. @@ -211,9 +232,9 @@ Fixpoint trans_pred (bound: nat) (p: pred_op) : - apply orb_true_intro. apply satFormula_mult2 in H. inv H. apply i in H0. auto. apply i0 in H0. auto. -Abort. +Qed. -(*Definition sat_pred (bound: nat) (p: pred_op) : +Definition sat_pred (bound: nat) (p: pred_op) : option ({al : alist | sat_predicate p (interp_alist al) = true} + {forall a : asgn, sat_predicate p a = false}). refine @@ -243,7 +264,7 @@ Definition sat_pred_temp (bound: nat) (p: pred_op) := match trans_pred_temp bound p with | Some fm => boundedSatSimple bound fm | None => None - end.*) + end. Inductive instr : Type := | RBnop : instr -- cgit From ea64d739af16952883abce536958ac4877698277 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 17 Sep 2021 18:21:19 +0100 Subject: Fix compilation with new CompCert version --- src/hls/RTLBlockInstr.v | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) (limited to 'src/hls/RTLBlockInstr.v') diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 5d9d578..3fab464 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -183,30 +183,30 @@ Fixpoint trans_pred (bound: nat) (p: pred_op) : | Pvar p' => Some (exist _ (((true, p') :: nil) :: nil) _) | Pand p1 p2 => match trans_pred n p1, trans_pred n p2 with - | Some (exist p1' _), Some (exist p2' _) => + | Some (exist _ p1' _), Some (exist _ p2' _) => Some (exist _ (p1' ++ p2') _) | _, _ => None end | Por p1 p2 => match trans_pred n p1, trans_pred n p2 with - | Some (exist p1' _), Some (exist p2' _) => + | Some (exist _ p1' _), Some (exist _ p2' _) => Some (exist _ (mult p1' p2') _) | _, _ => None end | Pnot (Pvar p') => Some (exist _ (((false, p') :: nil) :: nil) _) | Pnot (Pnot p') => match trans_pred n p' with - | Some (exist p1' _) => Some (exist _ p1' _) + | Some (exist _ p1' _) => Some (exist _ p1' _) | None => None end | Pnot (Pand p1 p2) => match trans_pred n (Por (Pnot p1) (Pnot p2)) with - | Some (exist p1' _) => Some (exist _ p1' _) + | Some (exist _ p1' _) => Some (exist _ p1' _) | None => None end | Pnot (Por p1 p2) => match trans_pred n (Pand (Pnot p1) (Pnot p2)) with - | Some (exist p1' _) => Some (exist _ p1' _) + | Some (exist _ p1' _) => Some (exist _ p1' _) | None => None end end @@ -239,9 +239,9 @@ Definition sat_pred (bound: nat) (p: pred_op) : + {forall a : asgn, sat_predicate p a = false}). refine ( match trans_pred bound p with - | Some (exist fm _) => + | Some (exist _ fm _) => match boundedSat bound fm with - | Some (inleft (exist a _)) => Some (inleft (exist _ a _)) + | Some (inleft (exist _ a _)) => Some (inleft (exist _ a _)) | Some (inright _) => Some (inright _) | None => None end @@ -255,7 +255,7 @@ Qed. Definition sat_pred_simple (bound: nat) (p: pred_op) := match sat_pred bound p with - | Some (inleft (exist al _)) => Some (Some al) + | Some (inleft (exist _ al _)) => Some (Some al) | Some (inright _) => Some None | None => None end. -- cgit