aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
-rw-r--r--lit/basic-block-generation.org378
-rw-r--r--lit/scheduler-languages.org689
-rw-r--r--lit/scheduler.org (renamed from lit/scheduling.org)1014
3 files changed, 1070 insertions, 1011 deletions
diff --git a/lit/basic-block-generation.org b/lit/basic-block-generation.org
new file mode 100644
index 0000000..4a3edc4
--- /dev/null
+++ b/lit/basic-block-generation.org
@@ -0,0 +1,378 @@
+#+title: Basic Block Generation
+#+author: Yann Herklotz
+#+email: yann [at] yannherklotz [dot] com
+
+* RTLBlockgen
+:PROPERTIES:
+:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgen.v
+:END:
+
+Refers to [[rtlblockgen-equalities][rtlblockgen-equalities]].
+
+#+begin_src coq :comments no :padline no :exports none
+<<license>>
+#+end_src
+
+#+name: rtlblockgen-imports
+#+begin_src coq
+Require compcert.backend.RTL.
+Require Import compcert.common.AST.
+Require Import compcert.lib.Maps.
+Require Import compcert.lib.Integers.
+Require Import compcert.lib.Floats.
+
+Require Import vericert.common.Vericertlib.
+Require Import vericert.hls.RTLBlockInstr.
+Require Import vericert.hls.RTLBlock.
+
+#[local] Open Scope positive.
+#+end_src
+
+#+name: rtlblockgen-equalities-insert
+#+begin_src coq
+<<rtlblockgen-equalities>>
+#+end_src
+
+#+name: rtlblockgen-main
+#+begin_src coq
+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. *)
+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.*)
+
+Definition check_instr (n: positive) (istr: RTL.instruction) (istr': instr) :=
+ match istr, istr' with
+ | RTL.Inop n', RBnop => (n' + 1 =? n)
+ | RTL.Iop op args dst n', RBop None op' args' dst' =>
+ ceq operation_eq op op' &&
+ ceq list_pos_eq args args' &&
+ ceq peq dst dst' && (n' + 1 =? n)
+ | RTL.Iload chunk addr args dst n', RBload None chunk' addr' args' dst' =>
+ ceq memory_chunk_eq chunk chunk' &&
+ ceq addressing_eq addr addr' &&
+ ceq list_pos_eq args args' &&
+ ceq peq dst dst' &&
+ (n' + 1 =? n)
+ | RTL.Istore chunk addr args src n', RBstore None chunk' addr' args' src' =>
+ ceq memory_chunk_eq chunk chunk' &&
+ ceq addressing_eq addr addr' &&
+ ceq list_pos_eq args args' &&
+ ceq peq src src' &&
+ (n' + 1 =? n)
+ | _, _ => false
+ end.
+
+Definition check_cf_instr_body (istr: RTL.instruction) (istr': instr): bool :=
+ match istr, istr' with
+ | RTL.Iop op args dst _, RBop None op' args' dst' =>
+ ceq operation_eq op op' &&
+ ceq list_pos_eq args args' &&
+ ceq peq dst dst'
+ | RTL.Iload chunk addr args dst _, RBload None chunk' addr' args' dst' =>
+ ceq memory_chunk_eq chunk chunk' &&
+ ceq addressing_eq addr addr' &&
+ ceq list_pos_eq args args' &&
+ ceq peq dst dst'
+ | RTL.Istore chunk addr args src _, RBstore None chunk' addr' args' src' =>
+ ceq memory_chunk_eq chunk chunk' &&
+ ceq addressing_eq addr addr' &&
+ ceq list_pos_eq args args' &&
+ ceq peq src src'
+ | RTL.Inop _, RBnop
+ | RTL.Icall _ _ _ _ _, RBnop
+ | RTL.Itailcall _ _ _, RBnop
+ | RTL.Ibuiltin _ _ _ _, RBnop
+ | RTL.Icond _ _ _ _, RBnop
+ | RTL.Ijumptable _ _, RBnop
+ | RTL.Ireturn _, RBnop => true
+ | _, _ => false
+ end.
+
+Definition check_cf_instr (istr: RTL.instruction) (istr': cf_instr) :=
+ match istr, istr' with
+ | RTL.Inop n, RBgoto n' => (n =? n')
+ | RTL.Iop _ _ _ n, RBgoto n' => (n =? n')
+ | RTL.Iload _ _ _ _ n, RBgoto n' => (n =? n')
+ | RTL.Istore _ _ _ _ n, RBgoto n' => (n =? n')
+ | RTL.Icall sig (inl r) args dst n, RBcall sig' (inl r') args' dst' n' =>
+ ceq signature_eq sig sig' &&
+ ceq peq r r' &&
+ ceq list_pos_eq args args' &&
+ ceq peq dst dst' &&
+ (n =? n')
+ | RTL.Icall sig (inr i) args dst n, RBcall sig' (inr i') args' dst' n' =>
+ ceq signature_eq sig sig' &&
+ ceq peq i i' &&
+ ceq list_pos_eq args args' &&
+ ceq peq dst dst' &&
+ (n =? n')
+ | RTL.Itailcall sig (inl r) args, RBtailcall sig' (inl r') args' =>
+ ceq signature_eq sig sig' &&
+ ceq peq r r' &&
+ ceq list_pos_eq args args'
+ | RTL.Itailcall sig (inr r) args, RBtailcall sig' (inr r') args' =>
+ ceq signature_eq sig sig' &&
+ ceq peq r r' &&
+ ceq list_pos_eq args args'
+ | RTL.Icond cond args n1 n2, RBcond cond' args' n1' n2' =>
+ ceq condition_eq cond cond' &&
+ ceq list_pos_eq args args' &&
+ ceq peq n1 n1' && ceq peq n2 n2'
+ | RTL.Ijumptable r ns, RBjumptable r' ns' =>
+ ceq peq r r' && ceq list_pos_eq ns ns'
+ | RTL.Ireturn (Some r), RBreturn (Some r') =>
+ ceq peq r r'
+ | RTL.Ireturn None, RBreturn None => true
+ | _, _ => false
+ end.
+
+Definition is_cf_instr (n: positive) (i: RTL.instruction) :=
+ match i with
+ | RTL.Inop n' => negb (n' + 1 =? n)
+ | RTL.Iop _ _ _ n' => negb (n' + 1 =? n)
+ | RTL.Iload _ _ _ _ n' => negb (n' + 1 =? n)
+ | RTL.Istore _ _ _ _ n' => negb (n' + 1 =? n)
+ | RTL.Icall _ _ _ _ _ => true
+ | RTL.Itailcall _ _ _ => true
+ | RTL.Ibuiltin _ _ _ _ => true
+ | RTL.Icond _ _ _ _ => true
+ | RTL.Ijumptable _ _ => true
+ | RTL.Ireturn _ => true
+ end.
+
+Definition check_present_blocks (c: code) (n: list positive) (max: positive) (i: positive) (istr: RTL.instruction) :=
+ let blockn := find_block max n i in
+ match c ! blockn with
+ | Some istrs =>
+ match List.nth_error istrs.(bb_body) (Pos.to_nat blockn - Pos.to_nat i)%nat with
+ | Some istr' =>
+ if is_cf_instr i istr
+ then check_cf_instr istr istrs.(bb_exit) && check_cf_instr_body istr istr'
+ else check_instr i istr istr'
+ | None => false
+ end
+ | None => false
+ end.
+
+Definition transl_function (f: RTL.function) :=
+ match partition f with
+ | Errors.OK f' =>
+ let blockids := map fst (PTree.elements f'.(fn_code)) in
+ if forall_ptree (check_present_blocks f'.(fn_code) blockids (fold_right Pos.max 1 blockids))
+ f.(RTL.fn_code) then
+ Errors.OK f'
+ else Errors.Error (Errors.msg "check_present_blocks failed")
+ | Errors.Error msg => Errors.Error msg
+ end.
+
+Definition transl_fundef := transf_partial_fundef transl_function.
+
+Definition transl_program : RTL.program -> Errors.res program :=
+ transform_partial_program transl_fundef.
+#+end_src
+
+** Equalities
+
+#+name: rtlblockgen-equalities
+#+begin_src coq :tangle no
+Lemma comparison_eq: forall (x y : comparison), {x = y} + {x <> y}.
+Proof.
+ decide equality.
+Defined.
+
+Lemma condition_eq: forall (x y : Op.condition), {x = y} + {x <> y}.
+Proof.
+ generalize comparison_eq; intro.
+ generalize Int.eq_dec; intro.
+ generalize Int64.eq_dec; intro.
+ decide equality.
+Defined.
+
+Lemma addressing_eq : forall (x y : Op.addressing), {x = y} + {x <> y}.
+Proof.
+ generalize Int.eq_dec; intro.
+ generalize AST.ident_eq; intro.
+ generalize Z.eq_dec; intro.
+ generalize Ptrofs.eq_dec; intro.
+ decide equality.
+Defined.
+
+Lemma typ_eq : forall (x y : AST.typ), {x = y} + {x <> y}.
+Proof.
+ decide equality.
+Defined.
+
+Lemma operation_eq: forall (x y : Op.operation), {x = y} + {x <> y}.
+Proof.
+ generalize Int.eq_dec; intro.
+ generalize Int64.eq_dec; intro.
+ generalize Float.eq_dec; intro.
+ generalize Float32.eq_dec; intro.
+ generalize AST.ident_eq; intro.
+ generalize condition_eq; intro.
+ generalize addressing_eq; intro.
+ generalize typ_eq; intro.
+ decide equality.
+Defined.
+
+Lemma memory_chunk_eq : forall (x y : AST.memory_chunk), {x = y} + {x <> y}.
+Proof.
+ decide equality.
+Defined.
+
+Lemma list_typ_eq: forall (x y : list AST.typ), {x = y} + {x <> y}.
+Proof.
+ generalize typ_eq; intro.
+ decide equality.
+Defined.
+
+Lemma option_typ_eq : forall (x y : option AST.typ), {x = y} + {x <> y}.
+Proof.
+ generalize typ_eq; intro.
+ decide equality.
+Defined.
+
+Lemma signature_eq: forall (x y : AST.signature), {x = y} + {x <> y}.
+Proof.
+ repeat decide equality.
+Defined.
+
+Lemma list_operation_eq : forall (x y : list Op.operation), {x = y} + {x <> y}.
+Proof.
+ generalize operation_eq; intro.
+ decide equality.
+Defined.
+
+Lemma list_pos_eq : forall (x y : list positive), {x = y} + {x <> y}.
+Proof.
+ generalize Pos.eq_dec; intros.
+ decide equality.
+Defined.
+
+Lemma sig_eq : forall (x y : AST.signature), {x = y} + {x <> y}.
+Proof.
+ repeat decide equality.
+Defined.
+
+Lemma instr_eq: forall (x y : instr), {x = y} + {x <> y}.
+Proof.
+ generalize Pos.eq_dec; intro.
+ generalize typ_eq; intro.
+ generalize Int.eq_dec; intro.
+ generalize memory_chunk_eq; intro.
+ generalize addressing_eq; intro.
+ generalize operation_eq; intro.
+ generalize condition_eq; intro.
+ generalize signature_eq; intro.
+ generalize list_operation_eq; intro.
+ generalize list_pos_eq; intro.
+ generalize AST.ident_eq; intro.
+ repeat decide equality.
+Defined.
+
+Lemma cf_instr_eq: forall (x y : cf_instr), {x = y} + {x <> y}.
+Proof.
+ generalize Pos.eq_dec; intro.
+ generalize typ_eq; intro.
+ generalize Int.eq_dec; intro.
+ generalize Int64.eq_dec; intro.
+ generalize Float.eq_dec; intro.
+ generalize Float32.eq_dec; intro.
+ generalize Ptrofs.eq_dec; intro.
+ generalize memory_chunk_eq; intro.
+ generalize addressing_eq; intro.
+ generalize operation_eq; intro.
+ generalize condition_eq; intro.
+ generalize signature_eq; intro.
+ generalize list_operation_eq; intro.
+ generalize list_pos_eq; intro.
+ generalize AST.ident_eq; intro.
+ repeat decide equality.
+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.
+#+end_src
+
+* RTLBlockgenproof
+:PROPERTIES:
+:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgenproof.v
+:END:
+
+#+begin_src coq :comments no :padline no :exports none
+<<license>>
+#+end_src
+
+** Imports
+
+#+name: rtlblockgenproof-imports
+#+begin_src coq
+Require compcert.backend.RTL.
+Require Import compcert.common.AST.
+Require Import compcert.lib.Maps.
+
+Require Import vericert.hls.RTLBlock.
+Require Import vericert.hls.RTLBlockgen.
+#+end_src
+
+** Match states
+
+The ~match_states~ predicate describes which states are equivalent between the two languages, in this
+case ~RTL~ and ~RTLBlock~.
+
+#+name: rtlblockgenproof-match-states
+#+begin_src coq
+Inductive match_states : RTL.state -> RTLBlock.state -> Prop :=
+| match_state :
+ forall stk f tf sp pc rs m
+ (TF: transl_function f = OK tf),
+ match_states (RTL.State stk f sp pc rs m)
+ (RTLBlock.State stk tf sp (find_block max n i) rs m).
+#+end_src
+
+** Correctness
+
+#+name: rtlblockgenproof-correctness
+#+begin_src coq
+Section CORRECTNESS.
+
+ Context (prog : RTL.program).
+ Context (tprog : RTLBlock.program).
+
+ Context (TRANSL : match_prog prog tprog).
+
+ Theorem transf_program_correct:
+ Smallstep.forward_simulation (RTL.semantics prog) (RTLBlock.semantics tprog).
+ Proof.
+ eapply Smallstep.forward_simulation_plus; eauto with htlproof.
+ apply senv_preserved.
+
+End CORRECTNESS.
+#+end_src
+
+* License
+
+#+name: license
+#+begin_src coq :tangle no
+(*
+ * Vericert: Verified high-level synthesis.
+ * Copyright (C) 2020-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/>.
+ *)
+#+end_src
diff --git a/lit/scheduler-languages.org b/lit/scheduler-languages.org
new file mode 100644
index 0000000..cf03b3a
--- /dev/null
+++ b/lit/scheduler-languages.org
@@ -0,0 +1,689 @@
+#+title: Scheduler Languages
+#+author: Yann Herklotz
+#+email: yann [at] yannherklotz [dot] com
+
+* RTLBlockInstr
+:PROPERTIES:
+:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockInstr.v
+:END:
+
+#+begin_src coq :comments no :padline no :exports none
+<<license>>
+#+end_src
+
+#+name: rtlblockinstr-imports
+#+begin_src coq
+Require Import Coq.micromega.Lia.
+
+Require Import compcert.backend.Registers.
+Require Import compcert.common.AST.
+Require Import compcert.common.Events.
+Require Import compcert.common.Globalenvs.
+Require Import compcert.common.Memory.
+Require Import compcert.common.Values.
+Require Import compcert.lib.Integers.
+Require Import compcert.lib.Maps.
+Require Import compcert.verilog.Op.
+
+Require Import Predicate.
+Require Import Vericertlib.
+#+end_src
+
+These instructions are used for ~RTLBlock~ and ~RTLPar~, so that they have consistent instructions,
+which greatly simplifies the proofs, as they will by default have the same instruction syntax and
+semantics. The only changes are therefore at the top-level of the instructions.
+
+** Instruction Definition
+
+First, we define the instructions that can be placed into a basic block, meaning they won't branch.
+The main changes to how instructions are defined in [RTL], is that these instructions don't have a
+next node, as they will be in a basic block, and they also have an optional predicate ([pred_op]).
+
+#+name: rtlblockinstr-instr-def
+#+begin_src coq
+Definition node := positive.
+
+Inductive instr : Type :=
+| RBnop : instr
+| RBop : option pred_op -> operation -> list reg -> reg -> instr
+| RBload : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr
+| RBstore : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr
+| RBsetpred : option pred_op -> condition -> list reg -> predicate -> instr.
+#+end_src
+
+** Control-Flow Instruction Definition
+
+These are the instructions that count as control-flow, and will be placed at the end of the basic
+blocks.
+
+#+name: rtlblockinstr-cf-instr-def
+#+begin_src coq
+Inductive cf_instr : Type :=
+| RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr
+| RBtailcall : signature -> reg + ident -> list reg -> cf_instr
+| RBbuiltin : external_function -> list (builtin_arg reg) ->
+ builtin_res reg -> node -> cf_instr
+| RBcond : condition -> list reg -> node -> node -> cf_instr
+| RBjumptable : reg -> list node -> cf_instr
+| RBreturn : option reg -> cf_instr
+| RBgoto : node -> cf_instr
+| RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr.
+#+end_src
+
+** Helper functions
+
+#+name: rtlblockinstr-helpers
+#+begin_src coq
+Fixpoint successors_instr (i : cf_instr) : list node :=
+ match i with
+ | RBcall sig ros args res s => s :: nil
+ | RBtailcall sig ros args => nil
+ | RBbuiltin ef args res s => s :: nil
+ | RBcond cond args ifso ifnot => ifso :: ifnot :: nil
+ | RBjumptable arg tbl => tbl
+ | RBreturn optarg => nil
+ | RBgoto n => n :: nil
+ | RBpred_cf p c1 c2 => concat (successors_instr c1 :: successors_instr c2 :: nil)
+ end.
+
+Definition max_reg_instr (m: positive) (i: instr) :=
+ match i with
+ | RBnop => m
+ | RBop p op args res =>
+ fold_left Pos.max args (Pos.max res m)
+ | RBload p chunk addr args dst =>
+ fold_left Pos.max args (Pos.max dst m)
+ | RBstore p chunk addr args src =>
+ fold_left Pos.max args (Pos.max src m)
+ | RBsetpred p' c args p =>
+ fold_left Pos.max args m
+ end.
+
+Fixpoint max_reg_cfi (m : positive) (i : cf_instr) :=
+ match i with
+ | RBcall sig (inl r) args res s =>
+ fold_left Pos.max args (Pos.max r (Pos.max res m))
+ | RBcall sig (inr id) args res s =>
+ fold_left Pos.max args (Pos.max res m)
+ | RBtailcall sig (inl r) args =>
+ fold_left Pos.max args (Pos.max r m)
+ | RBtailcall sig (inr id) args =>
+ fold_left Pos.max args m
+ | RBbuiltin ef args res s =>
+ fold_left Pos.max (params_of_builtin_args args)
+ (fold_left Pos.max (params_of_builtin_res res) m)
+ | RBcond cond args ifso ifnot => fold_left Pos.max args m
+ | RBjumptable arg tbl => Pos.max arg m
+ | RBreturn None => m
+ | RBreturn (Some arg) => Pos.max arg m
+ | RBgoto n => m
+ | RBpred_cf p c1 c2 => Pos.max (max_reg_cfi m c1) (max_reg_cfi m c2)
+ end.
+
+Definition regset := Regmap.t val.
+Definition predset := PMap.t bool.
+
+Definition eval_predf (pr: predset) (p: pred_op) :=
+ sat_predicate p (fun x => pr !! (Pos.of_nat x)).
+
+#[global]
+Instance eval_predf_Proper : Proper (eq ==> equiv ==> eq) eval_predf.
+Proof.
+ unfold Proper. simplify. unfold "==>".
+ intros.
+ unfold sat_equiv in *. intros. unfold eval_predf. subst. apply H0.
+Qed.
+
+#[local] Open Scope pred_op.
+
+Lemma eval_predf_Pand :
+ forall ps p p',
+ eval_predf ps (p ∧ p') = eval_predf ps p && eval_predf ps p'.
+Proof. unfold eval_predf; split; simplify; auto with bool. Qed.
+
+Lemma eval_predf_Por :
+ forall ps p p',
+ eval_predf ps (p ∨ p') = eval_predf ps p || eval_predf ps p'.
+Proof. unfold eval_predf; split; simplify; auto with bool. Qed.
+
+Lemma eval_predf_pr_equiv :
+ forall p ps ps',
+ (forall x, ps !! x = ps' !! x) ->
+ eval_predf ps p = eval_predf ps' p.
+Proof.
+ induction p; simplify; auto;
+ try (unfold eval_predf; simplify; repeat (destruct_match; []); inv Heqp0; rewrite <- H; auto);
+ [repeat rewrite eval_predf_Pand|repeat rewrite eval_predf_Por];
+ erewrite IHp1; try eassumption; erewrite IHp2; eauto.
+Qed.
+
+Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset :=
+ match rl, vl with
+ | r1 :: rs, v1 :: vs => Regmap.set r1 v1 (init_regs vs rs)
+ | _, _ => Regmap.init Vundef
+ end.
+#+end_src
+
+*** Instruction State
+
+Definition of the instruction state, which contains the following:
+
+- ~is_rs~ :: This is the current state of the registers.
+- ~is_ps~ :: This is the current state of the predicate registers, which is in a separate namespace
+ and area compared to the standard registers in [is_rs].
+- ~is_mem~ :: The current state of the memory.
+
+#+name: rtlblockinstr-instr-state
+#+begin_src coq
+Record instr_state := mk_instr_state {
+ is_rs: regset;
+ is_ps: predset;
+ is_mem: mem;
+}.
+#+end_src
+
+** Top-Level Type Definitions
+
+#+name: rtlblockinstr-type-def
+#+begin_src coq
+Section DEFINITION.
+
+ Context {bblock_body: Type}.
+
+ Record bblock : Type := mk_bblock {
+ bb_body: bblock_body;
+ bb_exit: cf_instr
+ }.
+
+ Definition code: Type := PTree.t bblock.
+
+ Record function: Type := mkfunction {
+ fn_sig: signature;
+ fn_params: list reg;
+ fn_stacksize: Z;
+ fn_code: code;
+ fn_entrypoint: node
+ }.
+
+ Definition fundef := AST.fundef function.
+
+ Definition program := AST.program fundef unit.
+
+ Definition funsig (fd: fundef) :=
+ match fd with
+ | Internal f => fn_sig f
+ | External ef => ef_sig ef
+ end.
+
+ Inductive stackframe : Type :=
+ | Stackframe:
+ forall (res: reg) (**r where to store the result *)
+ (f: function) (**r calling function *)
+ (sp: val) (**r stack pointer in calling function *)
+ (pc: node) (**r program point in calling function *)
+ (rs: regset) (**r register state in calling function *)
+ (pr: predset), (**r predicate state of the calling function *)
+ stackframe.
+
+ Inductive state : Type :=
+ | State:
+ forall (stack: list stackframe) (**r call stack *)
+ (f: function) (**r current function *)
+ (sp: val) (**r stack pointer *)
+ (pc: node) (**r current program point in [c] *)
+ (rs: regset) (**r register state *)
+ (pr: predset) (**r predicate register state *)
+ (m: mem), (**r memory state *)
+ state
+ | Callstate:
+ forall (stack: list stackframe) (**r call stack *)
+ (f: fundef) (**r function to call *)
+ (args: list val) (**r arguments to the call *)
+ (m: mem), (**r memory state *)
+ state
+ | Returnstate:
+ forall (stack: list stackframe) (**r call stack *)
+ (v: val) (**r return value for the call *)
+ (m: mem), (**r memory state *)
+ state.
+
+End DEFINITION.
+#+end_src
+
+** Semantics
+
+#+name: rtlblockinstr-semantics
+#+begin_src coq
+Section RELSEM.
+
+ Context {bblock_body : Type}.
+
+ Definition genv := Genv.t (@fundef bblock_body) unit.
+
+ Context (ge: genv).
+
+ Definition find_function
+ (ros: reg + ident) (rs: regset) : option fundef :=
+ match ros with
+ | inl r => Genv.find_funct ge rs#r
+ | inr symb =>
+ match Genv.find_symbol ge symb with
+ | None => None
+ | Some b => Genv.find_funct_ptr ge b
+ end
+ end.
+
+ Inductive eval_pred: option pred_op -> instr_state -> instr_state -> instr_state -> Prop :=
+ | eval_pred_true:
+ forall i i' p,
+ eval_predf (is_ps i) p = true ->
+ eval_pred (Some p) i i' i'
+ | eval_pred_false:
+ forall i i' p,
+ eval_predf (is_ps i) p = false ->
+ eval_pred (Some p) i i' i
+ | eval_pred_none:
+ forall i i', eval_pred None i i' i.
+
+ Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop :=
+ | exec_RBnop:
+ forall sp ist,
+ step_instr sp ist RBnop ist
+ | exec_RBop:
+ forall op v res args rs m sp p ist pr,
+ eval_operation ge sp op rs##args m = Some v ->
+ eval_pred p (mk_instr_state rs pr m) (mk_instr_state (rs#res <- v) pr m) ist ->
+ step_instr sp (mk_instr_state rs pr m) (RBop p op args res) ist
+ | exec_RBload:
+ forall addr rs args a chunk m v dst sp p pr ist,
+ eval_addressing ge sp addr rs##args = Some a ->
+ Mem.loadv chunk m a = Some v ->
+ eval_pred p (mk_instr_state rs pr m) (mk_instr_state (rs#dst <- v) pr m) ist ->
+ step_instr sp (mk_instr_state rs pr m) (RBload p chunk addr args dst) ist
+ | exec_RBstore:
+ forall addr rs args a chunk m src m' sp p pr ist,
+ eval_addressing ge sp addr rs##args = Some a ->
+ Mem.storev chunk m a rs#src = Some m' ->
+ eval_pred p (mk_instr_state rs pr m) (mk_instr_state rs pr m') ist ->
+ step_instr sp (mk_instr_state rs pr m) (RBstore p chunk addr args src) ist
+ | exec_RBsetpred:
+ forall sp rs pr m p c b args p' ist,
+ Op.eval_condition c rs##args m = Some b ->
+ eval_pred p' (mk_instr_state rs pr m) (mk_instr_state rs (pr#p <- b) m) ist ->
+ step_instr sp (mk_instr_state rs pr m) (RBsetpred p' c args p) ist.
+
+ Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop :=
+ | exec_RBcall:
+ forall s f sp rs m res fd ros sig args pc pc' pr,
+ find_function ros rs = Some fd ->
+ funsig fd = sig ->
+ step_cf_instr (State s f sp pc rs pr m) (RBcall sig ros args res pc')
+ E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m)
+ | exec_RBtailcall:
+ forall s f stk rs m sig ros args fd m' pc pr,
+ find_function ros rs = Some fd ->
+ funsig fd = sig ->
+ Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
+ step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) (RBtailcall sig ros args)
+ E0 (Callstate s fd rs##args m')
+ | exec_RBbuiltin:
+ forall s f sp rs m ef args res pc' vargs t vres m' pc pr,
+ eval_builtin_args ge (fun r => rs#r) sp m args vargs ->
+ external_call ef ge vargs m t vres m' ->
+ step_cf_instr (State s f sp pc rs pr m) (RBbuiltin ef args res pc')
+ t (State s f sp pc' (regmap_setres res vres rs) pr m')
+ | exec_RBcond:
+ forall s f sp rs m cond args ifso ifnot b pc pc' pr,
+ eval_condition cond rs##args m = Some b ->
+ pc' = (if b then ifso else ifnot) ->
+ step_cf_instr (State s f sp pc rs pr m) (RBcond cond args ifso ifnot)
+ E0 (State s f sp pc' rs pr m)
+ | exec_RBjumptable:
+ forall s f sp rs m arg tbl n pc pc' pr,
+ rs#arg = Vint n ->
+ list_nth_z tbl (Int.unsigned n) = Some pc' ->
+ step_cf_instr (State s f sp pc rs pr m) (RBjumptable arg tbl)
+ E0 (State s f sp pc' rs pr m)
+ | exec_RBreturn:
+ forall s f stk rs m or pc m' pr,
+ Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
+ step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) (RBreturn or)
+ E0 (Returnstate s (regmap_optget or Vundef rs) m')
+ | exec_RBgoto:
+ forall s f sp pc rs pr m pc',
+ step_cf_instr (State s f sp pc rs pr m) (RBgoto pc') E0 (State s f sp pc' rs pr m)
+ | exec_RBpred_cf:
+ forall s f sp pc rs pr m cf1 cf2 st' p t,
+ step_cf_instr (State s f sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' ->
+ step_cf_instr (State s f sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'.
+
+End RELSEM.
+#+end_src
+
+* RTLBlock
+:PROPERTIES:
+:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlock.v
+:END:
+
+#+begin_src coq :comments no :padline no :exports none
+<<license>>
+#+end_src
+
+#+name: rtlblock-main
+#+begin_src coq
+Require Import compcert.backend.Registers.
+Require Import compcert.common.AST.
+Require Import compcert.common.Events.
+Require Import compcert.common.Globalenvs.
+Require Import compcert.common.Memory.
+Require Import compcert.common.Smallstep.
+Require Import compcert.common.Values.
+Require Import compcert.lib.Coqlib.
+Require Import compcert.lib.Integers.
+Require Import compcert.lib.Maps.
+Require Import compcert.verilog.Op.
+
+Require Import vericert.hls.RTLBlockInstr.
+
+Definition bb := list instr.
+
+Definition bblock := @bblock bb.
+Definition code := @code bb.
+Definition function := @function bb.
+Definition fundef := @fundef bb.
+Definition program := @program bb.
+Definition funsig := @funsig bb.
+Definition stackframe := @stackframe bb.
+
+Definition genv := @genv bb.
+
+Inductive state : Type :=
+| State:
+ forall (stack: list stackframe) (**r call stack *)
+ (f: function) (**r current function *)
+ (b: bb) (**r current block being executed *)
+ (sp: val) (**r stack pointer *)
+ (pc: node) (**r current program point in [c] *)
+ (rs: regset) (**r register state *)
+ (pr: predset) (**r predicate register state *)
+ (m: mem), (**r memory state *)
+ state
+| Callstate:
+ forall (stack: list stackframe) (**r call stack *)
+ (f: fundef) (**r function to call *)
+ (args: list val) (**r arguments to the call *)
+ (m: mem), (**r memory state *)
+ state
+| Returnstate:
+ forall (stack: list stackframe) (**r call stack *)
+ (v: val) (**r return value for the call *)
+ (m: mem), (**r memory state *)
+ state.
+
+Section RELSEM.
+
+ Context (ge: genv).
+
+ Inductive step_instr_list: val -> instr_state -> list instr -> instr_state -> Prop :=
+ | exec_RBcons:
+ forall state i state' state'' instrs sp,
+ step_instr ge sp state i state' ->
+ step_instr_list sp state' instrs state'' ->
+ step_instr_list sp state (i :: instrs) state''
+ | exec_RBnil:
+ forall state sp,
+ step_instr_list sp state nil state.
+
+ Definition find_function
+ (ros: reg + ident) (rs: regset) : option fundef :=
+ match ros with
+ | inl r => Genv.find_funct ge rs#r
+ | inr symb =>
+ match Genv.find_symbol ge symb with
+ | None => None
+ | Some b => Genv.find_funct_ptr ge b
+ end
+ end.
+
+ Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop :=
+ | exec_RBcall:
+ forall s f b sp rs m res fd ros sig args pc pc' pr,
+ find_function ros rs = Some fd ->
+ funsig fd = sig ->
+ step_cf_instr (State s f b sp pc rs pr m) (RBcall sig ros args res pc')
+ E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m)
+ | exec_RBtailcall:
+ forall s f b stk rs m sig ros args fd m' pc pr,
+ find_function ros rs = Some fd ->
+ funsig fd = sig ->
+ Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
+ step_cf_instr (State s f b (Vptr stk Ptrofs.zero) pc rs pr m) (RBtailcall sig ros args)
+ E0 (Callstate s fd rs##args m')
+ | exec_RBbuiltin:
+ forall s f b sp rs m ef args res pc' vargs t vres m' pc pr,
+ eval_builtin_args ge (fun r => rs#r) sp m args vargs ->
+ external_call ef ge vargs m t vres m' ->
+ step_cf_instr (State s f b sp pc rs pr m) (RBbuiltin ef args res pc')
+ t (State s f b sp pc' (regmap_setres res vres rs) pr m')
+ | exec_RBcond:
+ forall s f block sp rs m cond args ifso ifnot b pc pc' pr,
+ eval_condition cond rs##args m = Some b ->
+ pc' = (if b then ifso else ifnot) ->
+ step_cf_instr (State s f block sp pc rs pr m) (RBcond cond args ifso ifnot)
+ E0 (State s f block sp pc' rs pr m)
+ | exec_RBjumptable:
+ forall s f b sp rs m arg tbl n pc pc' pr,
+ rs#arg = Vint n ->
+ list_nth_z tbl (Int.unsigned n) = Some pc' ->
+ step_cf_instr (State s f b sp pc rs pr m) (RBjumptable arg tbl)
+ E0 (State s f b sp pc' rs pr m)
+ | exec_RBreturn:
+ forall s f b stk rs m or pc m' pr,
+ Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
+ step_cf_instr (State s f b (Vptr stk Ptrofs.zero) pc rs pr m) (RBreturn or)
+ E0 (Returnstate s (regmap_optget or Vundef rs) m')
+ | exec_RBgoto:
+ forall s f b sp pc rs pr m pc',
+ step_cf_instr (State s f b sp pc rs pr m) (RBgoto pc') E0 (State s f b sp pc' rs pr m)
+ | exec_RBpred_cf:
+ forall s f b sp pc rs pr m cf1 cf2 st' p t,
+ step_cf_instr (State s f b sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' ->
+ step_cf_instr (State s f b sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'.
+
+ Inductive step: state -> trace -> state -> Prop :=
+ | exec_function_internal:
+ forall s f args m m' stk,
+ Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
+ step (Callstate s (Internal f) args m)
+ E0 (State s f
+ (Vptr stk Ptrofs.zero)
+ f.(fn_entrypoint)
+ (init_regs args f.(fn_params))
+ (PMap.init false)
+ m')
+ | exec_function_external:
+ forall s ef args res t m m',
+ external_call ef ge args m t res m' ->
+ step (Callstate s (External ef) args m)
+ t (Returnstate s res m')
+ | exec_return:
+ forall res f sp pc rs s vres m pr,
+ step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m)
+ E0 (State s f sp pc (rs#res <- vres) pr m).
+
+End RELSEM.
+
+Inductive initial_state (p: program): state -> Prop :=
+ | initial_state_intro: forall b f m0,
+ let ge := Genv.globalenv p in
+ Genv.init_mem p = Some m0 ->
+ Genv.find_symbol ge p.(prog_main) = Some b ->
+ Genv.find_funct_ptr ge b = Some f ->
+ funsig f = signature_main ->
+ initial_state p (Callstate nil f nil m0).
+
+Inductive final_state: state -> int -> Prop :=
+ | final_state_intro: forall r m,
+ final_state (Returnstate nil (Vint r) m) r.
+
+Definition semantics (p: program) :=
+ Semantics step (initial_state p) final_state (Genv.globalenv p).
+#+end_src
+
+* RTLPar
+:PROPERTIES:
+:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPar.v
+:END:
+
+#+begin_src coq :comments no :padline no :exports none
+<<license>>
+#+end_src
+
+#+name: rtlpar-main
+#+begin_src coq
+Require Import compcert.backend.Registers.
+Require Import compcert.common.AST.
+Require Import compcert.common.Events.
+Require Import compcert.common.Globalenvs.
+Require Import compcert.common.Memory.
+Require Import compcert.common.Smallstep.
+Require Import compcert.common.Values.
+Require Import compcert.lib.Coqlib.
+Require Import compcert.lib.Integers.
+Require Import compcert.lib.Maps.
+Require Import compcert.verilog.Op.
+
+Require Import vericert.hls.RTLBlockInstr.
+
+Definition bb := list (list (list instr)).
+
+Definition bblock := @bblock bb.
+Definition code := @code bb.
+Definition function := @function bb.
+Definition fundef := @fundef bb.
+Definition program := @program bb.
+Definition funsig := @funsig bb.
+Definition stackframe := @stackframe bb.
+Definition state := @state bb.
+Definition genv := @genv bb.
+
+Section RELSEM.
+
+ Context (ge: genv).
+
+ Inductive step_instr_list: val -> instr_state -> list instr -> instr_state -> Prop :=
+ | exec_RBcons:
+ forall state i state' state'' instrs sp,
+ step_instr ge sp state i state' ->
+ step_instr_list sp state' instrs state'' ->
+ step_instr_list sp state (i :: instrs) state''
+ | exec_RBnil:
+ forall state sp,
+ step_instr_list sp state nil state.
+
+ Inductive step_instr_seq (sp : val)
+ : instr_state -> list (list instr) -> instr_state -> Prop :=
+ | exec_instr_seq_cons:
+ forall state i state' state'' instrs,
+ step_instr_list sp state i state' ->
+ step_instr_seq sp state' instrs state'' ->
+ step_instr_seq sp state (i :: instrs) state''
+ | exec_instr_seq_nil:
+ forall state,
+ step_instr_seq sp state nil state.
+
+ Inductive step_instr_block (sp : val)
+ : instr_state -> bb -> instr_state -> Prop :=
+ | exec_instr_block_cons:
+ forall state i state' state'' instrs,
+ step_instr_seq sp state i state' ->
+ step_instr_block sp state' instrs state'' ->
+ step_instr_block sp state (i :: instrs) state''
+ | exec_instr_block_nil:
+ forall state,
+ step_instr_block sp state nil state.
+
+ Inductive step: state -> trace -> state -> Prop :=
+ | exec_bblock:
+ forall s f sp pc rs rs' m m' t s' bb pr pr',
+ f.(fn_code)!pc = Some bb ->
+ step_instr_block sp (mk_instr_state rs pr m) bb.(bb_body) (mk_instr_state rs' pr' m') ->
+ step_cf_instr ge (State s f sp pc rs' pr' m') bb.(bb_exit) t s' ->
+ step (State s f sp pc rs pr m) t s'
+ | exec_function_internal:
+ forall s f args m m' stk,
+ Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
+ step (Callstate s (Internal f) args m)
+ E0 (State s
+ f
+ (Vptr stk Ptrofs.zero)
+ f.(fn_entrypoint)
+ (init_regs args f.(fn_params))
+ (PMap.init false)
+ m')
+ | exec_function_external:
+ forall s ef args res t m m',
+ external_call ef ge args m t res m' ->
+ step (Callstate s (External ef) args m)
+ t (Returnstate s res m')
+ | exec_return:
+ forall res f sp pc rs s vres m pr,
+ step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m)
+ E0 (State s f sp pc (rs#res <- vres) pr m).
+
+End RELSEM.
+
+Inductive initial_state (p: program): state -> Prop :=
+ | initial_state_intro: forall b f m0,
+ let ge := Genv.globalenv p in
+ Genv.init_mem p = Some m0 ->
+ Genv.find_symbol ge p.(prog_main) = Some b ->
+ Genv.find_funct_ptr ge b = Some f ->
+ funsig f = signature_main ->
+ initial_state p (Callstate nil f nil m0).
+
+Inductive final_state: state -> int -> Prop :=
+ | final_state_intro: forall r m,
+ final_state (Returnstate nil (Vint r) m) r.
+
+Definition semantics (p: program) :=
+ Semantics step (initial_state p) final_state (Genv.globalenv p).
+
+Definition max_reg_bblock (m : positive) (pc : node) (bb : bblock) :=
+ let max_body := fold_left (fun x l => fold_left (fun x' l' => fold_left max_reg_instr l' x') l x) bb.(bb_body) m in
+ max_reg_cfi max_body bb.(bb_exit).
+
+Definition max_reg_function (f: function) :=
+ Pos.max
+ (PTree.fold max_reg_bblock f.(fn_code) 1%positive)
+ (fold_left Pos.max f.(fn_params) 1%positive).
+
+Definition max_pc_function (f: function) : positive :=
+ PTree.fold (fun m pc i => (Pos.max m
+ (pc + match Zlength i.(bb_body)
+ with Z.pos p => p | _ => 1 end))%positive)
+ f.(fn_code) 1%positive.
+#+end_src
+
+* License
+
+#+name: license
+#+begin_src coq :tangle no
+(*
+ * Vericert: Verified high-level synthesis.
+ * Copyright (C) 2020-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/>.
+ *)
+#+end_src
diff --git a/lit/scheduling.org b/lit/scheduler.org
index 728dfc5..018633c 100644
--- a/lit/scheduling.org
+++ b/lit/scheduler.org
@@ -1,671 +1,7 @@
-#+title: Scheduling
+#+title: Basic Block Generation
#+author: Yann Herklotz
-#+email: yann@yannherklotz.com
+#+email: yann [at] yannherklotz [dot] com
-* Language Definitions
-
-** RTLBlockInstr
-:PROPERTIES:
-:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockInstr.v
-:END:
-
-#+begin_src coq :comments no :padline no :exports none
-<<license>>
-#+end_src
-
-#+name: rtlblockinstr-imports
-#+begin_src coq
-Require Import Coq.micromega.Lia.
-
-Require Import compcert.backend.Registers.
-Require Import compcert.common.AST.
-Require Import compcert.common.Events.
-Require Import compcert.common.Globalenvs.
-Require Import compcert.common.Memory.
-Require Import compcert.common.Values.
-Require Import compcert.lib.Integers.
-Require Import compcert.lib.Maps.
-Require Import compcert.verilog.Op.
-
-Require Import Predicate.
-Require Import Vericertlib.
-#+end_src
-
-These instructions are used for ~RTLBlock~ and ~RTLPar~, so that they have consistent instructions,
-which greatly simplifies the proofs, as they will by default have the same instruction syntax and
-semantics. The only changes are therefore at the top-level of the instructions.
-
-*** Instruction Definition
-
-First, we define the instructions that can be placed into a basic block, meaning they won't branch.
-The main changes to how instructions are defined in [RTL], is that these instructions don't have a
-next node, as they will be in a basic block, and they also have an optional predicate ([pred_op]).
-
-#+name: rtlblockinstr-instr-def
-#+begin_src coq
-Definition node := positive.
-
-Inductive instr : Type :=
-| RBnop : instr
-| RBop : option pred_op -> operation -> list reg -> reg -> instr
-| RBload : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr
-| RBstore : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr
-| RBsetpred : option pred_op -> condition -> list reg -> predicate -> instr.
-#+end_src
-
-*** Control-Flow Instruction Definition
-
-These are the instructions that count as control-flow, and will be placed at the end of the basic
-blocks.
-
-#+name: rtlblockinstr-cf-instr-def
-#+begin_src coq
-Inductive cf_instr : Type :=
-| RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr
-| RBtailcall : signature -> reg + ident -> list reg -> cf_instr
-| RBbuiltin : external_function -> list (builtin_arg reg) ->
- builtin_res reg -> node -> cf_instr
-| RBcond : condition -> list reg -> node -> node -> cf_instr
-| RBjumptable : reg -> list node -> cf_instr
-| RBreturn : option reg -> cf_instr
-| RBgoto : node -> cf_instr
-| RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr.
-#+end_src
-
-*** Helper functions
-
-#+name: rtlblockinstr-helpers
-#+begin_src coq
-Fixpoint successors_instr (i : cf_instr) : list node :=
- match i with
- | RBcall sig ros args res s => s :: nil
- | RBtailcall sig ros args => nil
- | RBbuiltin ef args res s => s :: nil
- | RBcond cond args ifso ifnot => ifso :: ifnot :: nil
- | RBjumptable arg tbl => tbl
- | RBreturn optarg => nil
- | RBgoto n => n :: nil
- | RBpred_cf p c1 c2 => concat (successors_instr c1 :: successors_instr c2 :: nil)
- end.
-
-Definition max_reg_instr (m: positive) (i: instr) :=
- match i with
- | RBnop => m
- | RBop p op args res =>
- fold_left Pos.max args (Pos.max res m)
- | RBload p chunk addr args dst =>
- fold_left Pos.max args (Pos.max dst m)
- | RBstore p chunk addr args src =>
- fold_left Pos.max args (Pos.max src m)
- | RBsetpred p' c args p =>
- fold_left Pos.max args m
- end.
-
-Fixpoint max_reg_cfi (m : positive) (i : cf_instr) :=
- match i with
- | RBcall sig (inl r) args res s =>
- fold_left Pos.max args (Pos.max r (Pos.max res m))
- | RBcall sig (inr id) args res s =>
- fold_left Pos.max args (Pos.max res m)
- | RBtailcall sig (inl r) args =>
- fold_left Pos.max args (Pos.max r m)
- | RBtailcall sig (inr id) args =>
- fold_left Pos.max args m
- | RBbuiltin ef args res s =>
- fold_left Pos.max (params_of_builtin_args args)
- (fold_left Pos.max (params_of_builtin_res res) m)
- | RBcond cond args ifso ifnot => fold_left Pos.max args m
- | RBjumptable arg tbl => Pos.max arg m
- | RBreturn None => m
- | RBreturn (Some arg) => Pos.max arg m
- | RBgoto n => m
- | RBpred_cf p c1 c2 => Pos.max (max_reg_cfi m c1) (max_reg_cfi m c2)
- end.
-
-Definition regset := Regmap.t val.
-Definition predset := PMap.t bool.
-
-Definition eval_predf (pr: predset) (p: pred_op) :=
- sat_predicate p (fun x => pr !! (Pos.of_nat x)).
-
-#[global]
-Instance eval_predf_Proper : Proper (eq ==> equiv ==> eq) eval_predf.
-Proof.
- unfold Proper. simplify. unfold "==>".
- intros.
- unfold sat_equiv in *. intros. unfold eval_predf. subst. apply H0.
-Qed.
-
-#[local] Open Scope pred_op.
-
-Lemma eval_predf_Pand :
- forall ps p p',
- eval_predf ps (p ∧ p') = eval_predf ps p && eval_predf ps p'.
-Proof. unfold eval_predf; split; simplify; auto with bool. Qed.
-
-Lemma eval_predf_Por :
- forall ps p p',
- eval_predf ps (p ∨ p') = eval_predf ps p || eval_predf ps p'.
-Proof. unfold eval_predf; split; simplify; auto with bool. Qed.
-
-Lemma eval_predf_pr_equiv :
- forall p ps ps',
- (forall x, ps !! x = ps' !! x) ->
- eval_predf ps p = eval_predf ps' p.
-Proof.
- induction p; simplify; auto;
- try (unfold eval_predf; simplify; repeat (destruct_match; []); inv Heqp0; rewrite <- H; auto);
- [repeat rewrite eval_predf_Pand|repeat rewrite eval_predf_Por];
- erewrite IHp1; try eassumption; erewrite IHp2; eauto.
-Qed.
-
-Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset :=
- match rl, vl with
- | r1 :: rs, v1 :: vs => Regmap.set r1 v1 (init_regs vs rs)
- | _, _ => Regmap.init Vundef
- end.
-#+end_src
-
-**** Instruction State
-
-Definition of the instruction state, which contains the following:
-
-- ~is_rs~ :: This is the current state of the registers.
-- ~is_ps~ :: This is the current state of the predicate registers, which is in a separate namespace
- and area compared to the standard registers in [is_rs].
-- ~is_mem~ :: The current state of the memory.
-
-#+name: rtlblockinstr-instr-state
-#+begin_src coq
-Record instr_state := mk_instr_state {
- is_rs: regset;
- is_ps: predset;
- is_mem: mem;
-}.
-#+end_src
-
-*** Top-Level Type Definitions
-
-#+name: rtlblockinstr-type-def
-#+begin_src coq
-Section DEFINITION.
-
- Context {bblock_body: Type}.
-
- Record bblock : Type := mk_bblock {
- bb_body: bblock_body;
- bb_exit: cf_instr
- }.
-
- Definition code: Type := PTree.t bblock.
-
- Record function: Type := mkfunction {
- fn_sig: signature;
- fn_params: list reg;
- fn_stacksize: Z;
- fn_code: code;
- fn_entrypoint: node
- }.
-
- Definition fundef := AST.fundef function.
-
- Definition program := AST.program fundef unit.
-
- Definition funsig (fd: fundef) :=
- match fd with
- | Internal f => fn_sig f
- | External ef => ef_sig ef
- end.
-
- Inductive stackframe : Type :=
- | Stackframe:
- forall (res: reg) (**r where to store the result *)
- (f: function) (**r calling function *)
- (sp: val) (**r stack pointer in calling function *)
- (pc: node) (**r program point in calling function *)
- (rs: regset) (**r register state in calling function *)
- (pr: predset), (**r predicate state of the calling function *)
- stackframe.
-
- Inductive state : Type :=
- | State:
- forall (stack: list stackframe) (**r call stack *)
- (f: function) (**r current function *)
- (sp: val) (**r stack pointer *)
- (pc: node) (**r current program point in [c] *)
- (rs: regset) (**r register state *)
- (pr: predset) (**r predicate register state *)
- (m: mem), (**r memory state *)
- state
- | Callstate:
- forall (stack: list stackframe) (**r call stack *)
- (f: fundef) (**r function to call *)
- (args: list val) (**r arguments to the call *)
- (m: mem), (**r memory state *)
- state
- | Returnstate:
- forall (stack: list stackframe) (**r call stack *)
- (v: val) (**r return value for the call *)
- (m: mem), (**r memory state *)
- state.
-
-End DEFINITION.
-#+end_src
-
-*** Semantics
-
-#+name: rtlblockinstr-semantics
-#+begin_src coq
-Section RELSEM.
-
- Context {bblock_body : Type}.
-
- Definition genv := Genv.t (@fundef bblock_body) unit.
-
- Context (ge: genv).
-
- Definition find_function
- (ros: reg + ident) (rs: regset) : option fundef :=
- match ros with
- | inl r => Genv.find_funct ge rs#r
- | inr symb =>
- match Genv.find_symbol ge symb with
- | None => None
- | Some b => Genv.find_funct_ptr ge b
- end
- end.
-
- Inductive eval_pred: option pred_op -> instr_state -> instr_state -> instr_state -> Prop :=
- | eval_pred_true:
- forall i i' p,
- eval_predf (is_ps i) p = true ->
- eval_pred (Some p) i i' i'
- | eval_pred_false:
- forall i i' p,
- eval_predf (is_ps i) p = false ->
- eval_pred (Some p) i i' i
- | eval_pred_none:
- forall i i', eval_pred None i i' i.
-
- Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop :=
- | exec_RBnop:
- forall sp ist,
- step_instr sp ist RBnop ist
- | exec_RBop:
- forall op v res args rs m sp p ist pr,
- eval_operation ge sp op rs##args m = Some v ->
- eval_pred p (mk_instr_state rs pr m) (mk_instr_state (rs#res <- v) pr m) ist ->
- step_instr sp (mk_instr_state rs pr m) (RBop p op args res) ist
- | exec_RBload:
- forall addr rs args a chunk m v dst sp p pr ist,
- eval_addressing ge sp addr rs##args = Some a ->
- Mem.loadv chunk m a = Some v ->
- eval_pred p (mk_instr_state rs pr m) (mk_instr_state (rs#dst <- v) pr m) ist ->
- step_instr sp (mk_instr_state rs pr m) (RBload p chunk addr args dst) ist
- | exec_RBstore:
- forall addr rs args a chunk m src m' sp p pr ist,
- eval_addressing ge sp addr rs##args = Some a ->
- Mem.storev chunk m a rs#src = Some m' ->
- eval_pred p (mk_instr_state rs pr m) (mk_instr_state rs pr m') ist ->
- step_instr sp (mk_instr_state rs pr m) (RBstore p chunk addr args src) ist
- | exec_RBsetpred:
- forall sp rs pr m p c b args p' ist,
- Op.eval_condition c rs##args m = Some b ->
- eval_pred p' (mk_instr_state rs pr m) (mk_instr_state rs (pr#p <- b) m) ist ->
- step_instr sp (mk_instr_state rs pr m) (RBsetpred p' c args p) ist.
-
- Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop :=
- | exec_RBcall:
- forall s f sp rs m res fd ros sig args pc pc' pr,
- find_function ros rs = Some fd ->
- funsig fd = sig ->
- step_cf_instr (State s f sp pc rs pr m) (RBcall sig ros args res pc')
- E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m)
- | exec_RBtailcall:
- forall s f stk rs m sig ros args fd m' pc pr,
- find_function ros rs = Some fd ->
- funsig fd = sig ->
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) (RBtailcall sig ros args)
- E0 (Callstate s fd rs##args m')
- | exec_RBbuiltin:
- forall s f sp rs m ef args res pc' vargs t vres m' pc pr,
- eval_builtin_args ge (fun r => rs#r) sp m args vargs ->
- external_call ef ge vargs m t vres m' ->
- step_cf_instr (State s f sp pc rs pr m) (RBbuiltin ef args res pc')
- t (State s f sp pc' (regmap_setres res vres rs) pr m')
- | exec_RBcond:
- forall s f sp rs m cond args ifso ifnot b pc pc' pr,
- eval_condition cond rs##args m = Some b ->
- pc' = (if b then ifso else ifnot) ->
- step_cf_instr (State s f sp pc rs pr m) (RBcond cond args ifso ifnot)
- E0 (State s f sp pc' rs pr m)
- | exec_RBjumptable:
- forall s f sp rs m arg tbl n pc pc' pr,
- rs#arg = Vint n ->
- list_nth_z tbl (Int.unsigned n) = Some pc' ->
- step_cf_instr (State s f sp pc rs pr m) (RBjumptable arg tbl)
- E0 (State s f sp pc' rs pr m)
- | exec_RBreturn:
- forall s f stk rs m or pc m' pr,
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) (RBreturn or)
- E0 (Returnstate s (regmap_optget or Vundef rs) m')
- | exec_RBgoto:
- forall s f sp pc rs pr m pc',
- step_cf_instr (State s f sp pc rs pr m) (RBgoto pc') E0 (State s f sp pc' rs pr m)
- | exec_RBpred_cf:
- forall s f sp pc rs pr m cf1 cf2 st' p t,
- step_cf_instr (State s f sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' ->
- step_cf_instr (State s f sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'.
-
-End RELSEM.
-#+end_src
-
-** RTLBlock
-:PROPERTIES:
-:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlock.v
-:END:
-
-#+begin_src coq :comments no :padline no :exports none
-<<license>>
-#+end_src
-
-#+name: rtlblock-main
-#+begin_src coq
-Require Import compcert.backend.Registers.
-Require Import compcert.common.AST.
-Require Import compcert.common.Events.
-Require Import compcert.common.Globalenvs.
-Require Import compcert.common.Memory.
-Require Import compcert.common.Smallstep.
-Require Import compcert.common.Values.
-Require Import compcert.lib.Coqlib.
-Require Import compcert.lib.Integers.
-Require Import compcert.lib.Maps.
-Require Import compcert.verilog.Op.
-
-Require Import vericert.hls.RTLBlockInstr.
-
-Definition bb := list instr.
-
-Definition bblock := @bblock bb.
-Definition code := @code bb.
-Definition function := @function bb.
-Definition fundef := @fundef bb.
-Definition program := @program bb.
-Definition funsig := @funsig bb.
-Definition stackframe := @stackframe bb.
-
-Definition genv := @genv bb.
-
-Inductive state : Type :=
-| State:
- forall (stack: list stackframe) (**r call stack *)
- (f: function) (**r current function *)
- (b: bb) (**r current block being executed *)
- (sp: val) (**r stack pointer *)
- (pc: node) (**r current program point in [c] *)
- (rs: regset) (**r register state *)
- (pr: predset) (**r predicate register state *)
- (m: mem), (**r memory state *)
- state
-| Callstate:
- forall (stack: list stackframe) (**r call stack *)
- (f: fundef) (**r function to call *)
- (args: list val) (**r arguments to the call *)
- (m: mem), (**r memory state *)
- state
-| Returnstate:
- forall (stack: list stackframe) (**r call stack *)
- (v: val) (**r return value for the call *)
- (m: mem), (**r memory state *)
- state.
-
-Section RELSEM.
-
- Context (ge: genv).
-
- Inductive step_instr_list: val -> instr_state -> list instr -> instr_state -> Prop :=
- | exec_RBcons:
- forall state i state' state'' instrs sp,
- step_instr ge sp state i state' ->
- step_instr_list sp state' instrs state'' ->
- step_instr_list sp state (i :: instrs) state''
- | exec_RBnil:
- forall state sp,
- step_instr_list sp state nil state.
-
- Definition find_function
- (ros: reg + ident) (rs: regset) : option fundef :=
- match ros with
- | inl r => Genv.find_funct ge rs#r
- | inr symb =>
- match Genv.find_symbol ge symb with
- | None => None
- | Some b => Genv.find_funct_ptr ge b
- end
- end.
-
- Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop :=
- | exec_RBcall:
- forall s f b sp rs m res fd ros sig args pc pc' pr,
- find_function ros rs = Some fd ->
- funsig fd = sig ->
- step_cf_instr (State s f b sp pc rs pr m) (RBcall sig ros args res pc')
- E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m)
- | exec_RBtailcall:
- forall s f b stk rs m sig ros args fd m' pc pr,
- find_function ros rs = Some fd ->
- funsig fd = sig ->
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step_cf_instr (State s f b (Vptr stk Ptrofs.zero) pc rs pr m) (RBtailcall sig ros args)
- E0 (Callstate s fd rs##args m')
- | exec_RBbuiltin:
- forall s f b sp rs m ef args res pc' vargs t vres m' pc pr,
- eval_builtin_args ge (fun r => rs#r) sp m args vargs ->
- external_call ef ge vargs m t vres m' ->
- step_cf_instr (State s f b sp pc rs pr m) (RBbuiltin ef args res pc')
- t (State s f b sp pc' (regmap_setres res vres rs) pr m')
- | exec_RBcond:
- forall s f block sp rs m cond args ifso ifnot b pc pc' pr,
- eval_condition cond rs##args m = Some b ->
- pc' = (if b then ifso else ifnot) ->
- step_cf_instr (State s f block sp pc rs pr m) (RBcond cond args ifso ifnot)
- E0 (State s f block sp pc' rs pr m)
- | exec_RBjumptable:
- forall s f b sp rs m arg tbl n pc pc' pr,
- rs#arg = Vint n ->
- list_nth_z tbl (Int.unsigned n) = Some pc' ->
- step_cf_instr (State s f b sp pc rs pr m) (RBjumptable arg tbl)
- E0 (State s f b sp pc' rs pr m)
- | exec_RBreturn:
- forall s f b stk rs m or pc m' pr,
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step_cf_instr (State s f b (Vptr stk Ptrofs.zero) pc rs pr m) (RBreturn or)
- E0 (Returnstate s (regmap_optget or Vundef rs) m')
- | exec_RBgoto:
- forall s f b sp pc rs pr m pc',
- step_cf_instr (State s f b sp pc rs pr m) (RBgoto pc') E0 (State s f b sp pc' rs pr m)
- | exec_RBpred_cf:
- forall s f b sp pc rs pr m cf1 cf2 st' p t,
- step_cf_instr (State s f b sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' ->
- step_cf_instr (State s f b sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'.
-
- Inductive step: state -> trace -> state -> Prop :=
- | exec_function_internal:
- forall s f args m m' stk,
- Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
- step (Callstate s (Internal f) args m)
- E0 (State s f
- (Vptr stk Ptrofs.zero)
- f.(fn_entrypoint)
- (init_regs args f.(fn_params))
- (PMap.init false)
- m')
- | exec_function_external:
- forall s ef args res t m m',
- external_call ef ge args m t res m' ->
- step (Callstate s (External ef) args m)
- t (Returnstate s res m')
- | exec_return:
- forall res f sp pc rs s vres m pr,
- step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m)
- E0 (State s f sp pc (rs#res <- vres) pr m).
-
-End RELSEM.
-
-Inductive initial_state (p: program): state -> Prop :=
- | initial_state_intro: forall b f m0,
- let ge := Genv.globalenv p in
- Genv.init_mem p = Some m0 ->
- Genv.find_symbol ge p.(prog_main) = Some b ->
- Genv.find_funct_ptr ge b = Some f ->
- funsig f = signature_main ->
- initial_state p (Callstate nil f nil m0).
-
-Inductive final_state: state -> int -> Prop :=
- | final_state_intro: forall r m,
- final_state (Returnstate nil (Vint r) m) r.
-
-Definition semantics (p: program) :=
- Semantics step (initial_state p) final_state (Genv.globalenv p).
-#+end_src
-
-** RTLPar
-:PROPERTIES:
-:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPar.v
-:END:
-
-#+begin_src coq :comments no :padline no :exports none
-<<license>>
-#+end_src
-
-#+name: rtlpar-main
-#+begin_src coq
-Require Import compcert.backend.Registers.
-Require Import compcert.common.AST.
-Require Import compcert.common.Events.
-Require Import compcert.common.Globalenvs.
-Require Import compcert.common.Memory.
-Require Import compcert.common.Smallstep.
-Require Import compcert.common.Values.
-Require Import compcert.lib.Coqlib.
-Require Import compcert.lib.Integers.
-Require Import compcert.lib.Maps.
-Require Import compcert.verilog.Op.
-
-Require Import vericert.hls.RTLBlockInstr.
-
-Definition bb := list (list (list instr)).
-
-Definition bblock := @bblock bb.
-Definition code := @code bb.
-Definition function := @function bb.
-Definition fundef := @fundef bb.
-Definition program := @program bb.
-Definition funsig := @funsig bb.
-Definition stackframe := @stackframe bb.
-Definition state := @state bb.
-Definition genv := @genv bb.
-
-Section RELSEM.
-
- Context (ge: genv).
-
- Inductive step_instr_list: val -> instr_state -> list instr -> instr_state -> Prop :=
- | exec_RBcons:
- forall state i state' state'' instrs sp,
- step_instr ge sp state i state' ->
- step_instr_list sp state' instrs state'' ->
- step_instr_list sp state (i :: instrs) state''
- | exec_RBnil:
- forall state sp,
- step_instr_list sp state nil state.
-
- Inductive step_instr_seq (sp : val)
- : instr_state -> list (list instr) -> instr_state -> Prop :=
- | exec_instr_seq_cons:
- forall state i state' state'' instrs,
- step_instr_list sp state i state' ->
- step_instr_seq sp state' instrs state'' ->
- step_instr_seq sp state (i :: instrs) state''
- | exec_instr_seq_nil:
- forall state,
- step_instr_seq sp state nil state.
-
- Inductive step_instr_block (sp : val)
- : instr_state -> bb -> instr_state -> Prop :=
- | exec_instr_block_cons:
- forall state i state' state'' instrs,
- step_instr_seq sp state i state' ->
- step_instr_block sp state' instrs state'' ->
- step_instr_block sp state (i :: instrs) state''
- | exec_instr_block_nil:
- forall state,
- step_instr_block sp state nil state.
-
- Inductive step: state -> trace -> state -> Prop :=
- | exec_bblock:
- forall s f sp pc rs rs' m m' t s' bb pr pr',
- f.(fn_code)!pc = Some bb ->
- step_instr_block sp (mk_instr_state rs pr m) bb.(bb_body) (mk_instr_state rs' pr' m') ->
- step_cf_instr ge (State s f sp pc rs' pr' m') bb.(bb_exit) t s' ->
- step (State s f sp pc rs pr m) t s'
- | exec_function_internal:
- forall s f args m m' stk,
- Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
- step (Callstate s (Internal f) args m)
- E0 (State s
- f
- (Vptr stk Ptrofs.zero)
- f.(fn_entrypoint)
- (init_regs args f.(fn_params))
- (PMap.init false)
- m')
- | exec_function_external:
- forall s ef args res t m m',
- external_call ef ge args m t res m' ->
- step (Callstate s (External ef) args m)
- t (Returnstate s res m')
- | exec_return:
- forall res f sp pc rs s vres m pr,
- step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m)
- E0 (State s f sp pc (rs#res <- vres) pr m).
-
-End RELSEM.
-
-Inductive initial_state (p: program): state -> Prop :=
- | initial_state_intro: forall b f m0,
- let ge := Genv.globalenv p in
- Genv.init_mem p = Some m0 ->
- Genv.find_symbol ge p.(prog_main) = Some b ->
- Genv.find_funct_ptr ge b = Some f ->
- funsig f = signature_main ->
- initial_state p (Callstate nil f nil m0).
-
-Inductive final_state: state -> int -> Prop :=
- | final_state_intro: forall r m,
- final_state (Returnstate nil (Vint r) m) r.
-
-Definition semantics (p: program) :=
- Semantics step (initial_state p) final_state (Genv.globalenv p).
-
-Definition max_reg_bblock (m : positive) (pc : node) (bb : bblock) :=
- let max_body := fold_left (fun x l => fold_left (fun x' l' => fold_left max_reg_instr l' x') l x) bb.(bb_body) m in
- max_reg_cfi max_body bb.(bb_exit).
-
-Definition max_reg_function (f: function) :=
- Pos.max
- (PTree.fold max_reg_bblock f.(fn_code) 1%positive)
- (fold_left Pos.max f.(fn_params) 1%positive).
-
-Definition max_pc_function (f: function) : positive :=
- PTree.fold (fun m pc i => (Pos.max m
- (pc + match Zlength i.(bb_body)
- with Z.pos p => p | _ => 1 end))%positive)
- f.(fn_code) 1%positive.
-#+end_src
* Scheduler
:PROPERTIES:
:header-args:ocaml: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/Schedule.ml
@@ -1544,350 +880,6 @@ let schedule_fn (f : RTLBlock.coq_function) : RTLPar.coq_function =
}
#+end_src
-* RTLBlockgen
-:PROPERTIES:
-:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgen.v
-:END:
-
-Refers to [[rtlblockgen-equalities][rtlblockgen-equalities]].
-
-#+begin_src coq :comments no :padline no :exports none
-<<license>>
-#+end_src
-
-#+name: rtlblockgen-imports
-#+begin_src coq
-Require compcert.backend.RTL.
-Require Import compcert.common.AST.
-Require Import compcert.lib.Maps.
-Require Import compcert.lib.Integers.
-Require Import compcert.lib.Floats.
-
-Require Import vericert.common.Vericertlib.
-Require Import vericert.hls.RTLBlockInstr.
-Require Import vericert.hls.RTLBlock.
-
-#[local] Open Scope positive.
-#+end_src
-
-#+name: rtlblockgen-equalities-insert
-#+begin_src coq
-Lemma comparison_eq: forall (x y : comparison), {x = y} + {x <> y}.
-Proof.
- decide equality.
-Defined.
-
-Lemma condition_eq: forall (x y : Op.condition), {x = y} + {x <> y}.
-Proof.
- generalize comparison_eq; intro.
- generalize Int.eq_dec; intro.
- generalize Int64.eq_dec; intro.
- decide equality.
-Defined.
-
-Lemma addressing_eq : forall (x y : Op.addressing), {x = y} + {x <> y}.
-Proof.
- generalize Int.eq_dec; intro.
- generalize AST.ident_eq; intro.
- generalize Z.eq_dec; intro.
- generalize Ptrofs.eq_dec; intro.
- decide equality.
-Defined.
-
-Lemma typ_eq : forall (x y : AST.typ), {x = y} + {x <> y}.
-Proof.
- decide equality.
-Defined.
-
-Lemma operation_eq: forall (x y : Op.operation), {x = y} + {x <> y}.
-Proof.
- generalize Int.eq_dec; intro.
- generalize Int64.eq_dec; intro.
- generalize Float.eq_dec; intro.
- generalize Float32.eq_dec; intro.
- generalize AST.ident_eq; intro.
- generalize condition_eq; intro.
- generalize addressing_eq; intro.
- generalize typ_eq; intro.
- decide equality.
-Defined.
-
-Lemma memory_chunk_eq : forall (x y : AST.memory_chunk), {x = y} + {x <> y}.
-Proof.
- decide equality.
-Defined.
-
-Lemma list_typ_eq: forall (x y : list AST.typ), {x = y} + {x <> y}.
-Proof.
- generalize typ_eq; intro.
- decide equality.
-Defined.
-
-Lemma option_typ_eq : forall (x y : option AST.typ), {x = y} + {x <> y}.
-Proof.
- generalize typ_eq; intro.
- decide equality.
-Defined.
-
-Lemma signature_eq: forall (x y : AST.signature), {x = y} + {x <> y}.
-Proof.
- repeat decide equality.
-Defined.
-
-Lemma list_operation_eq : forall (x y : list Op.operation), {x = y} + {x <> y}.
-Proof.
- generalize operation_eq; intro.
- decide equality.
-Defined.
-
-Lemma list_pos_eq : forall (x y : list positive), {x = y} + {x <> y}.
-Proof.
- generalize Pos.eq_dec; intros.
- decide equality.
-Defined.
-
-Lemma sig_eq : forall (x y : AST.signature), {x = y} + {x <> y}.
-Proof.
- repeat decide equality.
-Defined.
-
-Lemma instr_eq: forall (x y : instr), {x = y} + {x <> y}.
-Proof.
- generalize Pos.eq_dec; intro.
- generalize typ_eq; intro.
- generalize Int.eq_dec; intro.
- generalize memory_chunk_eq; intro.
- generalize addressing_eq; intro.
- generalize operation_eq; intro.
- generalize condition_eq; intro.
- generalize signature_eq; intro.
- generalize list_operation_eq; intro.
- generalize list_pos_eq; intro.
- generalize AST.ident_eq; intro.
- repeat decide equality.
-Defined.
-
-Lemma cf_instr_eq: forall (x y : cf_instr), {x = y} + {x <> y}.
-Proof.
- generalize Pos.eq_dec; intro.
- generalize typ_eq; intro.
- generalize Int.eq_dec; intro.
- generalize Int64.eq_dec; intro.
- generalize Float.eq_dec; intro.
- generalize Float32.eq_dec; intro.
- generalize Ptrofs.eq_dec; intro.
- generalize memory_chunk_eq; intro.
- generalize addressing_eq; intro.
- generalize operation_eq; intro.
- generalize condition_eq; intro.
- generalize signature_eq; intro.
- generalize list_operation_eq; intro.
- generalize list_pos_eq; intro.
- generalize AST.ident_eq; intro.
- repeat decide equality.
-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.
-#+end_src
-
-#+name: rtlblockgen-main
-#+begin_src coq
-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. *)
-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.*)
-
-Definition check_instr (n: positive) (istr: RTL.instruction) (istr': instr) :=
- match istr, istr' with
- | RTL.Inop n', RBnop => (n' + 1 =? n)
- | RTL.Iop op args dst n', RBop None op' args' dst' =>
- ceq operation_eq op op' &&
- ceq list_pos_eq args args' &&
- ceq peq dst dst' && (n' + 1 =? n)
- | RTL.Iload chunk addr args dst n', RBload None chunk' addr' args' dst' =>
- ceq memory_chunk_eq chunk chunk' &&
- ceq addressing_eq addr addr' &&
- ceq list_pos_eq args args' &&
- ceq peq dst dst' &&
- (n' + 1 =? n)
- | RTL.Istore chunk addr args src n', RBstore None chunk' addr' args' src' =>
- ceq memory_chunk_eq chunk chunk' &&
- ceq addressing_eq addr addr' &&
- ceq list_pos_eq args args' &&
- ceq peq src src' &&
- (n' + 1 =? n)
- | _, _ => false
- end.
-
-Definition check_cf_instr_body (istr: RTL.instruction) (istr': instr): bool :=
- match istr, istr' with
- | RTL.Iop op args dst _, RBop None op' args' dst' =>
- ceq operation_eq op op' &&
- ceq list_pos_eq args args' &&
- ceq peq dst dst'
- | RTL.Iload chunk addr args dst _, RBload None chunk' addr' args' dst' =>
- ceq memory_chunk_eq chunk chunk' &&
- ceq addressing_eq addr addr' &&
- ceq list_pos_eq args args' &&
- ceq peq dst dst'
- | RTL.Istore chunk addr args src _, RBstore None chunk' addr' args' src' =>
- ceq memory_chunk_eq chunk chunk' &&
- ceq addressing_eq addr addr' &&
- ceq list_pos_eq args args' &&
- ceq peq src src'
- | RTL.Inop _, RBnop
- | RTL.Icall _ _ _ _ _, RBnop
- | RTL.Itailcall _ _ _, RBnop
- | RTL.Ibuiltin _ _ _ _, RBnop
- | RTL.Icond _ _ _ _, RBnop
- | RTL.Ijumptable _ _, RBnop
- | RTL.Ireturn _, RBnop => true
- | _, _ => false
- end.
-
-Definition check_cf_instr (istr: RTL.instruction) (istr': cf_instr) :=
- match istr, istr' with
- | RTL.Inop n, RBgoto n' => (n =? n')
- | RTL.Iop _ _ _ n, RBgoto n' => (n =? n')
- | RTL.Iload _ _ _ _ n, RBgoto n' => (n =? n')
- | RTL.Istore _ _ _ _ n, RBgoto n' => (n =? n')
- | RTL.Icall sig (inl r) args dst n, RBcall sig' (inl r') args' dst' n' =>
- ceq signature_eq sig sig' &&
- ceq peq r r' &&
- ceq list_pos_eq args args' &&
- ceq peq dst dst' &&
- (n =? n')
- | RTL.Icall sig (inr i) args dst n, RBcall sig' (inr i') args' dst' n' =>
- ceq signature_eq sig sig' &&
- ceq peq i i' &&
- ceq list_pos_eq args args' &&
- ceq peq dst dst' &&
- (n =? n')
- | RTL.Itailcall sig (inl r) args, RBtailcall sig' (inl r') args' =>
- ceq signature_eq sig sig' &&
- ceq peq r r' &&
- ceq list_pos_eq args args'
- | RTL.Itailcall sig (inr r) args, RBtailcall sig' (inr r') args' =>
- ceq signature_eq sig sig' &&
- ceq peq r r' &&
- ceq list_pos_eq args args'
- | RTL.Icond cond args n1 n2, RBcond cond' args' n1' n2' =>
- ceq condition_eq cond cond' &&
- ceq list_pos_eq args args' &&
- ceq peq n1 n1' && ceq peq n2 n2'
- | RTL.Ijumptable r ns, RBjumptable r' ns' =>
- ceq peq r r' && ceq list_pos_eq ns ns'
- | RTL.Ireturn (Some r), RBreturn (Some r') =>
- ceq peq r r'
- | RTL.Ireturn None, RBreturn None => true
- | _, _ => false
- end.
-
-Definition is_cf_instr (n: positive) (i: RTL.instruction) :=
- match i with
- | RTL.Inop n' => negb (n' + 1 =? n)
- | RTL.Iop _ _ _ n' => negb (n' + 1 =? n)
- | RTL.Iload _ _ _ _ n' => negb (n' + 1 =? n)
- | RTL.Istore _ _ _ _ n' => negb (n' + 1 =? n)
- | RTL.Icall _ _ _ _ _ => true
- | RTL.Itailcall _ _ _ => true
- | RTL.Ibuiltin _ _ _ _ => true
- | RTL.Icond _ _ _ _ => true
- | RTL.Ijumptable _ _ => true
- | RTL.Ireturn _ => true
- end.
-
-Definition check_present_blocks (c: code) (n: list positive) (max: positive) (i: positive) (istr: RTL.instruction) :=
- let blockn := find_block max n i in
- match c ! blockn with
- | Some istrs =>
- match List.nth_error istrs.(bb_body) (Pos.to_nat blockn - Pos.to_nat i)%nat with
- | Some istr' =>
- if is_cf_instr i istr
- then check_cf_instr istr istrs.(bb_exit) && check_cf_instr_body istr istr'
- else check_instr i istr istr'
- | None => false
- end
- | None => false
- end.
-
-Definition transl_function (f: RTL.function) :=
- match partition f with
- | Errors.OK f' =>
- let blockids := map fst (PTree.elements f'.(fn_code)) in
- if forall_ptree (check_present_blocks f'.(fn_code) blockids (fold_right Pos.max 1 blockids))
- f.(RTL.fn_code) then
- Errors.OK f'
- else Errors.Error (Errors.msg "check_present_blocks failed")
- | Errors.Error msg => Errors.Error msg
- end.
-
-Definition transl_fundef := transf_partial_fundef transl_function.
-
-Definition transl_program : RTL.program -> Errors.res program :=
- transform_partial_program transl_fundef.
-#+end_src
-
-* RTLBlockgenproof
-:PROPERTIES:
-:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgenproof.v
-:END:
-
-#+begin_src coq :comments no :padline no :exports none
-<<license>>
-#+end_src
-
-** Imports
-
-#+name: rtlblockgenproof-imports
-#+begin_src coq
-Require compcert.backend.RTL.
-Require Import compcert.common.AST.
-Require Import compcert.lib.Maps.
-
-Require Import vericert.hls.RTLBlock.
-Require Import vericert.hls.RTLBlockgen.
-#+end_src
-
-** Match states
-
-The ~match_states~ predicate describes which states are equivalent between the two languages, in this
-case ~RTL~ and ~RTLBlock~.
-
-#+name: rtlblockgenproof-match-states
-#+begin_src coq
-Inductive match_states : RTL.state -> RTLBlock.state -> Prop :=
-| match_state :
- forall stk f tf sp pc rs m
- (TF: transl_function f = OK tf),
- match_states (RTL.State stk f sp pc rs m)
- (RTLBlock.State stk tf sp (find_block max n i) rs m).
-#+end_src
-
-** Correctness
-
-#+name: rtlblockgenproof-correctness
-#+begin_src coq
-Section CORRECTNESS.
-
- Context (prog : RTL.program).
- Context (tprog : RTLBlock.program).
-
- Context (TRANSL : match_prog prog tprog).
-
- Theorem transf_program_correct:
- Smallstep.forward_simulation (RTL.semantics prog) (RTLBlock.semantics tprog).
- Proof.
- eapply Smallstep.forward_simulation_plus; eauto with htlproof.
- apply senv_preserved.
-
-End CORRECTNESS.
-#+end_src
* RTLPargen
:PROPERTIES:
:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargen.v
@@ -3235,7 +2227,7 @@ Proof. induction 2; try rewrite H; eauto with barg. Qed.
eauto. eauto. simplify. eauto. eauto. }
{ unfold bind in *. inv TRANSL0. clear Learn. inv H0. destruct_match; crush.
inv H2. unfold transl_function in Heqr. destruct_match; crush.
- inv Heqr.
+ inv Heqr.
repeat econstructor; eauto.
unfold bind in *. destruct_match; crush. }
{ inv TRANSL0. repeat econstructor; eauto using Events.external_call_symbols_preserved, symbols_preserved, senv_preserved, Events.E0_right. }