From dd8d4ae9c320668ac5fd70f72ea76b768edf8165 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 26 Mar 2022 15:48:47 +0000 Subject: Remove literal files again --- src/hls/RTLBlockgen.v | 15 ++++++++------- 1 file changed, 8 insertions(+), 7 deletions(-) (limited to 'src/hls/RTLBlockgen.v') diff --git a/src/hls/RTLBlockgen.v b/src/hls/RTLBlockgen.v index beca0ea..5d7376d 100644 --- a/src/hls/RTLBlockgen.v +++ b/src/hls/RTLBlockgen.v @@ -16,7 +16,6 @@ * along with this program. If not, see . *) -(* [[file:../../docs/basic-block-generation.org::rtlblockgen-imports][rtlblockgen-imports]] *) Require compcert.backend.RTL. Require Import compcert.common.AST. Require Import compcert.lib.Maps. @@ -28,7 +27,6 @@ Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLBlock. #[local] Open Scope positive. -(* rtlblockgen-imports ends here *) Lemma comparison_eq: forall (x y : comparison), {x = y} + {x <> y}. Proof. @@ -148,15 +146,19 @@ Defined. Definition ceq {A: Type} (eqd: forall a b: A, {a = b} + {a <> b}) (a b: A): bool := if eqd a b then true else false. -(* [[file:../../docs/basic-block-generation.org::rtlblockgen-main][rtlblockgen-main]] *) Parameter partition : RTL.function -> Errors.res function. -(** [find_block max nodes index]: Does not need to be sorted, because we use filter and the max fold - function to find the desired element. *) +(*| +``find_block max nodes index``: Does not need to be sorted, because we use +filter and the max fold function to find the desired element. + + .. coq:: +|*) + Definition find_block (max: positive) (nodes: list positive) (index: positive) : positive := List.fold_right Pos.min max (List.filter (fun x => (index <=? x)) nodes). -(*Compute find_block (2::94::28::40::19::nil) 40.*) +(*Compute find_block 100 (2::94::48::39::19::nil) 40.*) Definition check_instr (n: positive) (istr: RTL.instruction) (istr': instr) := match istr, istr' with @@ -287,4 +289,3 @@ Definition transl_fundef := transf_partial_fundef transl_function. Definition transl_program : RTL.program -> Errors.res program := transform_partial_program transl_fundef. -(* rtlblockgen-main ends here *) -- cgit