From c82132030fde68daeda7000fc9db0fa15e819904 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Tue, 3 May 2022 16:18:45 -0400 Subject: [PATCH 01/62] remove extraneous Search command --- compiler/src/compiler/FlatToRiscvFunctions.v | 1 - 1 file changed, 1 deletion(-) diff --git a/compiler/src/compiler/FlatToRiscvFunctions.v b/compiler/src/compiler/FlatToRiscvFunctions.v index 025ad77e0..e3fdea80a 100644 --- a/compiler/src/compiler/FlatToRiscvFunctions.v +++ b/compiler/src/compiler/FlatToRiscvFunctions.v @@ -1606,7 +1606,6 @@ Section Proofs. | x := _ |- _ => clearbody x end. - Search initialL_metrics. clear - word_ok RVM PRParams PR ext_spec word_riscv_ok locals_ok mem_ok fun_info_ok env_ok IHexec OC BC OL Exb GetMany Ext GE FS C V Mo Mo' Gra RaM GPC A GM. revert IHexec OC BC OL Exb GetMany Ext GE FS C V Mo Mo' Gra RaM GPC A GM. From 34513c0775df1ce6f7578e8b340a0a248042ad80 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Tue, 3 May 2022 20:01:05 -0400 Subject: [PATCH 02/62] add isReg to FlatImp semantics, other files still broken --- compiler/src/compiler/FlatImp.v | 171 +++++++++++++++++++++++--------- 1 file changed, 126 insertions(+), 45 deletions(-) diff --git a/compiler/src/compiler/FlatImp.v b/compiler/src/compiler/FlatImp.v index 4bf1c0e74..60972bfa5 100644 --- a/compiler/src/compiler/FlatImp.v +++ b/compiler/src/compiler/FlatImp.v @@ -28,7 +28,7 @@ Inductive bbinop: Type := Section Syntax. Context {varname: Type}. - + Inductive bcond: Type := | CondBinary (op: bbinop) (x y: varname) | CondNez (x: varname) @@ -268,6 +268,8 @@ Module exec. {env_ok: map.ok env} {ext_spec_ok: ext_spec.ok ext_spec}. + Variable (isReg: varname -> bool). + Variable (e: env). Local Notation metrics := MetricLog. @@ -275,6 +277,107 @@ Module exec. (* COQBUG(unification finds Type instead of Prop and fails to downgrade *) Implicit Types post : trace -> mem -> locals -> metrics -> Prop. + + (* Helper functions for computing costs of instructions *) + Definition cost_SLoad x a mc := + match (isReg x, isReg a) with + | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricStores 1 mc))) + | (false, true) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) + | ( true, false) => (addMetricInstructions 2 (addMetricLoads 4 mc)) + | ( true, true) => (addMetricInstructions 1 (addMetricLoads 2 mc)) + end. + + Definition cost_SStore a v mc := + match (isReg a, isReg v) with + | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricStores 1 mc))) + | (false, true) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) + | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) + | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc))) + end. + + Definition cost_SInlinetable x i mc := + match (isReg x, isReg i) with + | (false, false) => (addMetricInstructions 5 (addMetricLoads 7 (addMetricStores 1 (addMetricJumps 1 mc)))) + | (false, true) => (addMetricInstructions 4 (addMetricLoads 5 (addMetricStores 1 (addMetricJumps 1 mc)))) + | ( true, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) + | ( true, true) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) + end. + + Definition cost_SStackalloc x mc := + match isReg x with + | false => (addMetricInstructions 2 (addMetricLoads 2 (addMetricStores 1 mc))) + | true => (addMetricInstructions 1 (addMetricLoads 1 mc)) + end. + + Definition cost_SLit x mc := + match isReg x with + | false => (addMetricInstructions 9 (addMetricLoads 9 (addMetricStores 1 mc))) + | true => (addMetricInstructions 8 (addMetricLoads 8 mc)) + end. + + Definition cost_SOp x y z mc := + match (isReg x, isReg y, isReg z) with + | (false, false, false) => (addMetricInstructions 5 (addMetricLoads 7 (addMetricStores 1 mc))) + | (false, false, true) | (false, true, false) => (addMetricInstructions 4 (addMetricLoads 5 (addMetricStores 1 mc))) + | (false, true, true) => (addMetricInstructions 3 (addMetricLoads 3 (addMetricStores 1 mc))) + | ( true, false, false) => (addMetricInstructions 4 (addMetricLoads 6 mc)) + | ( true, false, true) | ( true, true, false) => (addMetricInstructions 3 (addMetricLoads 4 mc)) + | ( true, true, true) => (addMetricInstructions 2 (addMetricLoads 2 mc)) + end. + + Definition cost_SSet x y mc := + match (isReg x, isReg y) with + | (false, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricStores 1 mc))) + | (false, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricStores 1 mc))) + | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 mc)) + | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 mc)) + end. + + Definition cost_SIf bcond mc := + match bcond with + | CondBinary _ x y => + match (isReg x, isReg y) with + | (false, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) + | (false, true) | ( true, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) + | ( true, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) + end + | CondNez x => + match isReg x with + | false => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) + | true => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) + end + end. + + Definition cost_SLoop_true bcond mc := + match bcond with + | CondBinary _ x y => + match (isReg x, isReg y) with + | (false, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) + | (false, true) | ( true, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) + | ( true, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) + end + | CondNez x => + match isReg x with + | false => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) + | true => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) + end + end. + + Definition cost_SLoop_false bcond mc := + match bcond with + | CondBinary _ x y => + match (isReg x, isReg y) with + | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricJumps 1 mc))) + | (false, true) | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricJumps 1 mc))) + | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mc))) + end + | CondNez x => + match isReg x with + | false => (addMetricInstructions 2 (addMetricLoads 3 (addMetricJumps 1 mc))) + | true => (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mc))) + end + end. + (* alternative semantics which allow non-determinism *) Inductive exec: stmt varname -> @@ -294,6 +397,7 @@ Module exec. (addMetricStores 1 (addMetricLoads 2 mc)))) -> exec (SInteract resvars action argvars) t m l mc post + (* TODO check SInteract case costs *) | call: forall t m l mc binds fname args params rets fbody argvs st0 post outcome, map.get e fname = Some (params, rets, fbody) -> map.getmany_of_list l args = Some argvs -> @@ -310,35 +414,27 @@ Module exec. | load: forall t m l mc sz x a o v addr post, map.get l a = Some addr -> load sz m (word.add addr (word.of_Z o)) = Some v -> - post t m (map.put l x v) - (addMetricLoads 2 - (addMetricInstructions 1 mc)) -> + post t m (map.put l x v) (cost_SLoad x a mc)-> exec (SLoad sz x a o) t m l mc post | store: forall t m m' mc l sz a o addr v val post, map.get l a = Some addr -> map.get l v = Some val -> store sz m (word.add addr (word.of_Z o)) val = Some m' -> - post t m' l - (addMetricLoads 1 - (addMetricInstructions 1 - (addMetricStores 1 mc))) -> + post t m' l (cost_SStore a v mc) -> exec (SStore sz a v o) t m l mc post | inlinetable: forall sz x table i v index t m l mc post, (* compiled riscv code uses x as a tmp register and this shouldn't overwrite i *) x <> i -> map.get l i = Some index -> load sz (map.of_list_word table) index = Some v -> - post t m (map.put l x v) - (addMetricLoads 4 - (addMetricInstructions 3 - (addMetricJumps 1 mc))) -> + post t m (map.put l x v) (cost_SInlinetable x i mc) -> exec (SInlinetable sz x table i) t m l mc post | stackalloc: forall t mSmall l mc x n body post, n mod (bytes_per_word width) = 0 -> (forall a mStack mCombined, anybytes a n mStack -> map.split mCombined mSmall mStack -> - exec body t mCombined (map.put l x a) (addMetricLoads 1 (addMetricInstructions 1 mc)) + exec body t mCombined (map.put l x a) (cost_SStackalloc x mc) (fun t' mCombined' l' mc' => exists mSmall' mStack', anybytes a n mStack' /\ @@ -346,36 +442,24 @@ Module exec. post t' mSmall' l' mc')) -> exec (SStackalloc x n body) t mSmall l mc post | lit: forall t m l mc x v post, - post t m (map.put l x (word.of_Z v)) - (addMetricLoads 8 - (addMetricInstructions 8 mc)) -> + post t m (map.put l x (word.of_Z v)) (cost_SLit x mc) -> exec (SLit x v) t m l mc post | op: forall t m l mc x op y y' z z' post, map.get l y = Some y' -> map.get l z = Some z' -> - post t m (map.put l x (interp_binop op y' z')) - (addMetricLoads 2 - (addMetricInstructions 2 mc)) -> + post t m (map.put l x (interp_binop op y' z')) (cost_SOp x y z mc) -> exec (SOp x op y z) t m l mc post | set: forall t m l mc x y y' post, map.get l y = Some y' -> - post t m (map.put l x y') - (addMetricLoads 1 - (addMetricInstructions 1 mc)) -> + post t m (map.put l x y') (cost_SSet x y mc) -> exec (SSet x y) t m l mc post | if_true: forall t m l mc cond bThen bElse post, eval_bcond l cond = Some true -> - exec bThen t m l - (addMetricLoads 2 - (addMetricInstructions 2 - (addMetricJumps 1 mc))) post -> + exec bThen t m l (cost_SIf cond mc) post -> exec (SIf cond bThen bElse) t m l mc post | if_false: forall t m l mc cond bThen bElse post, eval_bcond l cond = Some false -> - exec bElse t m l - (addMetricLoads 2 - (addMetricInstructions 2 - (addMetricJumps 1 mc))) post -> + exec bElse t m l (cost_SIf cond mc) post -> exec (SIf cond bThen bElse) t m l mc post | loop: forall t m l mc cond body1 body2 mid1 mid2 post, (* This case is carefully crafted in such a way that recursive uses of exec @@ -388,10 +472,7 @@ Module exec. (forall t' m' l' mc', mid1 t' m' l' mc' -> eval_bcond l' cond = Some false -> - post t' m' l' - (addMetricLoads 1 - (addMetricInstructions 1 - (addMetricJumps 1 mc')))) -> + post t' m' l' (cost_SLoop_false cond mc')) -> (forall t' m' l' mc', mid1 t' m' l' mc' -> eval_bcond l' cond = Some true -> @@ -399,9 +480,7 @@ Module exec. (forall t'' m'' l'' mc'', mid2 t'' m'' l'' mc'' -> exec (SLoop body1 cond body2) t'' m'' l'' - (addMetricLoads 2 - (addMetricInstructions 2 - (addMetricJumps 1 mc''))) post) -> + (cost_SLoop_true cond mc'') post) -> exec (SLoop body1 cond body2) t m l mc post | seq: forall t m l mc s1 s2 mid post, exec s1 t m l mc mid -> @@ -448,15 +527,15 @@ Module exec. Lemma loop_cps: forall body1 cond body2 t m l mc post, exec body1 t m l mc (fun t m l mc => exists b, eval_bcond l cond = Some b /\ - (b = false -> post t m l (addMetricLoads 1 (addMetricInstructions 1 (addMetricJumps 1 mc)))) /\ + (b = false -> post t m l (cost_SLoop_false cond mc)) /\ (b = true -> exec body2 t m l mc (fun t m l mc => exec (SLoop body1 cond body2) t m l - (addMetricLoads 2 (addMetricInstructions 2 (addMetricJumps 1 mc))) post))) -> + (cost_SLoop_true cond mc) post))) -> exec (SLoop body1 cond body2) t m l mc post. Proof. intros. eapply loop. 1: eapply H. all: cbv beta; intros; simp. - congruence. - - replace b with false in * by congruence. clear b. eauto. + - replace b with false in * by congruence. clear b. eauto. - replace b with true in * by congruence. clear b. eauto. - assumption. Qed. @@ -604,14 +683,16 @@ Section FlatImp2. {env_ok: map.ok env} {ext_spec_ok: ext_spec.ok ext_spec}. + Variable (isReg: varname -> bool). + Definition SimState: Type := trace * mem * locals * MetricLog. Definition SimExec(e: env)(c: stmt varname): SimState -> (SimState -> Prop) -> Prop := fun '(t, m, l, mc) post => - exec e c t m l mc (fun t' m' l' mc' => post (t', m', l', mc')). + exec isReg e c t m l mc (fun t' m' l' mc' => post (t', m', l', mc')). Lemma modVarsSound: forall e s initialT (initialSt: locals) initialM (initialMc: MetricLog) post, - exec e s initialT initialM initialSt initialMc post -> - exec e s initialT initialM initialSt initialMc + exec isReg e s initialT initialM initialSt initialMc post -> + exec isReg e s initialT initialM initialSt initialMc (fun finalT finalM finalSt _ => map.only_differ initialSt (modVars s) finalSt). Proof. induction 1; @@ -631,7 +712,7 @@ Section FlatImp2. - eapply exec.stackalloc; try eassumption. intros. eapply exec.weaken. - + eapply exec.intersect. + + eapply exec.intersect; try eassumption. * eapply H0; eassumption. * eapply H1; eassumption. + simpl. intros. simp. @@ -651,7 +732,7 @@ Section FlatImp2. + intros. simp. eauto. + intros. simp. simpl. map_solver locals_ok. + intros. simp. simpl in *. - eapply exec.intersect; [eauto|]. + eapply exec.intersect; try eassumption; [eauto|]. eapply exec.weaken. * eapply H3; eassumption. * simpl. intros. map_solver locals_ok. From da28f4aa4785e4cc14d449a1a1cdc0c21d037aad Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Wed, 4 May 2022 17:24:47 -0400 Subject: [PATCH 03/62] start fixing other files to add isReg, low-hanging fruit first --- compiler/src/compiler/FlatImp.v | 9 +++++++++ compiler/src/compiler/FlatToRiscvCommon.v | 3 ++- compiler/src/compiler/Spilling.v | 3 ++- 3 files changed, 13 insertions(+), 2 deletions(-) diff --git a/compiler/src/compiler/FlatImp.v b/compiler/src/compiler/FlatImp.v index 60972bfa5..9ece4b535 100644 --- a/compiler/src/compiler/FlatImp.v +++ b/compiler/src/compiler/FlatImp.v @@ -750,3 +750,12 @@ Section FlatImp2. Qed. End FlatImp2. + +Definition isRegZ (var : Z) : bool := + Z.leb var 31. + +Definition isRegStr (var : String.string) : bool := + String.prefix "reg_" var. + + + diff --git a/compiler/src/compiler/FlatToRiscvCommon.v b/compiler/src/compiler/FlatToRiscvCommon.v index ffed74b16..336900f79 100644 --- a/compiler/src/compiler/FlatToRiscvCommon.v +++ b/compiler/src/compiler/FlatToRiscvCommon.v @@ -289,7 +289,8 @@ Section WithParameters. pos mod 4 = 0. Local Notation stmt := (stmt Z). - + Local Notation exec := (exec isRegZ). + (* note: [e_impl_reduced] and [funnames] will shrink one function at a time each time we enter a new function body, to make sure functions cannot call themselves, while [e_impl] and [e_pos] remain the same throughout because that's mandated by diff --git a/compiler/src/compiler/Spilling.v b/compiler/src/compiler/Spilling.v index 43f9ba279..aa5a09047 100644 --- a/compiler/src/compiler/Spilling.v +++ b/compiler/src/compiler/Spilling.v @@ -20,7 +20,8 @@ Open Scope Z_scope. Section Spilling. Notation stmt := (stmt Z). - + Notation exec := (exec isRegZ). + Definition zero := 0. Definition ra := 1. Definition sp := 2. From 143d4854a3a1407a793ed3fae522a6c04e2a83a3 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Tue, 10 May 2022 12:53:31 -0400 Subject: [PATCH 04/62] fix FlatToRiscvFunctions.v, could be a little tidier though --- compiler/src/compiler/FlatToRiscvFunctions.v | 103 +++++++++++++++++-- 1 file changed, 97 insertions(+), 6 deletions(-) diff --git a/compiler/src/compiler/FlatToRiscvFunctions.v b/compiler/src/compiler/FlatToRiscvFunctions.v index e3fdea80a..c5e39c985 100644 --- a/compiler/src/compiler/FlatToRiscvFunctions.v +++ b/compiler/src/compiler/FlatToRiscvFunctions.v @@ -202,6 +202,13 @@ Section Proofs. * subst. assumption. * eauto. Qed. + + Lemma valid_FlatImp_var_isRegZ : forall x, + valid_FlatImp_var x -> isRegZ x = true. + Proof. + unfold valid_FlatImp_var, isRegZ; blia. + Qed. + Hint Resolve valid_FlatImp_var_isRegZ. Ltac run1done := apply runsToDone; @@ -212,7 +219,14 @@ Section Proofs. end; ssplit; simpl_word_exprs word_ok; match goal with | |- _ => solve_word_eq word_ok - | |- (_ <= _)%metricsL => MetricsToRiscv.solve_MetricLog + | |- (_ <= _)%metricsL => + unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, + exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, + exec.cost_SLoop_true, exec.cost_SLoop_false in *; + repeat match goal with + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + end; + MetricsToRiscv.solve_MetricLog | |- iff1 ?x ?x => reflexivity (* `exists stack_trash frame_trash, ...` from goodMachine *) | |- exists _ _, _ = _ /\ _ = _ /\ (_ * _)%sep _ => @@ -416,8 +430,8 @@ Section Proofs. simpl in *; Simp.simp; repeat (simulate'; simpl_bools; simpl); try intuition congruence. Qed. - - + Local Notation exec := (exec isRegZ). + Lemma compile_function_body_correct: forall (e_impl_full : env) m l mc (argvs : list word) (st0 : locals) (post outcome : Semantics.trace -> mem -> locals -> MetricLog -> Prop) (argnames retnames : list Z) (body : stmt Z) (program_base : word) @@ -1629,7 +1643,8 @@ Section Proofs. } inline_iff1. run1det. clear H0. (* <-- TODO this should not be needed *) run1done. - + + - idtac "Case compile_stmt_correct/SStore". inline_iff1. simpl_MetricRiscvMachine_get_set. @@ -1840,6 +1855,12 @@ Section Proofs. } run1done. cbn. + unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, + exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, + exec.cost_SLoop_true, exec.cost_SLoop_false in *. + repeat match goal with + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + end. remember (updateMetricsForLiteral v initialL_metrics) as finalMetrics; symmetry in HeqfinalMetrics; pose proof update_metrics_for_literal_bounded (width := width) as Hlit; @@ -1900,7 +1921,25 @@ Section Proofs. simpl_MetricRiscvMachine_get_set. intros. destruct_RiscvMachine mid. fwd. run1done. - + destruct cond eqn:E; + unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, + exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, + exec.cost_SLoop_true, exec.cost_SLoop_false in *; + unfold ForallVars_bcond in *; + repeat match goal with + | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H + end; + repeat match goal with + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + end; + try MetricsToRiscv.solve_MetricLog. + destruct H6p0. + repeat match goal with + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + end. + MetricsToRiscv.solve_MetricLog. + + - idtac "Case compile_stmt_correct/SIf/Else". (* execute branch instruction, which will jump over then-branch *) eapply runsToStep. @@ -1925,6 +1964,23 @@ Section Proofs. * (* at end of else-branch, i.e. also at end of if-then-else, just prove that computed post satisfies required post *) simpl. intros. destruct_RiscvMachine middle. fwd. subst. run1done. + destruct cond eqn:E; + unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, + exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, + exec.cost_SLoop_true, exec.cost_SLoop_false in *; + unfold ForallVars_bcond in *; + repeat match goal with + | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H + end; + repeat match goal with + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + end; + try MetricsToRiscv.solve_MetricLog. + destruct H6p0. + repeat match goal with + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + end. + MetricsToRiscv.solve_MetricLog. - idtac "Case compile_stmt_correct/SLoop". match goal with @@ -1997,6 +2053,24 @@ Section Proofs. } (* at end of loop, just prove that computed post satisfies required post *) simpl. intros. destruct_RiscvMachine middle. fwd. run1done. + destruct cond eqn:blah; + unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, + exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, + exec.cost_SLoop_true, exec.cost_SLoop_false in *; + unfold ForallVars_bcond in *; + repeat match goal with + | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H + end; + repeat match goal with + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + end; + try MetricsToRiscv.solve_MetricLog. + destruct H11p0. + repeat match goal with + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + end. + MetricsToRiscv.solve_MetricLog. + * (* false: done, jump over body2 *) eapply runsToStep. { eapply compile_bcond_by_inverting_correct with (l := lH') (b := false); @@ -2004,7 +2078,24 @@ Section Proofs. try safe_sidecond. } simpl_MetricRiscvMachine_get_set. - intros. destruct_RiscvMachine mid. fwd. run1done. + intros. destruct_RiscvMachine mid. fwd. run1done. + destruct cond eqn:blah; + unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, + exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, + exec.cost_SLoop_true, exec.cost_SLoop_false in *; + unfold ForallVars_bcond in *; + repeat match goal with + | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H + end; + repeat match goal with + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + end; + try MetricsToRiscv.solve_MetricLog. + destruct H11p0. + repeat match goal with + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + end. + MetricsToRiscv.solve_MetricLog. - idtac "Case compile_stmt_correct/SSeq". on hyp[(FlatImpConstraints.uses_standard_arg_regs s1); runsTo] From 97206648d9b34df461a83bf43fa265e23dd5f41f Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Mon, 16 May 2022 12:49:12 -0400 Subject: [PATCH 05/62] progress on RegAlloc.v, not done yet --- compiler/src/compiler/RegAlloc.v | 100 ++++++++++++++++++++----------- 1 file changed, 65 insertions(+), 35 deletions(-) diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index ddd1f0f2c..657c75d36 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -456,9 +456,15 @@ Definition assert(b: bool)(els: result unit): result unit := if b then Success t Definition mapping_eqb: srcvar * impvar -> srcvar * impvar -> bool := fun '(x, x') '(y, y') => andb (String.eqb x y) (Z.eqb x' y'). +Definition check_regs (x: srcvar) (x': impvar) : bool := + negb (andb (isRegStr x) (negb (isRegZ x'))). + Definition assert_in(y: srcvar)(y': impvar)(m: list (srcvar * impvar)): result unit := match List.find (mapping_eqb (y, y')) m with - | Some _ => Success tt + | Some _ => if check_regs y y' then Success tt else + error:("The register allocator found a mapping of source register variable" y + "to target stack variable" y' + "but source register variables must be in registers in the target.") | None => error:("The register allocator replaced source variable" y "by target variable" y' "but when the checker encountered this pair," @@ -474,7 +480,12 @@ Definition assert_ins(args: list srcvar)(args': list impvar)(m: list (srcvar * i (List.combine args args')) error:("Register allocation checker got a source variable list" args "and a target variable list" args' - "that are incompatible with its current mapping of source to target variables" m). + "that are incompatible with its current mapping of source to target variables" m);; + assert (List.forallb (fun '(x, x') => if check_regs x x' then true else false) + (List.combine args args')) + error:("Register allocation checker got a source variable list" args + "and a target variable list" args' + "in which at least one source register variable is on the stack in the target."). Definition check_bcond(m: list (srcvar * impvar))(c: bcond)(c': bcond'): result unit := match c, c' with @@ -488,13 +499,20 @@ Definition check_bcond(m: list (srcvar * impvar))(c: bcond)(c': bcond'): result | _, _ => error:("Register allocation checker cannot match" c "and" c') end. -Definition assignment(m: list (srcvar * impvar))(x: srcvar)(x': impvar): list (srcvar * impvar) := - (x, x') :: (remove_by_snd Z.eqb x' (remove_by_fst String.eqb x m)). - +Definition assignment(m: list (srcvar * impvar))(x: srcvar)(x': impvar): result (list (srcvar * impvar)) := + if check_regs x x' then + Success ((x, x') :: (remove_by_snd Z.eqb x' (remove_by_fst String.eqb x m))) + else + error:("Register allocation checker got an assignment of source register variable" x + "to target stack variable" x' + "but source register variables must be in registers in the target."). + Fixpoint assignments(m: list (srcvar * impvar))(xs: list srcvar)(xs': list impvar): result (list (srcvar * impvar)) := match xs, xs' with - | x :: xs0, x' :: xs0' => assignments (assignment m x x') xs0 xs0' + | x :: xs0, x' :: xs0' => + a <- assignment m x x';; + assignments a xs0 xs0' | nil, nil => Success m | _, _ => error:("Register allocator checker got variable lists of different length") end. @@ -530,7 +548,8 @@ Fixpoint check(m: list (srcvar * impvar))(s: stmt)(s': stmt'){struct s}: result assert_in y y' m;; assert (Z.eqb ofs ofs') err;; assert (access_size_beq sz sz') err;; - Success (assignment m x x') + a <- assignment m x x';; + Success a | SStore sz x y ofs, SStore sz' x' y' ofs' => assert_in x x' m;; assert_in y y' m;; @@ -543,20 +562,26 @@ Fixpoint check(m: list (srcvar * impvar))(s: stmt)(s': stmt'){struct s}: result assert (negb (Z.eqb x' y')) err;; assert (access_size_beq sz sz') err;; assert (List.list_eqb Byte.eqb bs bs') err;; - Success (assignment m x x') + a <- (assignment m x x');; + Success a | SStackalloc x n body, SStackalloc x' n' body' => assert (Z.eqb n n') err;; - check (assignment m x x') body body' + a <- assignment m x x';; + check a body body' | SLit x z, SLit x' z' => assert (Z.eqb z z') err;; - Success (assignment m x x') + a <- assignment m x x';; + Success a | SOp x op y z, SOp x' op' y' z' => assert_in y y' m;; assert_in z z' m;; assert (bopname_beq op op') err;; - Success (assignment m x x') + a <- assignment m x x';; + Success a | SSet x y, SSet x' y' => - assert_in y y' m;; Success (assignment m x x') + assert_in y y' m;; + a <- assignment m x x';; + Success a | SIf c s1 s2, SIf c' s1' s2' => check_bcond m c c';; m1 <- check m s1 s1';; @@ -633,8 +658,8 @@ Lemma extends_cons: forall a l1 l2, Proof. unfold extends, assert_in. simpl. intros. simp. destruct_one_match_hyp. 1: reflexivity. - epose proof (H _ _ _) as A. destruct_one_match_hyp. 2: discriminate. rewrite E1. reflexivity. - Unshelve. rewrite E. reflexivity. + epose proof (H _ _ _) as A. destruct_one_match_hyp. 2: discriminate. rewrite E2. reflexivity. + Unshelve. rewrite E. rewrite E0. reflexivity. Qed. Lemma extends_cons_l: forall a l, @@ -650,13 +675,14 @@ Lemma extends_cons_r: forall a l1 l2, extends l1 l2 -> extends l1 (a :: l2). Proof. - unfold extends, assert_in. simpl. intros. simp. + unfold extends, assert_in. simpl. intros. simp. destruct_one_match_hyp. 2: { - eapply H0. rewrite E. reflexivity. + specialize (H0 x x'). rewrite E in H0. rewrite E0 in H0. + eapply H0. reflexivity. } destruct_one_match. 1: reflexivity. - eapply find_none in E1. 2: eassumption. - simpl in E1. exfalso. congruence. + eapply find_none in E2. 2: eassumption. + simpl in E2. exfalso. congruence. Qed. Lemma extends_intersect_l: forall l1 l2, @@ -840,25 +866,26 @@ Section RegAlloc. induction ys; intros. - unfold assert_ins in *. cbn in *. simp. destruct ys'. 2: discriminate. reflexivity. - cbn in *. unfold assert_ins, assert in H0. simp. - autoforward with typeclass_instances in E3. destruct ys' as [|a' ys']. 1: discriminate. - inversion E3. clear E3. - cbn in *. simp. simpl in E2. + autoforward with typeclass_instances in E5. destruct ys' as [|a' ys']. 1: discriminate. + inversion E5. clear E5. + cbn in *. simp. simpl in E3. erewrite states_compat_get; try eassumption. 2: { - unfold assert_in. unfold mapping_eqb. rewrite E1. reflexivity. + unfold assert_in. unfold mapping_eqb. rewrite E2. rewrite E1. reflexivity. } unfold map.getmany_of_list in *. erewrite IHys; eauto. unfold assert_ins. rewrite H1. rewrite Nat.eqb_refl. simpl. - unfold assert. rewrite E2. reflexivity. + unfold assert. rewrite E3. cbn in E4. rewrite E4. reflexivity. Qed. - Lemma states_compat_put: forall lH corresp lL x x' v, + Lemma states_compat_put: forall lH corresp lL x x' v m, + assignment corresp x x' = Success m -> states_compat lH corresp lL -> - states_compat (map.put lH x v) (assignment corresp x x') (map.put lL x' v). + states_compat (map.put lH x v) m (map.put lL x' v). Proof. intros. unfold states_compat in *. intros k k'. intros. - rewrite map.get_put_dec. rewrite map.get_put_dec in H1. - unfold assert_in, assignment in H0. simp. simpl in E. + rewrite map.get_put_dec. rewrite map.get_put_dec in H2. + unfold assert_in, assignment in H, H0, H1. simp. simpl in E. rewrite String.eqb_sym, Z.eqb_sym in E. destr (Z.eqb x' k'). - destr (String.eqb x k). @@ -877,10 +904,10 @@ Section RegAlloc. eapply In_remove_by_fst in E. destruct E. destr (String.eqb x k). + exfalso. congruence. - + eapply H. 2: eassumption. unfold assert_in. - destruct_one_match. 1: reflexivity. - eapply find_none in E1. 2: eassumption. - simpl in E1. rewrite String.eqb_refl, Z.eqb_refl in E1. discriminate. + + eapply H0. 2: eassumption. unfold assert_in. + destruct_one_match. 1: rewrite E0; trivial. + eapply find_none in E3. 2: eassumption. + simpl in E3. rewrite String.eqb_refl, Z.eqb_refl in E3. discriminate. Qed. Lemma putmany_of_list_zip_states_compat: forall binds binds' resvals lL lH l' corresp corresp', @@ -897,7 +924,7 @@ Section RegAlloc. - simpl in *. simp. specialize IHbinds with (1 := H). rename l' into lH'. - edestruct IHbinds as (lL' & P & C). 1: eassumption. 1: eapply states_compat_put. 1: eassumption. + edestruct IHbinds as (lL' & P & C). 1: eassumption. 1: eapply states_compat_put. 1: eassumption. 1: eassumption. simpl. rewrite P. eauto. Qed. @@ -906,7 +933,8 @@ Section RegAlloc. Proof. induction xs; intros; destruct xs'; try discriminate. - reflexivity. - - simpl in *. f_equal. eapply IHxs. eassumption. + - simpl in *. f_equal. destruct_one_match_hyp; try discriminate. + specialize (IHxs xs' a0 m2). eauto. Qed. Lemma assert_ins_same_length: forall xs xs' m u, @@ -921,11 +949,11 @@ Section RegAlloc. Lemma checker_correct: forall (e: srcEnv) (e': impEnv) s t m lH mc post, check_funcs e e' = Success tt -> - exec e s t m lH mc post -> + exec isRegStr e s t m lH mc post -> forall lL corresp corresp' s', check corresp s s' = Success corresp' -> states_compat lH (precond corresp s s') lL -> - exec e' s' t m lL mc (fun t' m' lL' mc' => + exec isRegZ e' s' t m lL mc (fun t' m' lL' mc' => exists lH', states_compat lH' corresp' lL' /\ post t' m' lH' mc'). Proof. induction 2; intros; @@ -977,6 +1005,8 @@ Section RegAlloc. * exact L4. * eexists. split. 2: eassumption. exact SC. - (* Case exec.load *) + econstructor; eauto 10 with checker_hints. + eauto 10 with checker_hints. - (* Case exec.store *) eauto 10 with checker_hints. From d157885aa68ee36ec6808e0b856d37106e21b113 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Tue, 14 Jun 2022 18:32:24 -0400 Subject: [PATCH 06/62] fixed but slow version of regalloc proof; TODO improve performance --- compiler/src/compiler/FlatToRiscvMetric.v | 4 +- compiler/src/compiler/FlattenExpr.v | 22 +- compiler/src/compiler/RegAlloc.v | 313 +++++++++++++++++++--- 3 files changed, 284 insertions(+), 55 deletions(-) diff --git a/compiler/src/compiler/FlatToRiscvMetric.v b/compiler/src/compiler/FlatToRiscvMetric.v index ced38fd4d..e2fde7171 100644 --- a/compiler/src/compiler/FlatToRiscvMetric.v +++ b/compiler/src/compiler/FlatToRiscvMetric.v @@ -91,7 +91,7 @@ Section Proofs. ext_spec t mGive action argvals outcome -> False. Hypothesis stackalloc_always_0: forall x n body t m (l: locals) mc post, - FlatImp.exec map.empty (SStackalloc x n body) t m l mc post -> n = 0. + FlatImp.exec isRegZ map.empty (SStackalloc x n body) t m l mc post -> n = 0. Hypothesis sp_always_set: forall l: locals, map.get l RegisterNames.sp = Some (word.of_Z 42). @@ -101,7 +101,7 @@ Section Proofs. Lemma compile_stmt_correct: forall (s: stmt Z) t initialMH initialRegsH postH initialMetricsH, - FlatImp.exec map.empty s t initialMH initialRegsH initialMetricsH postH -> + FlatImp.exec isRegZ map.empty s t initialMH initialRegsH initialMetricsH postH -> forall R Rexec (initialL: RiscvMachineL) insts pos, compile_stmt iset compile_ext_call map.empty pos 12345678 s = insts -> stmt_not_too_big s -> diff --git a/compiler/src/compiler/FlattenExpr.v b/compiler/src/compiler/FlattenExpr.v index 5b38785e9..81e6126b2 100644 --- a/compiler/src/compiler/FlattenExpr.v +++ b/compiler/src/compiler/FlattenExpr.v @@ -337,13 +337,15 @@ Section FlattenExpr1. simpl (disjoint _ _) in *; map_solver locals_ok. + Local Notation exec := (FlatImp.exec FlatImp.isRegStr). + Lemma seq_with_modVars: forall env t m (l: locals) mc s1 s2 mid post, - FlatImp.exec env s1 t m l mc mid -> + exec env s1 t m l mc mid -> (forall t' m' l' mc', mid t' m' l' mc' -> map.only_differ l (FlatImp.modVars s1) l' -> - FlatImp.exec env s2 t' m' l' mc' post) -> - FlatImp.exec env (FlatImp.SSeq s1 s2) t m l mc post. + exec env s2 t' m' l' mc' post) -> + exec env (FlatImp.SSeq s1 s2) t m l mc post. Proof. intros *. intros E1 E2. eapply @FlatImp.exec.seq. - eapply FlatImp.exec.intersect. @@ -365,7 +367,7 @@ Section FlattenExpr1. map.undef_on initialH (allFreshVars ngs1) -> disjoint (union (ExprImp.allVars_expr e) (of_option oResVar)) (allFreshVars ngs1) -> eval_expr initialM initialH e initialMcH = Some (res, finalMcH) -> - FlatImp.exec fenv s t initialM initialL initialMcL (fun t' finalM finalL finalMcL => + exec fenv s t initialM initialL initialMcL (fun t' finalM finalL finalMcL => t' = t /\ finalM = initialM /\ map.get finalL resVar = Some res /\ (finalMcL - initialMcL <= finalMcH - initialMcH)%metricsH). Proof. @@ -498,7 +500,7 @@ Section FlattenExpr1. map.undef_on lH (allFreshVars ngs1) -> disjoint (union (ExprImp.allVars_expr e) (of_option oResVar)) (allFreshVars ngs1) -> eval_expr m lH e initialMcH = Some (res, finalMcH) -> - FlatImp.exec fenv s t m lL initialMcL (fun t' m' lL' finalMcL => + exec fenv s t m lL initialMcL (fun t' m' lL' finalMcL => map.only_differ lL (FlatImp.modVars s) lL' /\ t' = t /\ m' = m /\ map.get lL' resVar = Some res /\ (finalMcL - initialMcL <= finalMcH - initialMcH)%metricsH). @@ -517,7 +519,7 @@ Section FlattenExpr1. disjoint (ExprImp.allVars_exprs es) (allFreshVars ngs1) -> evaluate_call_args_log m lH es initialMcH = Some (resVals, finalMcH) -> (* List.option_all (List.map (eval_expr m lH) es) = Some resVals -> *) - FlatImp.exec fenv s t m lL initialMcL (fun t' m' lL' finalMcL => + exec fenv s t m lL initialMcL (fun t' m' lL' finalMcL => t' = t /\ m' = m /\ map.getmany_of_list lL' resVars = Some resVals /\ map.only_differ lL (FlatImp.modVars s) lL' /\ @@ -590,7 +592,7 @@ Section FlattenExpr1. map.undef_on initialH (allFreshVars ngs1) -> disjoint (ExprImp.allVars_expr e) (allFreshVars ngs1) -> eval_expr initialM initialH e initialMcH = Some (res, finalMcH) -> - FlatImp.exec fenv s t initialM initialL initialMcL (fun t' finalM finalL finalMcL => + exec fenv s t initialM initialL initialMcL (fun t' finalM finalL finalMcL => t' = t /\ finalM = initialM /\ FlatImp.eval_bcond finalL resCond = Some (negb (word.eqb res (word.of_Z 0))) /\ (finalMcL - initialMcL <= finalMcH - initialMcH)%metricsH). @@ -627,7 +629,7 @@ Section FlattenExpr1. map.undef_on initialH (allFreshVars ngs1) -> disjoint (ExprImp.allVars_expr e) (allFreshVars ngs1) -> eval_expr initialM initialH e initialMcH = Some (res, finalMcH) -> - FlatImp.exec fenv s t initialM initialL initialMcL (fun t' finalM finalL finalMcL => + exec fenv s t initialM initialL initialMcL (fun t' finalM finalL finalMcL => (t' = t /\ finalM = initialM /\ FlatImp.eval_bcond finalL resCond = Some (negb (word.eqb res (word.of_Z 0))) /\ (finalMcL - initialMcL <= finalMcH - initialMcH)%metricsH) /\ @@ -684,7 +686,7 @@ Section FlattenExpr1. map.extends lL lH -> map.undef_on lH (allFreshVars ngs) -> disjoint (ExprImp.allVars_cmd sH) (allFreshVars ngs) -> - FlatImp.exec eL sL t m lL mcL (fun t' m' lL' mcL' => exists lH' mcH', + exec eL sL t m lL mcL (fun t' m' lL' mcL' => exists lH' mcH', post t' m' lH' mcH' /\ (* <-- put first so that eassumption will instantiate lH' correctly *) map.extends lL' lH' /\ (* this one is a property purely about ExprImp (it's the conclusion of @@ -924,7 +926,7 @@ Section FlattenExpr1. flatten_functions eH = Success eL -> ExprImp2FlatImp sH = sL -> Semantics.exec eH sH t m map.empty mc post -> - FlatImp.exec eL sL t m lL mc (fun t' m' lL' mcL' => exists lH' mcH', + exec eL sL t m lL mc (fun t' m' lL' mcL' => exists lH' mcH', post t' m' lH' mcH' /\ map.extends lL' lH' /\ (mcL' - mc <= mcH' - mc)%metricsH). diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index 657c75d36..03cd38930 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -943,18 +943,186 @@ Section RegAlloc. unfold assert_ins, assert. intros. simp. apply Nat.eqb_eq. assumption. Qed. + Ltac unfold_metrics := + unfold + MetricLogging.withInstructions, MetricLogging.withLoads, + MetricLogging.withStores, MetricLogging.withJumps, + MetricLogging.addMetricInstructions, MetricLogging.addMetricLoads, + MetricLogging.addMetricStores, MetricLogging.addMetricJumps, + MetricLogging.subMetricInstructions, MetricLogging.subMetricLoads, + MetricLogging.subMetricStores, MetricLogging.subMetricJumps, + MetricLogging.metricsOp, MetricLogging.metricSub, MetricLogging.metricsSub, + MetricLogging.metricLeq, MetricLogging.metricsLeq + in *. + + Opaque isRegStr. + Opaque isRegZ. + + Lemma check_regs_cost_SLoad: forall x x' a a' mc mc', + check_regs a a' = true -> + check_regs x x' = true -> + (MetricLogging.metricsLeq + (MetricLogging.metricsSub (exec.cost_SLoad isRegZ x' a' mc') mc') + (MetricLogging.metricsSub (exec.cost_SLoad isRegStr x a mc) mc)). + Proof. + intros; unfold check_regs in *; cbn in *; + destr (isRegStr x); destr (isRegStr a); destr (isRegZ x'); destr (isRegZ a'); + try discriminate; + unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; blia. + Qed. + + Lemma check_regs_cost_SStore: forall x x' a a' mc mc', + check_regs a a' = true -> + check_regs x x' = true -> + (MetricLogging.metricsLeq + (MetricLogging.metricsSub (exec.cost_SStore isRegZ x' a' mc') mc') + (MetricLogging.metricsSub (exec.cost_SStore isRegStr x a mc) mc)). + Proof. + intros; unfold check_regs in *; cbn in *; + destr (isRegStr x); destr (isRegStr a); destr (isRegZ x'); destr (isRegZ a'); + try discriminate; + unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; blia. + Qed. + + Lemma check_regs_cost_SInlinetable: forall x x' a a' mc mc', + check_regs a a' = true -> + check_regs x x' = true -> + (MetricLogging.metricsLeq + (MetricLogging.metricsSub (exec.cost_SInlinetable isRegZ x' a' mc') mc') + (MetricLogging.metricsSub (exec.cost_SInlinetable isRegStr x a mc) mc)). + Proof. + intros; unfold check_regs in *; cbn in *; + destr (isRegStr x); destr (isRegStr a); destr (isRegZ x'); destr (isRegZ a'); + try discriminate; + unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; blia. + Qed. + + Lemma check_regs_cost_SStackalloc: forall x x' mcL mcL' mcH mcH', + check_regs x x' = true -> + (MetricLogging.metricsLeq + (MetricLogging.metricsSub mcL' (exec.cost_SStackalloc isRegZ x' mcL)) + (MetricLogging.metricsSub mcH' (exec.cost_SStackalloc isRegStr x mcH))) -> + (MetricLogging.metricsLeq + (MetricLogging.metricsSub mcL' mcL) + (MetricLogging.metricsSub mcH' mcH)). + Proof. + intros; unfold check_regs in *; cbn in *; unfold exec.cost_SStackalloc in *; + destr (isRegStr x); destr (isRegZ x'); + try discriminate; + unfold_metrics; cbn; repeat split; destruct mcL; destruct mcH; destruct mcL'; destruct mcH'; cbn in *; unfold_metrics; cbn in *; try blia. + Qed. + + Lemma check_regs_cost_SLit: forall x x' mc mc', + check_regs x x' = true -> + (MetricLogging.metricsLeq + (MetricLogging.metricsSub (exec.cost_SLit isRegZ x' mc') mc') + (MetricLogging.metricsSub (exec.cost_SLit isRegStr x mc) mc)). + Proof. + intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLit in *; + destr (isRegStr x); destr (isRegZ x'); + try discriminate; + unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; try blia. + Qed. + + Lemma check_regs_cost_SOp: forall x x' y y' z z' mc mc', + check_regs x x' = true -> + check_regs y y' = true -> + check_regs z z' = true -> + (MetricLogging.metricsLeq + (MetricLogging.metricsSub (exec.cost_SOp isRegZ x' y' z' mc') mc') + (MetricLogging.metricsSub (exec.cost_SOp isRegStr x y z mc) mc)). + Proof. + intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLit in *; + destr (isRegStr x); destr (isRegZ x'); destr (isRegStr y); destr (isRegZ y'); destr (isRegStr z); destr (isRegZ z'); + try discriminate; + unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; try blia. + Qed. + + Lemma check_regs_cost_SSet: forall x x' y y' mc mc', + check_regs x x' = true -> + check_regs y y' = true -> + (MetricLogging.metricsLeq + (MetricLogging.metricsSub (exec.cost_SSet isRegZ x' y' mc') mc') + (MetricLogging.metricsSub (exec.cost_SSet isRegStr x y mc) mc)). + Proof. + intros; unfold check_regs in *; cbn in *; + destr (isRegStr x); destr (isRegStr y); destr (isRegZ x'); destr (isRegZ y'); + try discriminate; + unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; blia. + Qed. + + Ltac discr_match_success := + match goal with + | H: match ?expr with _ => _ end = Success _ |- _ => destr expr; discriminate + end. + + Lemma check_regs_cost_SIf: + forall (cond : bcond) (bThen bElse : stmt) (corresp : list (string * Z)) + (cond0 : bcond') (s'1 s'2 : stmt') (a0 a1 : list (string * Z)) + (mc mcL : MetricLogging.MetricLog), + check_bcond corresp cond cond0 = Success tt -> + check corresp bThen s'1 = Success a0 -> + check corresp bElse s'2 = Success a1 -> + forall mc' mcH' : MetricLogging.MetricLog, + MetricLogging.metricsLeq + (MetricLogging.metricsSub mc' (exec.cost_SIf isRegZ cond0 mcL)) + (MetricLogging.metricsSub mcH' (exec.cost_SIf isRegStr cond mc)) -> + MetricLogging.metricsLeq (MetricLogging.metricsSub mc' mcL) + (MetricLogging.metricsSub mcH' mc). + Proof. + intros. + repeat (unfold check_bcond, assert_in, assignment in *; simp). + intros; unfold check_regs in *; cbn in *; unfold exec.cost_SIf in *; + destr cond; destr cond0; destr (isRegStr x); destr (isRegZ x0); try (destr (isRegStr y)); try (destr (isRegZ y0)); + try discriminate; + unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; simp; + destr mcL; destr mc'; destr mc; destr mcH'; try blia. + all: try discr_match_success. + all: cbn in *; simp; try blia. + Qed. + + Lemma check_regs_cost_SLoop_false: + forall (cond : bcond) (body1 body2 : stmt) (corresp' : list (string * Z)) (s'1 : stmt') + (cond0 : bcond') (s'2 : stmt') (mc mcL : MetricLogging.MetricLog) (a : list (string * Z)), + check a body1 s'1 = Success corresp' -> + check_bcond corresp' cond cond0 = Success tt -> + forall a2 : list (string * Z), + check corresp' body2 s'2 = Success a2 -> + forall mc' mcH' : MetricLogging.MetricLog, + MetricLogging.metricsLeq (MetricLogging.metricsSub mc' mcL) (MetricLogging.metricsSub mcH' mc) -> + MetricLogging.metricsLeq (MetricLogging.metricsSub (exec.cost_SLoop_false isRegZ cond0 mc') mcL) + (MetricLogging.metricsSub (exec.cost_SLoop_false isRegStr cond mcH') mc). + Proof. + intros. + repeat (unfold check_bcond, assert_in, assignment in *; simp). + intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLoop_false in *; + destr cond; destr cond0; destr (isRegStr x); destr (isRegZ x0); try (destr (isRegStr y)); try (destr (isRegZ y0)); + try discriminate; + unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; simp; + destr mcL; destr mc'; destr mc; destr mcH'; try blia. + all: try discr_match_success. + all: cbn in *; simp; try blia. + Qed. + Hint Constructors exec.exec : checker_hints. Hint Resolve states_compat_get : checker_hints. Hint Resolve states_compat_put : checker_hints. - - Lemma checker_correct: forall (e: srcEnv) (e': impEnv) s t m lH mc post, + Hint Resolve + check_regs_cost_SLoad check_regs_cost_SStore check_regs_cost_SInlinetable + check_regs_cost_SStackalloc check_regs_cost_SLit check_regs_cost_SOp + check_regs_cost_SSet check_regs_cost_SIf check_regs_cost_SLoop_false + : checker_hints. + + Lemma checker_correct: forall (e: srcEnv) (e': impEnv) s t m lH mcH post, check_funcs e e' = Success tt -> - exec isRegStr e s t m lH mc post -> - forall lL corresp corresp' s', + exec isRegStr e s t m lH mcH post -> + forall lL corresp corresp' s' mcL, check corresp s s' = Success corresp' -> states_compat lH (precond corresp s s') lL -> - exec isRegZ e' s' t m lL mc (fun t' m' lL' mc' => - exists lH', states_compat lH' corresp' lL' /\ post t' m' lH' mc'). + exec isRegZ e' s' t m lL mcL (fun t' m' lL' mcL' => + exists lH' mcH', states_compat lH' corresp' lL' /\ + MetricLogging.metricsLeq (MetricLogging.metricsSub mcL' mcL) (MetricLogging.metricsSub mcH' mcH) /\ + post t' m' lH' mcH'). Proof. induction 2; intros; match goal with @@ -979,6 +1147,8 @@ Section RegAlloc. intros. edestruct H3 as (l' & P & F). 1: eassumption. eapply putmany_of_list_zip_states_compat in P. 2-3: eassumption. destruct P as (lL' & P & SC). eexists. split. 1: eassumption. intros. eauto. + repeat eexists; eauto. + all: repeat (unfold_metrics; cbn; try blia). - (* Case exec.call *) rename binds0 into binds', args0 into args'. unfold check_funcs in H. @@ -1003,41 +1173,58 @@ Section RegAlloc. do 2 eexists. ssplit. * eapply states_compat_getmany; eassumption. * exact L4. - * eexists. split. 2: eassumption. exact SC. + * repeat eexists. 6: eassumption. 1: exact SC. + all: repeat (unfold_metrics; cbn in *; try blia). - (* Case exec.load *) - econstructor; eauto 10 with checker_hints. - - eauto 10 with checker_hints. + repeat econstructor; eauto 10 with checker_hints. + all: unfold assert_in, assignment in *; simp. + all: eapply check_regs_cost_SLoad; eauto. - (* Case exec.store *) - eauto 10 with checker_hints. + repeat econstructor; eauto 10 with checker_hints. + all: unfold assert_in, assignment in *; simp. + all: eapply check_regs_cost_SStore; eauto. - (* Case exec.inlinetable *) - eauto 10 with checker_hints. - - (* Case exec.stackalloc *) + repeat econstructor; eauto 10 with checker_hints. + all: unfold assert_in, assignment in *; simp. + all: eapply check_regs_cost_SInlinetable; eauto. + - (* Case exec.stackalloc *) eapply exec.stackalloc. 1: assumption. intros. eapply exec.weaken. + eapply H2; try eassumption. - eapply states_compat_precond. eapply states_compat_put. assumption. - + cbv beta. intros. simp. eauto 10 with checker_hints. + eapply states_compat_precond. eapply states_compat_put. all: eassumption. + + cbv beta. intros. simp. + eexists. eexists. repeat (split; try eassumption). + eexists. eexists. repeat (split; try eassumption). + all: unfold assert_in, assignment in *; simp. + all: eapply check_regs_cost_SStackalloc; eauto. - (* Case exec.lit *) - eauto 10 with checker_hints. + repeat econstructor; eauto 10 with checker_hints. + all: unfold assert_in, assignment in *; simp. + all: eapply check_regs_cost_SLit; eauto. - (* Case exec.op *) - eauto 10 with checker_hints. + repeat econstructor; eauto 10 with checker_hints. + all: unfold assert_in, assignment in *; simp. + all: eapply check_regs_cost_SOp; eauto. - (* Case exec.set *) - eauto 10 with checker_hints. + repeat econstructor; eauto 10 with checker_hints. + all: unfold assert_in, assignment in *; simp. + all: eapply check_regs_cost_SSet; eauto. - (* Case exec.if_true *) eapply exec.if_true. 1: eauto using states_compat_eval_bcond. eapply exec.weaken. + eapply IHexec. 1: eassumption. eapply states_compat_precond. eassumption. - + cbv beta. intros. simp. eexists. split. 2: eassumption. - eapply states_compat_extends. 2: eassumption. eapply extends_intersect_l. + + cbv beta. intros. simp. eexists. eexists. split. 2: split. 3: eassumption. + 1: eapply states_compat_extends. 2: eassumption. 1: eapply extends_intersect_l. + eapply check_regs_cost_SIf; eauto. - (* Case exec.if_false *) eapply exec.if_false. 1: eauto using states_compat_eval_bcond. eapply exec.weaken. + eapply IHexec. 1: eassumption. eapply states_compat_precond. eassumption. - + cbv beta. intros. simp. eexists. split. 2: eassumption. - eapply states_compat_extends. 2: eassumption. eapply extends_intersect_r. + + cbv beta. intros. simp. eexists. eexists. split. 2: split. 3: eassumption. + 1: eapply states_compat_extends. 2: eassumption. 1: eapply extends_intersect_r. + eapply check_regs_cost_SIf; eauto. - (* Case exec.loop *) rename H4 into IH2, IHexec into IH1, H6 into IH12. match goal with @@ -1045,34 +1232,74 @@ Section RegAlloc. end. pose proof SC as SC0. unfold loop_inv in SC. - rewrite E in SC. - eapply exec.loop. + rewrite E in SC. + eapply exec.loop + with + (* (mid3 := (fun (t' : Semantics.trace) (m' : mem) (lL' : impLocals) (mcL' : MetricLogging.MetricLog) => *) + (* exists (lH' : srcLocals) (mcH' : MetricLogging.MetricLog), *) + (* states_compat lH' corresp' lL' /\ *) + (* (MetricLogging.metricsLeq (MetricLogging.metricsSub mcL' mcL) (MetricLogging.metricsSub mcH' mc)) /\ *) + (* mid1 t' m' lH' mcH')) *) + (mid4 := (fun (t'0 : Semantics.trace) (m'0 : mem) (lL' : impLocals) (mcL' : MetricLogging.MetricLog) => + exists (lH' : srcLocals) (mcH' : MetricLogging.MetricLog), + states_compat lH' a2 lL' /\ + (exists mcHmid mcLmid, + MetricLogging.metricsLeq (MetricLogging.metricsSub mcLmid mcL) (MetricLogging.metricsSub mcHmid mc) /\ + MetricLogging.metricsLeq (MetricLogging.metricsSub mcL' mcLmid) (MetricLogging.metricsSub mcH' mcHmid)) /\ + mid2 t'0 m'0 lH' mcH')). + eapply IH1. 1: eassumption. eapply states_compat_precond. exact SC. + cbv beta. intros. simp. eauto using states_compat_eval_bcond_None. - + cbv beta. intros. simp. eexists. split. 2: eauto using states_compat_eval_bcond_bw. assumption. - + cbv beta. intros. simp. eapply IH2; eauto using states_compat_eval_bcond_bw. - eapply states_compat_precond. assumption. - + cbv beta. intros. simp. eapply IH12. 1: eassumption. 1: eassumption. - eapply states_compat_extends. 2: eassumption. - pose proof defuel_loop_inv as P. - specialize P with (2 := E0). - specialize P with (2 := E2). - specialize (P corresp). - unfold loop_inv in P|-*. - rewrite E in P. rewrite E. - specialize (P eq_refl). - rewrite P. - eapply extends_intersect_r. + + cbv beta. intros. simp. eexists. eexists. (* exists (exec.cost_SLoop_false isRegStr cond mcH'). *) + split. 2: split. 3: eauto using states_compat_eval_bcond_bw. 1: assumption. + eapply check_regs_cost_SLoop_false. 1: apply E0. all: eauto. + + cbv beta. intros. simp. eapply exec.weaken. 1: eapply IH2; eauto using states_compat_eval_bcond_bw. + 1: eapply states_compat_precond; eassumption. + cbv beta. intros. simp. eexists. eexists. split. 2: split. 1,3: eauto. + exists mcH'. exists mc'. + split; eauto. + (* repeat (unfold check_bcond, assert_in, assignment in *; simp). *) + (* intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLoop_true in *; *) + (* destr cond; destr cond0; destr (isRegStr x); destr (isRegZ x0); try (destr (isRegStr y)); try (destr (isRegZ y0)); *) + (* try discriminate; *) + (* unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; simp; *) + (* destr mcL; destr mc'; destr mc; destr mcH'; destr mc'0; destr mcH'0; try blia. *) + (* all: try discr_match_success. *) + (* all: cbn in *; simp; try blia. *) + + cbv beta. intros. simp. eapply exec.weaken. 1: eapply IH12. 1: eassumption. 1: eassumption. + * eapply states_compat_extends. 2: eassumption. + pose proof defuel_loop_inv as P. + specialize P with (2 := E0). + specialize P with (2 := E2). + specialize (P corresp). + unfold loop_inv in P|-*. + rewrite E in P. rewrite E. + specialize (P eq_refl). + rewrite P. + eapply extends_intersect_r. + * cbv beta. intros. simp. eexists. eexists. split. 2: split. 1: eauto. 2: eauto. + intros. + repeat (unfold check_bcond, assert_in, assignment in *; simp). + clear -E0 E1 E2 H4p1p0 H4p1p1 H4p1 H4p3. + intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLoop_true in *; + destr cond; destr cond0; destr (isRegStr x); destr (isRegZ x0); try (destr (isRegStr y)); try (destr (isRegZ y0)); + try discriminate; try discr_match_success; + unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; simp. + all: destr mc''; destr mcH'; destr mcHmid; destr mcLmid; destr mc'; destr mcH'0; destr mc; destr mcL. + all: try discr_match_success. + all: cbn in *; simp; try blia. - (* Case exec.seq *) rename H2 into IH2, IHexec into IH1. eapply exec.seq. + eapply IH1. 1: eassumption. eapply states_compat_precond. assumption. - + cbv beta. intros. simp. - eapply IH2. 1: eassumption. 1: eassumption. - eapply states_compat_precond. assumption. - - (* Case exec.skip *) - eapply exec.skip. eauto. + + cbv beta. intros t' m' l' mcblah ?. simp. + eapply IH2 in H2p2. 2,3: eauto using states_compat_precond. + eapply exec.weaken; eauto. + cbv beta. intros. simp. exists lH'0. exists mcH'0. split. 2:split. 1,3: eauto. + destr mc; destr mcL; unfold_metrics; cbn in *; unfold_metrics; cbn in *; blia. + - (* case exec.skip *) + eapply exec.skip. eexists. exists mc. split. 2: split. 1,3: eauto. + destr mc; destr mcL; unfold_metrics; cbn in *; unfold_metrics; cbn in *; blia. Qed. End RegAlloc. From 803a13780f9739c6738059ae41c1ae64b56f23af Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Tue, 21 Jun 2022 14:06:24 -0700 Subject: [PATCH 07/62] perf improvements for RegAlloc.v --- compiler/src/compiler/RegAlloc.v | 28 +++++++--------------------- 1 file changed, 7 insertions(+), 21 deletions(-) diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index 03cd38930..ac6c27665 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -1233,13 +1233,7 @@ Section RegAlloc. pose proof SC as SC0. unfold loop_inv in SC. rewrite E in SC. - eapply exec.loop - with - (* (mid3 := (fun (t' : Semantics.trace) (m' : mem) (lL' : impLocals) (mcL' : MetricLogging.MetricLog) => *) - (* exists (lH' : srcLocals) (mcH' : MetricLogging.MetricLog), *) - (* states_compat lH' corresp' lL' /\ *) - (* (MetricLogging.metricsLeq (MetricLogging.metricsSub mcL' mcL) (MetricLogging.metricsSub mcH' mc)) /\ *) - (* mid1 t' m' lH' mcH')) *) + eapply exec.loop with (mid4 := (fun (t'0 : Semantics.trace) (m'0 : mem) (lL' : impLocals) (mcL' : MetricLogging.MetricLog) => exists (lH' : srcLocals) (mcH' : MetricLogging.MetricLog), states_compat lH' a2 lL' /\ @@ -1257,14 +1251,6 @@ Section RegAlloc. cbv beta. intros. simp. eexists. eexists. split. 2: split. 1,3: eauto. exists mcH'. exists mc'. split; eauto. - (* repeat (unfold check_bcond, assert_in, assignment in *; simp). *) - (* intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLoop_true in *; *) - (* destr cond; destr cond0; destr (isRegStr x); destr (isRegZ x0); try (destr (isRegStr y)); try (destr (isRegZ y0)); *) - (* try discriminate; *) - (* unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; simp; *) - (* destr mcL; destr mc'; destr mc; destr mcH'; destr mc'0; destr mcH'0; try blia. *) - (* all: try discr_match_success. *) - (* all: cbn in *; simp; try blia. *) + cbv beta. intros. simp. eapply exec.weaken. 1: eapply IH12. 1: eassumption. 1: eassumption. * eapply states_compat_extends. 2: eassumption. pose proof defuel_loop_inv as P. @@ -1280,13 +1266,13 @@ Section RegAlloc. intros. repeat (unfold check_bcond, assert_in, assignment in *; simp). clear -E0 E1 E2 H4p1p0 H4p1p1 H4p1 H4p3. - intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLoop_true in *; + intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLoop_true in *; try discr_match_success; destr cond; destr cond0; destr (isRegStr x); destr (isRegZ x0); try (destr (isRegStr y)); try (destr (isRegZ y0)); - try discriminate; try discr_match_success; - unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; simp. - all: destr mc''; destr mcH'; destr mcHmid; destr mcLmid; destr mc'; destr mcH'0; destr mc; destr mcL. - all: try discr_match_success. - all: cbn in *; simp; try blia. + try discriminate; try discr_match_success. + all: unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; simp. + all: destr mc''; destr mcH'; destr mcHmid; destr mcLmid; destr mc'; destr mcH'0; destr mc; destr mcL; + try discr_match_success. + all: cbn in *; try blia. - (* Case exec.seq *) rename H2 into IH2, IHexec into IH1. eapply exec.seq. From 276bb7b8e573a9b3c9279dfac1ec5d476a5a0332 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Tue, 12 Jul 2022 00:38:04 -0400 Subject: [PATCH 08/62] fix LowerPipeline.v --- compiler/src/compiler/LowerPipeline.v | 2 +- compiler/src/compiler/RegAlloc.v | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/compiler/src/compiler/LowerPipeline.v b/compiler/src/compiler/LowerPipeline.v index 2a5f95812..43acbdd50 100644 --- a/compiler/src/compiler/LowerPipeline.v +++ b/compiler/src/compiler/LowerPipeline.v @@ -561,7 +561,7 @@ Section LowerPipeline. (exists argnames retnames fbody, map.get p1 fname = Some (argnames, retnames, fbody) /\ forall l mc, map.of_list_zip argnames argvals = Some l -> - FlatImp.exec p1 fbody t m l mc (fun t' m' l' mc' => + FlatImp.exec isRegZ p1 fbody t m l mc (fun t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals)) -> riscv_call p2 fname t m argvals post. diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index ac6c27665..c5d2cd7bd 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -1234,7 +1234,7 @@ Section RegAlloc. unfold loop_inv in SC. rewrite E in SC. eapply exec.loop with - (mid4 := (fun (t'0 : Semantics.trace) (m'0 : mem) (lL' : impLocals) (mcL' : MetricLogging.MetricLog) => + (mid2 := (fun (t'0 : Semantics.trace) (m'0 : mem) (lL' : impLocals) (mcL' : MetricLogging.MetricLog) => exists (lH' : srcLocals) (mcH' : MetricLogging.MetricLog), states_compat lH' a2 lL' /\ (exists mcHmid mcLmid, From db9c19bb0a33f1517d3639ac77c146812aa3ace6 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Tue, 12 Jul 2022 00:46:08 -0400 Subject: [PATCH 09/62] start work on Spilling.v (#1) --- compiler/src/compiler/Spilling.v | 26 +++++++++++++++----------- 1 file changed, 15 insertions(+), 11 deletions(-) diff --git a/compiler/src/compiler/Spilling.v b/compiler/src/compiler/Spilling.v index aa5a09047..9f32aa376 100644 --- a/compiler/src/compiler/Spilling.v +++ b/compiler/src/compiler/Spilling.v @@ -1351,13 +1351,14 @@ Section Spilling. (fun (t2' : Semantics.trace) (m2' : mem) (l2' : locals) (mc2' : MetricLog) => exists t1' m1' l1' mc1', related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ - post t1' m1' l1' mc1'). + post t1' m1' l1' mc1' /\ + MetricLogging.metricsLeq (MetricLogging.metricsSub mc2' mc2) (MetricLogging.metricsSub mc1' mc1)). Proof. induction 1; intros; cbn [spill_stmt valid_vars_src Forall_vars_stmt] in *; fwd. - (* exec.interact *) eapply exec.seq_cps. eapply set_reg_range_to_vars_correct; try eassumption; try (unfold a0, a7; blia). - intros *. intros R GM. clear l2 mc2 H4. + intros *. intros R GM. clear l2 H4. unfold related in R. fwd. spec (subst_split (ok := mem_ok) m) as A. 1: eassumption. 1: ecancel_assumption. @@ -1395,12 +1396,15 @@ Section Spilling. { reflexivity. } { unfold a0, a7. blia. } { eassumption. } - { intros. do 4 eexists. split. 1: eassumption. eapply H2p1. - unfold map.split. split; [reflexivity|]. - move C at bottom. - unfold sep at 1 in C. destruct C as (mKeepL' & mRest & SC & ? & _). subst mKeepL'. - move H2 at bottom. unfold map.split in H2. fwd. - eapply map.shrink_disjoint_l; eassumption. } + { intros. do 4 eexists. split. 1: eassumption. split. + { eapply H2p1. + unfold map.split. split; [reflexivity|]. + move C at bottom. + unfold sep at 1 in C. destruct C as (mKeepL' & mRest & SC & ? & _). subst mKeepL'. + move H2 at bottom. unfold map.split in H2. fwd. + eapply map.shrink_disjoint_l; eassumption. } + admit. + } (* related for set_vars_to_reg_range_correct: *) unfold related. eexists _, _, _. ssplit. @@ -1531,7 +1535,7 @@ Section Spilling. } cbv beta. subst maxvar'. blia. } - cbv beta. intros tL5 mL5 lFL5 mcL5 (tH5 & mH5 & lFH5 & mcH5 & R5 & OC). + cbv beta. intros tL5 mL5 lFL5 mcL5 (tH5 & mH5 & lFH5 & mcH5 & R5 & OC & Hmetrics). match goal with | H: context[outcome], A: context[outcome] |- _ => specialize H with (1 := A); move H at bottom; rename H into Q @@ -1615,12 +1619,12 @@ Section Spilling. { unfold a0, a7. blia. } { eassumption. } { intros m22 l22 mc22 R22. do 4 eexists. split. 1: eassumption. - eassumption. } + split; try eassumption. Fail solve_MetricLog. admit. } - (* exec.load *) eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2 H3. intros. + intros. eapply exec.seq_cps. pose proof H2 as A. unfold related in A. fwd. unfold Memory.load, Memory.load_Z, Memory.load_bytes in *. fwd. From 6670dcd91cbed4c0bd2550f023b51f3387b15c42 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Wed, 20 Jul 2022 16:42:46 -0400 Subject: [PATCH 10/62] finish Spilling pass except some admits relating to SInteract --- compiler/src/compiler/Spilling.v | 307 +++++++++++++++++++------------ 1 file changed, 192 insertions(+), 115 deletions(-) diff --git a/compiler/src/compiler/Spilling.v b/compiler/src/compiler/Spilling.v index 9f32aa376..14ab6c828 100644 --- a/compiler/src/compiler/Spilling.v +++ b/compiler/src/compiler/Spilling.v @@ -450,13 +450,15 @@ Section Spilling. related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> fp < r <= maxvar /\ (r < a0 \/ a7 < r) -> map.get l1 r = Some v -> - (forall mc2, - related maxvar frame fpval t1 m1 l1 t2 m2 (map.put l2 (iarg_reg i r) v) -> - post t2 m2 (map.put l2 (iarg_reg i r) v) mc2) -> + (related maxvar frame fpval t1 m1 l1 t2 m2 (map.put l2 (iarg_reg i r) v) -> + post t2 m2 (map.put l2 (iarg_reg i r) v) + (if isRegZ r then mc2 else (addMetricInstructions 1 (addMetricLoads 2 mc2)))) -> exec e2 (load_iarg_reg i r) t2 m2 l2 mc2 post. Proof. intros. unfold load_iarg_reg, stack_loc, iarg_reg, related in *. fwd. + assert (isRegZ (9 + i) = true) by (unfold isRegZ; blia). + assert (isRegZ fp = true) by (unfold isRegZ; (assert (fp = 5) by auto); blia). destr (32 <=? r). - eapply exec.load. + eapply get_sep. ecancel_assumption. @@ -466,7 +468,10 @@ Section Spilling. eapply map.get_split_r. 1,3: eassumption. destr (map.get mp r); [exfalso|reflexivity]. specialize H0p2 with (1 := E0). blia. - + eapply H3. + + unfold exec.cost_SLoad. + assert (isRegZ r = false) by (unfold isRegZ; blia); rewrite H4 in H3. + unfold spill_tmp in H3. rewrite H0; rewrite H1. + eapply H3. repeat match goal with | |- exists _, _ => eexists | |- _ /\ _ => split @@ -483,6 +488,7 @@ Section Spilling. destr (map.get lStack r); [exfalso|reflexivity]. specialize H0p3 with (1 := E0). blia. } + assert (isRegZ r = true) by (unfold isRegZ; blia); rewrite H4 in H3. eapply H3. repeat match goal with | |- exists _, _ => eexists @@ -490,50 +496,50 @@ Section Spilling. | |- _ => eassumption || reflexivity end. Qed. - - Lemma load_iarg_reg_correct'(i: Z): forall r e2 t1 t2 m1 m2 l1 l2 mc1 mc2 post frame maxvar v fpval, - i = 1 \/ i = 2 -> - related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> - fp < r <= maxvar /\ (r < a0 \/ a7 < r) -> - map.get l1 r = Some v -> - post t1 m1 l1 mc1 -> - exec e2 (load_iarg_reg i r) t2 m2 l2 mc2 - (fun t2' m2' l2' mc2' => exists t1' m1' l1' mc1', - related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ post t1' m1' l1' mc1'). - Proof. - intros. - unfold load_iarg_reg, stack_loc, iarg_reg, related in *. fwd. - destr (32 <=? r). - - eapply exec.load. - + eapply get_sep. ecancel_assumption. - + eapply load_from_word_array. 1: ecancel_assumption. 2: blia. - eapply H0p6. 1: blia. - unfold sep in H0p4. fwd. - eapply map.get_split_r. 1,3: eassumption. - destr (map.get mp r); [exfalso|reflexivity]. - specialize H0p2 with (1 := E0). blia. - + repeat match goal with - | |- exists _, _ => eexists - | |- _ /\ _ => split - | |- _ => eassumption || reflexivity - end. - eapply put_tmp; eassumption. - - eapply exec.skip. - replace l2 with (map.put l2 r v) in H0p5|-*. 2: { - apply map.put_idemp. - edestruct (eq_sep_to_split l2) as (l2Rest & S22 & SP22). 1: ecancel_assumption. - eapply map.get_split_grow_r. 1: eassumption. - unfold sep in H0p4. destruct H0p4 as (lRegs' & lStack' & S2 & ? & ?). subst lRegs' lStack'. - eapply map.get_split_l. 1: exact S2. 2: assumption. - destr (map.get lStack r); [exfalso|reflexivity]. - specialize H0p3 with (1 := E0). blia. - } - repeat match goal with - | |- exists _, _ => eexists - | |- _ /\ _ => split - | |- _ => eassumption || reflexivity - end. - Qed. + + (* Lemma load_iarg_reg_correct'(i: Z): forall r e2 t1 t2 m1 m2 l1 l2 mc1 mc2 post frame maxvar v fpval, *) + (* i = 1 \/ i = 2 -> *) + (* related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> *) + (* fp < r <= maxvar /\ (r < a0 \/ a7 < r) -> *) + (* map.get l1 r = Some v -> *) + (* post t1 m1 l1 mc1 -> *) + (* exec e2 (load_iarg_reg i r) t2 m2 l2 mc2 *) + (* (fun t2' m2' l2' mc2' => exists t1' m1' l1' mc1', *) + (* related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ post t1' m1' l1' mc1'). *) + (* Proof. *) + (* intros. *) + (* unfold load_iarg_reg, stack_loc, iarg_reg, related in *. fwd. *) + (* destr (32 <=? r). *) + (* - eapply exec.load. *) + (* + eapply get_sep. ecancel_assumption. *) + (* + eapply load_from_word_array. 1: ecancel_assumption. 2: blia. *) + (* eapply H0p6. 1: blia. *) + (* unfold sep in H0p4. fwd. *) + (* eapply map.get_split_r. 1,3: eassumption. *) + (* destr (map.get mp r); [exfalso|reflexivity]. *) + (* specialize H0p2 with (1 := E0). blia. *) + (* + repeat match goal with *) + (* | |- exists _, _ => eexists *) + (* | |- _ /\ _ => split *) + (* | |- _ => eassumption || reflexivity *) + (* end. *) + (* eapply put_tmp; eassumption. *) + (* - eapply exec.skip. *) + (* replace l2 with (map.put l2 r v) in H0p5|-*. 2: { *) + (* apply map.put_idemp. *) + (* edestruct (eq_sep_to_split l2) as (l2Rest & S22 & SP22). 1: ecancel_assumption. *) + (* eapply map.get_split_grow_r. 1: eassumption. *) + (* unfold sep in H0p4. destruct H0p4 as (lRegs' & lStack' & S2 & ? & ?). subst lRegs' lStack'. *) + (* eapply map.get_split_l. 1: exact S2. 2: assumption. *) + (* destr (map.get lStack r); [exfalso|reflexivity]. *) + (* specialize H0p3 with (1 := E0). blia. *) + (* } *) + (* repeat match goal with *) + (* | |- exists _, _ => eexists *) + (* | |- _ /\ _ => split *) + (* | |- _ => eassumption || reflexivity *) + (* end. *) + (* Qed. *) (* Note: if we wanted to use this lemma in subgoals created by exec.loop, new postcondition must not mention the original t2, m2, l2, mc2, (even though @@ -548,7 +554,8 @@ Section Spilling. map.get l1 r = Some v -> exec e2 (load_iarg_reg i r) t2 m2 l2 mc2 (fun t2' m2' l2' mc2' => t2' = t2 /\ m2' = m2 /\ l2' = map.put l2 (iarg_reg i r) v /\ - related maxvar frame fpval t1 m1 l1 t2' m2' l2'). + related maxvar frame fpval t1 m1 l1 t2' m2' l2' /\ + (mc2' <= (if isRegZ r then mc2 else (addMetricInstructions 1 (addMetricLoads 2 mc2))))%metricsH). Proof. intros. unfold load_iarg_reg, stack_loc, iarg_reg, related in *. fwd. @@ -566,7 +573,10 @@ Section Spilling. | |- _ /\ _ => split | |- _ => eassumption || reflexivity end. - eapply put_tmp; eassumption. + 1: eapply put_tmp; eassumption. + unfold exec.cost_SLoad. assert (isRegZ (9+i) = true) by (unfold isRegZ; blia); rewrite H0. + assert (fp = 5) by auto; rewrite H1; cbn. + destr (isRegZ r); solve_MetricLog. - eapply exec.skip. assert (l2 = map.put l2 r v) as F. { symmetry. apply map.put_idemp. @@ -582,20 +592,24 @@ Section Spilling. | |- _ /\ _ => split | |- _ => eassumption || reflexivity end. + destr (isRegZ r); solve_MetricLog. Qed. (* SOp does not create an up-to-date `related` before we invoke this one, because after SOp, `related` does not hold: the result is already in l1 and lStack, but not yet in stackwords. So we request the `related` that held *before* SOp, i.e. the one where the result is not yet in l1 and l2. *) - Lemma save_ires_reg_correct: forall e t1 t2 m1 m2 l1 l2 mc1 mc2 x v maxvar frame post fpval, - post t1 m1 (map.put l1 x v) mc1 -> + Lemma save_ires_reg_correct: forall e t1 t2 m1 m2 l1 l2 mc1 mc1' mc2 mc2' x v maxvar frame post fpval, + post t1 m1 (map.put l1 x v) mc1' -> related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> fp < x <= maxvar /\ (x < a0 \/ a7 < x) -> - exec e (save_ires_reg x) t2 m2 (map.put l2 (ires_reg x) v) mc2 - (fun t2' m2' l2' mc2' => exists t1' m1' l1' mc1', - related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ post t1' m1' l1' mc1'). - Proof. + (mc2' - mc2 <= mc1' - (if isRegZ x then mc1 else (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc1)))))%metricsH -> + exec e (save_ires_reg x) t2 m2 (map.put l2 (ires_reg x) v) mc2' + (fun t2' m2' l2' mc2'' => exists t1' m1' l1' mc1'', + related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ + post t1' m1' l1' mc1'' /\ + (mc2'' - mc2 <= mc1'' - mc1)%metricsH). + Proof. intros. unfold save_ires_reg, stack_loc, ires_reg, related in *. fwd. destr (32 <=? x). @@ -639,12 +653,16 @@ Section Spilling. - eapply Nj. 1: blia. eauto. } 1: { unfold spill_tmp. eapply put_tmp; eauto. } - blia. + 1: blia. + unfold exec.cost_SStore. unfold spill_tmp; cbn. + destr (isRegZ x). + { solve_MetricLog. } + { solve_MetricLog. } } blia. - eapply exec.skip. - (* even though we did nothing, we have to reconstruct the `related` from the `related` that - held *before* the SOp *) + (* even though we did nothing, we have to reconstruct the `related` from the `related` that *) + (* held *before* the SOp *) repeat match goal with | |- exists _, _ => eexists | |- _ /\ _ => split @@ -680,12 +698,13 @@ Section Spilling. specialize H0p8 with (1 := H1). blia. } all: try eassumption. + destr (isRegZ x); solve_MetricLog. Qed. - (* SOp does not create an up-to-date `related` before we invoke this one, because after SOp, - `related` does not hold: the result is already in l1 and lStack, but not yet in stackwords. - So we request the `related` that held *before* SOp, i.e. the one where the result is not - yet in l1 and l2. *) + (* SOp does not create an up-to-date `related` before we invoke this one, because after SOp, *) + (* `related` does not hold: the result is already in l1 and lStack, but not yet in stackwords. *) + (* So we request the `related` that held *before* SOp, i.e. the one where the result is not *) + (* yet in l1 and l2. *) Lemma save_ires_reg_correct'': forall e t1 t2 m1 m2 l1 l2 mc2 x v maxvar frame post fpval, related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> fp < x <= maxvar /\ (x < a0 \/ a7 < x) -> @@ -742,8 +761,8 @@ Section Spilling. blia. - eapply exec.skip. eapply H1. - (* even though we did nothing, we have to reconstruct the `related` from the `related` that - held *before* the SOp *) + (* even though we did nothing, we have to reconstruct the `related` from the `related` that *) + (* held *before* the SOp *) repeat match goal with | |- exists _, _ => eexists | |- _ /\ _ => split @@ -1153,7 +1172,8 @@ Section Spilling. (fun (t2' : Semantics.trace) (m2' : mem) (l2' : locals) (mc2' : MetricLog) => exists t1' m1' l1' mc1', related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ - post t1' m1' l1' mc1'). + post t1' m1' l1' mc1' /\ + (mc2' - mc2 <= mc1' - mc1)%metricsH). Definition call_spec(e: env) '(argnames, retnames, fbody) (t: Semantics.trace)(m: mem)(argvals: list word) @@ -1335,6 +1355,30 @@ Section Spilling. all: try assumption. Qed. + + Lemma iarg_reg_isReg: forall i a, + (i <= 10) -> + (isRegZ (iarg_reg i a) = true). + Proof. + intros. unfold isRegZ, iarg_reg. destr (32 <=? a); unfold spill_tmp; blia. + Qed. + + Lemma ires_reg_isReg: forall r, + (isRegZ (ires_reg r) = true). + Proof. + intros. unfold isRegZ, ires_reg. destr (32 <=? r); unfold spill_tmp; blia. + Qed. + + Ltac isReg_helper := + match goal with + | |- context[(isRegZ (iarg_reg _ _))] => rewrite iarg_reg_isReg by blia + | H: context[(isRegZ (iarg_reg _ _))] |- _ => rewrite iarg_reg_isReg in H by blia + | |- context[(isRegZ (ires_reg _))] => rewrite ires_reg_isReg + | H: context[(isRegZ (ires_reg _))] |- _ => rewrite ires_reg_isReg in H + end. + + + Lemma spilling_correct (e1 e2 : env) (Ev : spill_functions e1 = Success e2) (s1 : stmt) (t1 : Semantics.trace) @@ -1352,7 +1396,7 @@ Section Spilling. exists t1' m1' l1' mc1', related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ post t1' m1' l1' mc1' /\ - MetricLogging.metricsLeq (MetricLogging.metricsSub mc2' mc2) (MetricLogging.metricsSub mc1' mc1)). + (mc2' - mc2 <= mc1' - mc1)%metricsH). Proof. induction 1; intros; cbn [spill_stmt valid_vars_src Forall_vars_stmt] in *; fwd. - (* exec.interact *) @@ -1403,6 +1447,7 @@ Section Spilling. unfold sep at 1 in C. destruct C as (mKeepL' & mRest & SC & ? & _). subst mKeepL'. move H2 at bottom. unfold map.split in H2. fwd. eapply map.shrink_disjoint_l; eassumption. } + admit. } (* related for set_vars_to_reg_range_correct: *) @@ -1624,7 +1669,7 @@ Section Spilling. - (* exec.load *) eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - intros. + intros. eapply exec.seq_cps. pose proof H2 as A. unfold related in A. fwd. unfold Memory.load, Memory.load_Z, Memory.load_bytes in *. fwd. @@ -1638,12 +1683,12 @@ Section Spilling. + eassumption. + eassumption. + blia. + + unfold exec.cost_SLoad in *; repeat isReg_helper. + destr (isRegZ x); destr (isRegZ a); solve_MetricLog. - (* exec.store *) - eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2 H4. intros. - eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2 H3. intros. - pose proof H3 as A. unfold related in A. fwd. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. + pose proof H5 as A. unfold related in A. fwd. unfold Memory.store, Memory.store_Z, Memory.store_bytes in *. fwd. edestruct (@sep_def _ _ _ m2 (eq m)) as (m' & m2Rest & Sp & ? & ?). 1: ecancel_assumption. unfold map.split in Sp. subst. fwd. @@ -1653,18 +1698,19 @@ Section Spilling. { unfold Memory.store, Memory.store_Z, Memory.store_bytes. unfold Memory.load_bytes in *. erewrite map.getmany_of_tuple_in_disjoint_putmany; eauto. } - do 4 eexists. split. 2: eassumption. - unfold related. - repeat match goal with - | |- exists _, _ => eexists - | |- _ /\ _ => split - end. - all: try eassumption || reflexivity. - spec store_bytes_sep_hi2lo as A. 1: eassumption. - all: ecancel_assumption. + do 4 eexists. split. 2: split. 2: eassumption. + + unfold related. + repeat match goal with + | |- exists _, _ => eexists + | |- _ /\ _ => split + end. + all: try eassumption || reflexivity. + spec store_bytes_sep_hi2lo as A. 1: eassumption. + all: ecancel_assumption. + + unfold exec.cost_SStore; repeat isReg_helper. + destr (isRegZ v); destr (isRegZ a); solve_MetricLog. - (* exec.inlinetable *) - eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2 H4. intros. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.seq_cps. eapply exec.inlinetable. { unfold ires_reg, iarg_reg, spill_tmp, fp, a0, a7 in *. destr (32 <=? x); destr (32 <=? i); try blia. } @@ -1674,6 +1720,8 @@ Section Spilling. + eassumption. + eassumption. + blia. + + unfold exec.cost_SInlinetable in *; repeat isReg_helper. + destr (isRegZ x); destr (isRegZ i); solve_MetricLog. - (* exec.stackalloc *) rename H1 into IH. eapply exec.stackalloc. 1: assumption. @@ -1691,17 +1739,18 @@ Section Spilling. | |- _ /\ _ => split end. 1,4,3,2: eassumption. + unfold exec.cost_SStackalloc in *; repeat isReg_helper. + destr (isRegZ x); try solve_MetricLog. all: admit. - (* exec.lit *) eapply exec.seq_cps. eapply exec.lit. eapply save_ires_reg_correct. + eassumption. + eassumption. + blia. + + unfold exec.cost_SLit; repeat isReg_helper. destr (isRegZ x); solve_MetricLog. - (* exec.op *) - eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2 H3. intros. - eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2 H2. intros. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.seq_cps. eapply exec.op. 1: eapply get_iarg_reg_1; eauto with zarith. @@ -1710,49 +1759,62 @@ Section Spilling. + eassumption. + eassumption. + blia. + + unfold exec.cost_SOp; repeat isReg_helper. + destr (isRegZ z); destr (isRegZ x); destr (isRegZ y); solve_MetricLog. - (* exec.set *) - eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2 H2. intros. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.seq_cps. eapply exec.set. 1: apply map.get_put_same. eapply save_ires_reg_correct. + eassumption. + eassumption. + blia. + + unfold exec.cost_SSet; repeat isReg_helper. + destr (isRegZ x); destr (isRegZ y); solve_MetricLog. - (* exec.if_true *) unfold prepare_bcond. destr cond; cbn [ForallVars_bcond eval_bcond spill_bcond] in *; fwd. + eapply exec.seq_assoc. - eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2 H2. intros. - eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2. intros. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.if_true. { cbn. erewrite get_iarg_reg_1 by eauto with zarith. rewrite map.get_put_same. congruence. } - eapply IHexec; eassumption. - + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2 H2. intros. + eapply exec.weaken. + * eapply IHexec; eassumption. + * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. + unfold exec.cost_SIf in *; repeat isReg_helper. + destr (isRegZ y); destr (isRegZ x); solve_MetricLog. + + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.if_true. { cbn. rewrite map.get_put_same. congruence. } - eapply IHexec; eassumption. + eapply exec.weaken. + * eapply IHexec; eassumption. + * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. + unfold exec.cost_SIf in *; repeat isReg_helper. + destr (isRegZ x); solve_MetricLog. - (* exec.if_false *) unfold prepare_bcond. destr cond; cbn [ForallVars_bcond eval_bcond spill_bcond] in *; fwd. + eapply exec.seq_assoc. - eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2 H2. intros. - eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2. intros. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.if_false. { cbn. erewrite get_iarg_reg_1 by eauto with zarith. rewrite map.get_put_same. congruence. } - eapply IHexec; eassumption. - + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear mc2 H2. intros. + eapply exec.weaken. + * eapply IHexec; eassumption. + * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. + unfold exec.cost_SIf in *; repeat isReg_helper. + destr (isRegZ x); destr (isRegZ y); solve_MetricLog. + + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.if_false. { cbn. rewrite map.get_put_same. congruence. } - eapply IHexec; eassumption. + eapply exec.weaken. + * eapply IHexec; eassumption. + * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. + unfold exec.cost_SIf in *; repeat isReg_helper. + destr (isRegZ x); solve_MetricLog. - (* exec.loop *) rename IHexec into IH1, H3 into IH2, H5 into IH12. eapply exec.loop_cps. @@ -1772,15 +1834,21 @@ Section Spilling. erewrite get_iarg_reg_1 by eauto with zarith. rewrite map.get_put_same. eexists. split; [reflexivity|]. split; intros. - * do 4 eexists. split. - -- exact H3p6. + * do 4 eexists. split. 2: split. + -- exact H3p8. -- eapply H1. 1: eassumption. cbn. rewrite E, E0. congruence. + -- unfold exec.cost_SLoop_false; repeat isReg_helper. + destr (isRegZ x); destr (isRegZ y); solve_MetricLog. * eapply exec.weaken. 1: eapply IH2. -- eassumption. -- cbn. rewrite E, E0. congruence. -- eassumption. -- eassumption. - -- cbv beta. intros. fwd. eauto 10. (* IH12 *) + -- cbv beta. intros. fwd. eapply exec.weaken. + ++ eapply IH12; try eassumption. repeat split; eauto; blia. + ++ cbv beta; intros; fwd. exists t1'1, m1'1, l1'1, mc1'1. split. 2:split. all: try eassumption. + unfold exec.cost_SLoop_true in *; repeat isReg_helper. + destr (isRegZ x); destr (isRegZ y); solve_MetricLog. + specialize H0 with (1 := H3p1). cbn in H0. fwd. eapply exec.weaken. { eapply load_iarg_reg_correct''; (blia || eassumption || idtac). @@ -1788,24 +1856,33 @@ Section Spilling. cbv beta. intros. fwd. cbn [eval_bcond spill_bcond]. rewrite map.get_put_same. eexists. split; [reflexivity|]. split; intros. - * do 4 eexists. split. - -- exact H3p5. + * do 4 eexists. split. 2: split. + -- exact H3p6. -- eapply H1. 1: eassumption. cbn. rewrite E. congruence. + -- unfold exec.cost_SLoop_false; repeat isReg_helper. + destr (isRegZ x); solve_MetricLog. * eapply exec.weaken. 1: eapply IH2. -- eassumption. -- cbn. rewrite E. congruence. -- eassumption. -- eassumption. - -- cbv beta. intros. fwd. eauto 10. (* IH12 *) + -- cbv beta. intros. fwd. eapply exec.weaken. + ++ eapply IH12; try eassumption. repeat split; eauto; blia. + ++ cbv beta; intros; fwd. exists t1'1, m1'1, l1'1, mc1'1. split. 2:split. all: try eassumption. + unfold exec.cost_SLoop_true in *; repeat isReg_helper. + destr (isRegZ x); solve_MetricLog. - (* exec.seq *) cbn in *. fwd. rename H1 into IH2, IHexec into IH1. eapply exec.seq. + eapply IH1. 1: eassumption. eauto 15. - + cbn. intros. fwd. eapply IH2. 1,2: eassumption. eauto 15. + + cbn. intros. fwd. eapply exec.weaken. + * eapply IH2; eassumption. + * cbv beta. intros. fwd. exists t1'0, m1'0, l1'0, mc1'0. split. 2:split. all: try eassumption. + solve_MetricLog. - (* exec.skip *) - eapply exec.skip. eauto 20. - Qed. + eapply exec.skip. exists t, m, l, mc. repeat split; eauto; solve_MetricLog. + Admitted. Lemma spill_fun_correct: forall e1 e2 argnames1 retnames1 body1 argnames2 retnames2 body2, spill_functions e1 = Success e2 -> From 7d1d671a98d2e6eccaaca32bc20414d69b29e649 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Wed, 20 Jul 2022 17:10:17 -0400 Subject: [PATCH 11/62] FlattenExpr pass with admits for metrics goals, because bedrock2 semantics not updated yet --- compiler/src/compiler/FlattenExpr.v | 34 +++++++++++++++-------------- compiler/src/compiler/Pipeline.v | 8 ++++--- 2 files changed, 23 insertions(+), 19 deletions(-) diff --git a/compiler/src/compiler/FlattenExpr.v b/compiler/src/compiler/FlattenExpr.v index 81e6126b2..fd99c277b 100644 --- a/compiler/src/compiler/FlattenExpr.v +++ b/compiler/src/compiler/FlattenExpr.v @@ -373,19 +373,21 @@ Section FlattenExpr1. Proof. induction e; intros *; intros F Ex U D Ev; simpl in *; simp. + (* Ltac solve_MetricLog ::= try solve [solve_MetricLog; idtac " Admitted metrics goal"; admit]. *) + - (* expr.literal *) - eapply @FlatImp.exec.lit; t_safe; solve_MetricLog. + eapply @FlatImp.exec.lit; t_safe. Fail solve_MetricLog. admit. - (* expr.var *) destruct oResVar; simp. - + eapply @FlatImp.exec.set; t_safe; [maps | solve_MetricLog]. + + eapply @FlatImp.exec.set; t_safe; [maps | try solve_MetricLog]. admit. + eapply @FlatImp.exec.skip; t_safe. solve_MetricLog. - (* expr.load *) eapply @FlatImp.exec.seq. + eapply IHe; try eassumption. maps. + intros. simpl in *. simp. - eapply @FlatImp.exec.load; t_safe; rewrite ?add_0_r; try eassumption; solve_MetricLog. + eapply @FlatImp.exec.load; t_safe; rewrite ?add_0_r; try eassumption. try solve_MetricLog. admit. - (* expr.inlinetable *) repeat match goal with @@ -396,7 +398,7 @@ Section FlattenExpr1. + eapply IHe; try eassumption. 1: maps. set_solver; destr (String.eqb s0 x); subst; tauto. (* TODO improve set_solver? *) + intros. simpl in *. simp. - eapply @FlatImp.exec.inlinetable; t_safe; try eassumption. 2: solve_MetricLog. + eapply @FlatImp.exec.inlinetable; t_safe; try eassumption. 2: try solve_MetricLog; admit. apply_in_hyps flattenExpr_uses_Some_resVar. subst s0. intro C. subst s2. destruct oResVar. @@ -412,7 +414,7 @@ Section FlattenExpr1. * eapply IHe2. 1: eassumption. 4: eassumption. 1,2: solve [maps]. clear IHe1 IHe2. pose_flatten_var_ineqs. set_solver. * intros. simpl in *. simp. clear IHe1 IHe2. - eapply @FlatImp.exec.op; t_safe; t_safe. 2 : solve_MetricLog. + eapply @FlatImp.exec.op; t_safe; t_safe. 2 : try solve_MetricLog; admit. eapply flattenExpr_valid_resVar in E1; maps. - (* expr.ite *) @@ -454,7 +456,7 @@ Section FlattenExpr1. eapply subset_union_rr. eapply subset_refl. } - ++ cbv beta. intros. simp. t_safe. 2: solve_MetricLog. + ++ cbv beta. intros. simp. t_safe. 2: try solve_MetricLog; admit. clear IHe1 IHe2 IHe3. apply_in_hyps flattenExpr_uses_Some_resVar. subst. assumption. * eapply FlatImp.exec.if_true. @@ -488,10 +490,10 @@ Section FlattenExpr1. eapply subset_union_rl. eapply subset_refl. } - ++ cbv beta. intros. simp. t_safe. 2: solve_MetricLog. + ++ cbv beta. intros. simp. t_safe. 2: try solve_MetricLog; admit. clear IHe1 IHe2 IHe3. apply_in_hyps flattenExpr_uses_Some_resVar. subst. assumption. - Qed. + Admitted. Goal True. idtac "FlattenExpr: flattenExpr_correct_aux done". Abort. Lemma flattenExpr_correct_with_modVars : forall e fenv oResVar ngs1 ngs2 resVar s t m lH lL initialMcH initialMcL finalMcH res, @@ -613,7 +615,7 @@ Section FlattenExpr1. | intros; simpl in *; simp; default_flattenBooleanExpr ]. - 2, 4, 6 : solve_MetricLog. + 2, 4, 6 : try solve_MetricLog. all: rewrite bool_to_word_to_bool_id; destruct_one_match; @@ -734,7 +736,7 @@ Section FlattenExpr1. * intros. simpl in *. simp. eapply @FlatImp.exec.store; rewrite ?add_0_r; try eassumption. { eapply flattenExpr_valid_resVar in E; maps. } - { repeat eexists; repeat (split || eassumption || solve_MetricLog); maps. } + { repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all:admit. } - (* exec.stackalloc *) eapply @FlatImp.exec.stackalloc. 1: eassumption. @@ -742,7 +744,7 @@ Section FlattenExpr1. eapply @FlatImp.exec.weaken. { eapply IHexec; try reflexivity; try eassumption; maps. } { intros. simpl in *. simp. do 2 eexists. ssplit; try eassumption. - do 2 eexists. ssplit; try eassumption; try solve_MetricLog. maps. } + do 2 eexists. ssplit; try eassumption; try solve_MetricLog; try maps. admit. } - (* if_true *) eapply @FlatImp.exec.seq. @@ -759,7 +761,7 @@ Section FlattenExpr1. eapply @FlatImp.exec.weaken. { eapply IHexec; try reflexivity; try eassumption; maps. } { intros. simpl in *. simp. - repeat eexists; repeat (split || eassumption || solve_MetricLog). maps. } + repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all:admit. } - (* if_false *) eapply @FlatImp.exec.seq. @@ -776,7 +778,7 @@ Section FlattenExpr1. eapply @FlatImp.exec.weaken. { eapply IHexec; try reflexivity; try eassumption; maps. } { intros. simpl in *. simp. - repeat eexists; repeat (split || eassumption || solve_MetricLog). maps. } + repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all:admit. } - (* seq *) eapply seq_with_modVars. @@ -808,7 +810,7 @@ Section FlattenExpr1. | intros; simpl in *; simp .. ]. + maps. + congruence. - + repeat eexists; repeat (split || eassumption || solve_MetricLog); maps. + + repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all:admit. + exfalso. match goal with | H: context [word.eqb _ _] |- _ => rewrite word.eqb_eq in H @@ -853,7 +855,7 @@ Section FlattenExpr1. 1,3: solve [maps]. pose proof (ExprImp.modVars_subset_allVars c). maps. * simpl. intros. simp. - repeat eexists; repeat (split || eassumption || solve_MetricLog). maps. + repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all:admit. - (* call *) unfold flattenCall in *. simp. @@ -919,7 +921,7 @@ Section FlattenExpr1. simple apply conj; [eassumption|]. split; [simple eapply map.only_differ_putmany; eassumption|]. solve_MetricLog. - Qed. + Admitted. Goal True. idtac "FlattenExpr: flattenStmt_correct_aux done". Abort. Lemma flattenStmt_correct: forall eH eL sH sL lL t m mc post, diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index 0ae2185a6..71b8ef8ec 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -172,7 +172,7 @@ Section WithWordAndMem. GetArgCount := get_argcount; GetRetCount := get_retcount; Valid := map.forall_values ParamsNoDup; - Call := locals_based_call_spec FlatImp.exec; + Call := locals_based_call_spec (FlatImp.exec isRegStr); |}. (* | *) (* | RegAlloc *) @@ -182,7 +182,7 @@ Section WithWordAndMem. GetArgCount := get_argcount; GetRetCount := get_retcount; Valid := map.forall_values ParamsNoDup; - Call := locals_based_call_spec FlatImp.exec; + Call := locals_based_call_spec (FlatImp.exec isRegZ); |}. (* | *) (* | Spilling *) @@ -192,7 +192,7 @@ Section WithWordAndMem. GetArgCount := get_argcount; GetRetCount := get_retcount; Valid := map.forall_values FlatToRiscvDef.valid_FlatImp_fun; - Call := locals_based_call_spec FlatImp.exec; + Call := locals_based_call_spec (FlatImp.exec isRegZ); |}. (* | *) (* | FlatToRiscv *) @@ -362,6 +362,8 @@ Section WithWordAndMem. rewrite H1 in P'. inversion P'. exact Cp. - simpl. intros. fwd. eexists. split. 2: eassumption. eauto using states_compat_getmany. + Unshelve. + all: repeat constructor. Qed. Ltac debool := From 402f67fc3a1b07dca75c9ec3e87f271890905f0f Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Wed, 20 Jul 2022 17:31:12 -0400 Subject: [PATCH 12/62] cleaning up --- deps/coq-record-update | 2 +- deps/coqutil | 2 +- deps/kami | 2 +- deps/riscv-coq | 2 +- 4 files changed, 4 insertions(+), 4 deletions(-) diff --git a/deps/coq-record-update b/deps/coq-record-update index 57ff53940..a0c038f56 160000 --- a/deps/coq-record-update +++ b/deps/coq-record-update @@ -1 +1 @@ -Subproject commit 57ff539400ed22fd7c881ff66b96dde1c13710ed +Subproject commit a0c038f5691ee5038d9e83e58c1fe7f541d2501d diff --git a/deps/coqutil b/deps/coqutil index 5d391eb1e..4d109fc0a 160000 --- a/deps/coqutil +++ b/deps/coqutil @@ -1 +1 @@ -Subproject commit 5d391eb1e32c5eb27ff1a3f1bfb0e14bad53931f +Subproject commit 4d109fc0a342cb235c2e99b76104c5bf0475e418 diff --git a/deps/kami b/deps/kami index 66fd05345..5bbc402f9 160000 --- a/deps/kami +++ b/deps/kami @@ -1 +1 @@ -Subproject commit 66fd05345a4993f5eb5a36d5756cf43ea1bfa074 +Subproject commit 5bbc402f9580b1086fc5798572394d9a6964aad8 diff --git a/deps/riscv-coq b/deps/riscv-coq index 12128792b..0fd5faa1f 160000 --- a/deps/riscv-coq +++ b/deps/riscv-coq @@ -1 +1 @@ -Subproject commit 12128792b1d01700eb4ee3cc8e470193d14c457b +Subproject commit 0fd5faa1fb8721ece3924d81d5b8d14b5b54cdf9 From ca1bd2a1d4ba529eb17404188d148fe8279fdb07 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Wed, 20 Jul 2022 18:00:42 -0400 Subject: [PATCH 13/62] update submodules --- deps/coq-record-update | 2 +- deps/coqutil | 2 +- deps/kami | 2 +- deps/riscv-coq | 2 +- 4 files changed, 4 insertions(+), 4 deletions(-) diff --git a/deps/coq-record-update b/deps/coq-record-update index a0c038f56..57ff53940 160000 --- a/deps/coq-record-update +++ b/deps/coq-record-update @@ -1 +1 @@ -Subproject commit a0c038f5691ee5038d9e83e58c1fe7f541d2501d +Subproject commit 57ff539400ed22fd7c881ff66b96dde1c13710ed diff --git a/deps/coqutil b/deps/coqutil index 4d109fc0a..5d391eb1e 160000 --- a/deps/coqutil +++ b/deps/coqutil @@ -1 +1 @@ -Subproject commit 4d109fc0a342cb235c2e99b76104c5bf0475e418 +Subproject commit 5d391eb1e32c5eb27ff1a3f1bfb0e14bad53931f diff --git a/deps/kami b/deps/kami index 5bbc402f9..66fd05345 160000 --- a/deps/kami +++ b/deps/kami @@ -1 +1 @@ -Subproject commit 5bbc402f9580b1086fc5798572394d9a6964aad8 +Subproject commit 66fd05345a4993f5eb5a36d5756cf43ea1bfa074 diff --git a/deps/riscv-coq b/deps/riscv-coq index 0fd5faa1f..12128792b 160000 --- a/deps/riscv-coq +++ b/deps/riscv-coq @@ -1 +1 @@ -Subproject commit 0fd5faa1fb8721ece3924d81d5b8d14b5b54cdf9 +Subproject commit 12128792b1d01700eb4ee3cc8e470193d14c457b From f8af06b1d585156c89ee0684abcf24c652bad738 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Wed, 20 Jul 2022 18:24:56 -0400 Subject: [PATCH 14/62] fix merge issues --- compiler/src/compiler/FlatToRiscvFunctions.v | 4 +- compiler/src/compiler/FlatToRiscvMetric.v | 344 ------------------- compiler/src/compiler/LowerPipeline.v | 6 +- 3 files changed, 4 insertions(+), 350 deletions(-) delete mode 100644 compiler/src/compiler/FlatToRiscvMetric.v diff --git a/compiler/src/compiler/FlatToRiscvFunctions.v b/compiler/src/compiler/FlatToRiscvFunctions.v index a56a139ce..c36d50d66 100644 --- a/compiler/src/compiler/FlatToRiscvFunctions.v +++ b/compiler/src/compiler/FlatToRiscvFunctions.v @@ -1621,9 +1621,6 @@ Section Proofs. repeat match goal with | x := _ |- _ => clearbody x end. - - clear - word_ok RVM PRParams PR ext_spec word_riscv_ok locals_ok mem_ok fun_info_ok env_ok - IHexec OC BC OL Exb GetMany Ext GE FS C V Mo Mo' Gra RaM GPC A GM. revert IHexec OC BC OL Exb GetMany Ext GE FS C V Mo Mo' Gra RaM GPC A GM. eapply compile_function_body_correct. } @@ -2128,3 +2125,4 @@ Section Proofs. End Proofs. + diff --git a/compiler/src/compiler/FlatToRiscvMetric.v b/compiler/src/compiler/FlatToRiscvMetric.v deleted file mode 100644 index 395d14cde..000000000 --- a/compiler/src/compiler/FlatToRiscvMetric.v +++ /dev/null @@ -1,344 +0,0 @@ -Require Import riscv.Spec.Primitives. -Require Import riscv.Platform.RiscvMachine. -Require Import riscv.Platform.MetricRiscvMachine. -Require Import riscv.Platform.MetricLogging. -Require Import riscv.Utility.Utility. -Require Import riscv.Utility.runsToNonDet. -Require Import riscv.Utility.InstructionCoercions. -Require Import compiler.util.Common. -Require Import compiler.eqexact. -Require Import coqutil.Tactics.Simp. -Require Import compiler.on_hyp_containing. -Require Import compiler.SeparationLogic. -Require Export coqutil.Word.SimplWordExpr. -Require Import compiler.GoFlatToRiscv. -Require Import compiler.DivisibleBy4. -Require Import compiler.MetricsToRiscv. -Require Import compiler.FlatImp. -Require Import compiler.RiscvWordProperties. -Require Import compiler.FlatToRiscvDef. -Require Import compiler.FlatToRiscvCommon. -Require Import compiler.FlatToRiscvLiterals. - -Open Scope ilist_scope. - -Local Arguments Z.mul: simpl never. -Local Arguments Z.add: simpl never. -Local Arguments Z.of_nat: simpl never. -Local Arguments Z.modulo : simpl never. -Local Arguments Z.pow: simpl never. -Local Arguments Z.sub: simpl never. - -Section Proofs. - Context {iset: Decode.InstructionSet}. - Context {pos_map: map.map String.string Z}. - Context (compile_ext_call: pos_map -> Z -> Z -> stmt Z -> list Decode.Instruction). - Context {width: Z} {BW: Bitwidth width} {word: word.word width} {word_ok: word.ok word}. - Context {word_riscv_ok: RiscvWordProperties.word.riscv_ok word}. - Context {locals: map.map Z word} {locals_ok: map.ok locals}. - Context {mem: map.map word byte} {mem_ok: map.ok mem}. - Context {env: map.map String.string (list Z * list Z * stmt Z)} {env_ok: map.ok env}. - Context {M: Type -> Type}. - Context {MM: Monads.Monad M}. - Context {RVM: Machine.RiscvProgram M word}. - Context {PRParams: PrimitivesParams M MetricRiscvMachine}. - Context {ext_spec: Semantics.ExtSpec} {ext_spec_ok: Semantics.ext_spec.ok ext_spec}. - Context {PR: MetricPrimitives.MetricPrimitives PRParams}. - Context {BWM: bitwidth_iset width iset}. - - Add Ring wring : (word.ring_theory (word := word)) - (preprocess [autorewrite with rew_word_morphism], - morphism (word.ring_morph (word := word)), - constants [word_cst]). - - Local Notation RiscvMachineL := (MetricRiscvMachine (width := width)). - - Ltac run1done := - apply runsToDone; - simpl_MetricRiscvMachine_get_set; - rewrite ?word.sru_ignores_hibits, - ?word.slu_ignores_hibits, - ?word.srs_ignores_hibits, - ?word.mulhuu_simpl, - ?word.divu0_simpl, - ?word.modu0_simpl; - simpl in *; - eexists; (* finalMH *) - eexists; (* finalMetricsH *) - repeat split; - simpl_word_exprs word_ok; - first - [ solve [eauto] - | solve_MetricLog - | solve_word_eq word_ok - | solve [wcancel_assumption] - | eapply rearrange_footpr_subset; [ eassumption | wwcancel ] - | solve [solve_valid_machine word_ok] - | idtac ]. - - Ltac IH_sidecondition := - simpl_word_exprs word_ok; - try solve - [ reflexivity - | auto - | solve_word_eq word_ok - | simpl; solve_divisibleBy4 - | solve_valid_machine word_ok - | eapply rearrange_footpr_subset; [ eassumption | wwcancel ] - | wcancel_assumption ]. - - Hypothesis no_ext_calls: forall t mGive action argvals outcome, - ext_spec t mGive action argvals outcome -> False. - - Hypothesis stackalloc_always_0: forall x n body t m (l: locals) mc post, - FlatImp.exec isRegZ map.empty (SStackalloc x n body) t m l mc post -> n = 0. - - Hypothesis sp_always_set: forall l: locals, - map.get l RegisterNames.sp = Some (word.of_Z 42). - - (* not needed any more *) - Definition stmt_not_too_big(s: stmt Z): Prop := True. - - Lemma compile_stmt_correct: - forall (s: stmt Z) t initialMH initialRegsH postH initialMetricsH, - FlatImp.exec isRegZ map.empty s t initialMH initialRegsH initialMetricsH postH -> - forall R Rexec (initialL: RiscvMachineL) insts pos, - compile_stmt iset compile_ext_call map.empty pos 12345678 s = insts -> - stmt_not_too_big s -> - valid_FlatImp_vars s -> - divisibleBy4 initialL.(getPc) -> - initialL.(getRegs) = initialRegsH -> - subset (footpr (program iset initialL.(getPc) insts * Rexec)%sep) (of_list initialL.(getXAddrs)) -> - (program iset initialL.(getPc) insts * Rexec * eq initialMH * R)%sep initialL.(getMem) -> - initialL.(getLog) = t -> - initialL.(getNextPc) = add initialL.(getPc) (word.of_Z 4) -> - valid_machine initialL -> - runsTo initialL (fun finalL => exists finalMH finalMetricsH, - postH finalL.(getLog) finalMH finalL.(getRegs) finalMetricsH /\ - subset (footpr (program iset initialL.(getPc) insts * Rexec)%sep) - (of_list finalL.(getXAddrs)) /\ - (program iset initialL.(getPc) insts * Rexec * eq finalMH * R)%sep finalL.(getMem) /\ - finalL.(getPc) = add initialL.(getPc) - (word.mul (word.of_Z 4) (word.of_Z (Z.of_nat (length insts)))) /\ - finalL.(getNextPc) = word.add finalL.(getPc) (word.of_Z 4) /\ - (finalL.(getMetrics) - initialL.(getMetrics) <= - lowerMetrics (finalMetricsH - initialMetricsH))%metricsL /\ - valid_machine finalL). - Proof. - induction 1; intros; - repeat match goal with - | m: _ |- _ => destruct_RiscvMachine m; simpl_MetricRiscvMachine_get_set - end; - simpl in *; - subst; - simp. - - - (* SInteract *) - exfalso. eapply no_ext_calls. eassumption. - - - (* SCall *) - lazymatch goal with - | A: map.get map.empty _ = Some _ |- _ => - exfalso; simpl in *; - rewrite map.get_empty in A - end. - discriminate. - - - (* SLoad *) - unfold Memory.load, Memory.load_Z in *. simp. subst_load_bytes_for_eq. - run1det. run1done. - - - (* SStore *) - simpl_MetricRiscvMachine_get_set. - assert ((eq m * (program iset initialL_pc [[compile_store iset sz a v o]] * Rexec * R))%sep - initialL_mem) as A by ecancel_assumption. - match goal with - | H: _ |- _ => pose proof (store_bytes_frame H A) as P; move H at bottom; - unfold Memory.store, Memory.store_Z, Memory.store_bytes in H - end. - destruct P as (finalML & P1 & P2). - simp. - destruct (eq_sym (LittleEndianList.length_le_split (Memory.bytes_per(width:=width) sz) (word.unsigned val))) in t0, E. - subst_load_bytes_for_eq. - run1det. run1done. - eapply preserve_subset_of_xAddrs. 1: assumption. - ecancel_assumption. - - - (* SInlinetable *) - run1det. - assert (map.get (map.put l x (word.add initialL_pc (word.of_Z 4))) i = Some index). { - rewrite map.get_put_diff by congruence. assumption. - } - run1det. - assert (Memory.load sz initialL_mem - (word.add (word.add (word.add initialL_pc (word.of_Z 4)) index) (word.of_Z 0)) - = Some v). { - rewrite add_0_r. - eapply load_from_compile_byte_list. 1: eassumption. - wcancel_assumption. - } - run1det. - rewrite !map.put_put_same in *. - run1done. - - - (* SStackalloc *) - assert (valid_register RegisterNames.sp) by (cbv; auto). - specialize (stackalloc_always_0 x n body t mSmall l mc post). move stackalloc_always_0 at bottom. - assert (n = 0). { - eapply stackalloc_always_0. econstructor; eauto. - } - subst n. - run1det. - eapply runsTo_weaken. { - eapply H1 with (mStack := map.empty) (mCombined := mSmall). - { unfold Memory.anybytes. exists nil. reflexivity. } - { rewrite map.split_empty_r. reflexivity. } - all: IH_sidecondition. - } - simpl. - intros. - unfold Memory.anybytes, Memory.ftprint, map.of_disjoint_list_zip in *. simpl in *. - simp. - rewrite map.split_empty_r in H6p0p1. subst mSmall'. - repeat match goal with - | m: _ |- _ => destruct_RiscvMachine m; simpl_MetricRiscvMachine_get_set - end. - eexists. eexists. - split; [eassumption|]. - split; [solve [sidecondition]|]. - split; [solve [sidecondition]|]. - split. { - subst. solve_word_eq word_ok. - } - split; [solve [sidecondition]|]. - split. { - solve_MetricLog. - } - assumption. - - - (* SLit *) - RunInstruction.get_runsTo_valid_for_free. - eapply compile_lit_correct_full. - + sidecondition. - + use_sep_assumption. cbn. - (* ecancel. (* The term "Tree.Leaf (subset (footpr (program iset initialL_pc (compile_lit x v) * Rexec)%sep))" has type "Tree.Tree (set word -> Prop)" while it is expected to have type "Tree.Tree (?map -> Prop)". *) *) - eapply RelationClasses.reflexivity. - + unfold compile_stmt. simpl. ecancel_assumption. - + sidecondition. - + assumption. - + simpl. run1done; - remember (updateMetricsForLiteral v initialL_metrics) as finalMetrics; - symmetry in HeqfinalMetrics; - pose proof update_metrics_for_literal_bounded (width := width) as Hlit; - specialize Hlit with (1 := HeqfinalMetrics); - solve_MetricLog. - - - (* SOp *) - match goal with - | o: Syntax.bopname.bopname |- _ => destruct o - end; - simpl in *; run1det; - rewrite ?word.sru_ignores_hibits, - ?word.slu_ignores_hibits, - ?word.srs_ignores_hibits, - ?word.mulhuu_simpl, - ?word.divu0_simpl, - ?word.modu0_simpl in *; - try solve [run1done]. - simpl_MetricRiscvMachine_get_set. - run1det. run1done; - [match goal with - | H: ?post _ _ _ |- ?post _ _ _ => eqexact H - end | solve_MetricLog..]. - rewrite reduce_eq_to_sub_and_lt. - symmetry. apply map.put_put_same. - - - (* SSet *) - run1det. run1done. - - - (* SIf/Then *) - (* execute branch instruction, which will not jump *) - eapply runsTo_det_step_with_valid_machine; simpl in *; subst. - + assumption. - + simulate'. simpl_MetricRiscvMachine_get_set. - destruct cond; [destruct op | ]; - simpl in *; simp; repeat (simulate'; simpl_bools; simpl); try reflexivity. - + intro V. eapply runsTo_trans; simpl_MetricRiscvMachine_get_set. - * (* use IH for then-branch *) - eapply IHexec; IH_sidecondition. - * (* jump over else-branch *) - simpl. intros. destruct_RiscvMachine middle. simp. subst. - run1det. run1done. - - - (* SIf/Else *) - (* execute branch instruction, which will jump over then-branch *) - eapply runsTo_det_step_with_valid_machine; simpl in *; subst. - + assumption. - + simulate'. - destruct cond; [destruct op | ]; - simpl in *; simp; repeat (simulate'; simpl_bools; simpl); try reflexivity. - + intro V. eapply runsTo_trans; simpl_MetricRiscvMachine_get_set. - * (* use IH for else-branch *) - eapply IHexec; IH_sidecondition. - * (* at end of else-branch, i.e. also at end of if-then-else, just prove that - computed post satisfies required post *) - simpl. intros. destruct_RiscvMachine middle. simp. subst. run1done. - - - (* SLoop/again *) - on hyp[(stmt_not_too_big body1); runsTo] do (fun H => rename H into IH1). - on hyp[(stmt_not_too_big body2); runsTo] do (fun H => rename H into IH2). - on hyp[(stmt_not_too_big (SLoop body1 cond body2)); runsTo] do (fun H => rename H into IH12). - eapply runsTo_trans; simpl_MetricRiscvMachine_get_set. - + (* 1st application of IH: part 1 of loop body *) - eapply IH1; IH_sidecondition. - + simpl in *. simpl. intros. destruct_RiscvMachine middle. simp. subst. - destruct (eval_bcond middle_regs cond) as [condB|] eqn: E. - 2: exfalso; - match goal with - | H: context [_ <> None] |- _ => solve [eapply H; eauto] - end. - destruct condB. - * (* true: iterate again *) - eapply runsTo_det_step_with_valid_machine; simpl in *; subst. - { assumption. } - { simulate'. - destruct cond; [destruct op | ]; - simpl in *; simp; repeat (simulate'; simpl_bools; simpl); try reflexivity. } - { intro V. eapply runsTo_trans; simpl_MetricRiscvMachine_get_set. - - (* 2nd application of IH: part 2 of loop body *) - eapply IH2; IH_sidecondition; simpl_MetricRiscvMachine_get_set; - try eassumption; IH_sidecondition. - - simpl in *. simpl. intros. destruct_RiscvMachine middle. simp. subst. - (* jump back to beginning of loop: *) - run1det. - eapply runsTo_trans; simpl_MetricRiscvMachine_get_set. - + (* 3rd application of IH: run the whole loop again *) - eapply IH12 with (pos := pos); IH_sidecondition; simpl_MetricRiscvMachine_get_set; - try eassumption; IH_sidecondition. - + (* at end of loop, just prove that computed post satisfies required post *) - simpl. intros. destruct_RiscvMachine middle. simp. subst. - run1done. } - * (* false: done, jump over body2 *) - eapply runsTo_det_step_with_valid_machine; simpl in *; subst. - { assumption. } - { simulate'. - destruct cond; [destruct op | ]; - simpl in *; simp; repeat (simulate'; simpl_bools; simpl); try reflexivity. } - { intro V. simpl in *. run1done. } - - - (* SSeq *) - on hyp[(stmt_not_too_big s1); runsTo] do (fun H => rename H into IH1). - on hyp[(stmt_not_too_big s2); runsTo] do (fun H => rename H into IH2). - eapply runsTo_trans. - + eapply IH1; IH_sidecondition. - + simpl. intros. destruct_RiscvMachine middle. simp. subst. - eapply runsTo_trans. - * eapply IH2; IH_sidecondition; simpl_MetricRiscvMachine_get_set; - try eassumption; IH_sidecondition. - * simpl. intros. destruct_RiscvMachine middle. simp. subst. run1done. - - - (* SSkip *) - run1done. - Qed. - -End Proofs. diff --git a/compiler/src/compiler/LowerPipeline.v b/compiler/src/compiler/LowerPipeline.v index dd55cc8ca..8e3b46fef 100644 --- a/compiler/src/compiler/LowerPipeline.v +++ b/compiler/src/compiler/LowerPipeline.v @@ -468,13 +468,13 @@ Section LowerPipeline. Qed. Lemma flat_to_riscv_correct: forall p1 p2, - map.forall_values FlatToRiscvDef.valid_FlatImp_fun p1 -> + map.forall_values FlatToRiscvDef.valid_FlatImp_fun p1 -> riscvPhase p1 = Success p2 -> forall fname t m argvals post, (exists argnames retnames fbody l, map.get p1 fname = Some (argnames, retnames, fbody) /\ - forall l mc, map.of_list_zip argnames argvals = Some l -> - FlatImp.exec isRegZ p1 fbody t m l mc (fun t' m' l' mc' => + map.of_list_zip argnames argvals = Some l /\ + forall mc, FlatImp.exec isRegZ p1 fbody t m l mc (fun t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals)) -> riscv_call p2 fname t m argvals post. From 832919718a7d19ab5f7a6e2b43dbe635ec7b5886 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Thu, 21 Jul 2022 10:53:03 -0400 Subject: [PATCH 15/62] make compilerExamples/SoftmulCompile.v compatible with new FlatImp.exec to fix build error --- compiler/src/compilerExamples/SoftmulCompile.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/compiler/src/compilerExamples/SoftmulCompile.v b/compiler/src/compilerExamples/SoftmulCompile.v index 8e5da0042..ad2a72e3f 100644 --- a/compiler/src/compilerExamples/SoftmulCompile.v +++ b/compiler/src/compilerExamples/SoftmulCompile.v @@ -805,7 +805,7 @@ Section Riscv. unfold FlatToRiscvCommon.compiles_FlatToRiscv_correctly. intros. match goal with - | H: FlatImp.exec.exec _ (FlatImp.SInteract _ _ _) _ _ _ _ _ |- _ => inversion H + | H: FlatImp.exec.exec _ _ (FlatImp.SInteract _ _ _) _ _ _ _ _ |- _ => inversion H end. contradiction. } { intros. reflexivity. } From 24eb5f51a7e97721d55dd52b8302364bba6d62ad Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Thu, 11 Aug 2022 15:36:54 -0400 Subject: [PATCH 16/62] work on spilling pass, current costs for call and interact don't work --- compiler/src/compiler/FlatImp.v | 8 +- compiler/src/compiler/FlattenExpr.v | 2 +- compiler/src/compiler/Spilling.v | 165 +++++++++++++++++++--------- 3 files changed, 117 insertions(+), 58 deletions(-) diff --git a/compiler/src/compiler/FlatImp.v b/compiler/src/compiler/FlatImp.v index 9ece4b535..ecd78280a 100644 --- a/compiler/src/compiler/FlatImp.v +++ b/compiler/src/compiler/FlatImp.v @@ -393,11 +393,11 @@ Module exec. exists l', map.putmany_of_list_zip resvars resvals l = Some l' /\ forall m', map.split m' mKeep mReceive -> post (((mGive, action, argvals), (mReceive, resvals)) :: t) m' l' - (addMetricInstructions 1 - (addMetricStores 1 - (addMetricLoads 2 mc)))) -> + (addMetricInstructions 100 + (addMetricJumps 100 + (addMetricStores 100 + (addMetricLoads 100 mc))))) -> exec (SInteract resvars action argvars) t m l mc post - (* TODO check SInteract case costs *) | call: forall t m l mc binds fname args params rets fbody argvs st0 post outcome, map.get e fname = Some (params, rets, fbody) -> map.getmany_of_list l args = Some argvs -> diff --git a/compiler/src/compiler/FlattenExpr.v b/compiler/src/compiler/FlattenExpr.v index fd99c277b..415b780b9 100644 --- a/compiler/src/compiler/FlattenExpr.v +++ b/compiler/src/compiler/FlattenExpr.v @@ -920,7 +920,7 @@ Section FlattenExpr1. split; eauto. simple apply conj; [eassumption|]. split; [simple eapply map.only_differ_putmany; eassumption|]. - solve_MetricLog. + Fail solve_MetricLog. admit. Admitted. Goal True. idtac "FlattenExpr: flattenStmt_correct_aux done". Abort. diff --git a/compiler/src/compiler/Spilling.v b/compiler/src/compiler/Spilling.v index 14ab6c828..7428ac74f 100644 --- a/compiler/src/compiler/Spilling.v +++ b/compiler/src/compiler/Spilling.v @@ -655,9 +655,7 @@ Section Spilling. 1: { unfold spill_tmp. eapply put_tmp; eauto. } 1: blia. unfold exec.cost_SStore. unfold spill_tmp; cbn. - destr (isRegZ x). - { solve_MetricLog. } - { solve_MetricLog. } + destr (isRegZ x); solve_MetricLog. } blia. - eapply exec.skip. @@ -708,9 +706,9 @@ Section Spilling. Lemma save_ires_reg_correct'': forall e t1 t2 m1 m2 l1 l2 mc2 x v maxvar frame post fpval, related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> fp < x <= maxvar /\ (x < a0 \/ a7 < x) -> - (forall t2' m2' l2' mc2', + (forall t2' m2' l2', related maxvar frame fpval t1 m1 (map.put l1 x v) t2' m2' l2' -> - post t2' m2' l2' mc2') -> + post t2' m2' l2' (if isRegZ x then mc2 else (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc2))))) -> exec e (save_ires_reg x) t2 m2 (map.put l2 (ires_reg x) v) mc2 post. Proof. intros. @@ -724,7 +722,12 @@ Section Spilling. eapply get_sep. ecancel_assumption. - rewrite map.get_put_same. reflexivity. - exact St. - - eapply H1. + - unfold exec.cost_SStore. + assert (isRegZ (spill_tmp 1) = true) by auto; rewrite H. + assert (isRegZ fp = true) by auto; rewrite H0. + clear H H0. + destr (isRegZ x); try blia. + eapply H1. repeat match goal with | |- exists _, _ => eexists | |- _ /\ _ => split @@ -760,6 +763,7 @@ Section Spilling. } blia. - eapply exec.skip. + destr (isRegZ x); try blia. eapply H1. (* even though we did nothing, we have to reconstruct the `related` from the `related` that *) (* held *before* the SOp *) @@ -844,6 +848,41 @@ Section Spilling. intros. apply H. eapply hide_ll_arg_reg_ptsto_core; eassumption. Qed. + Fixpoint cost_set_vars_to_reg_range (args: list Z) (start : Z) (mc : MetricLog) : MetricLog := + match args with + | [] => mc + | x :: xs => if isRegZ x then + addMetricInstructions 1 + (addMetricLoads 1 + (cost_set_vars_to_reg_range xs (start + 1) mc)) + else (addMetricInstructions 1 + (addMetricLoads 1 + (addMetricStores 1 + (cost_set_vars_to_reg_range xs (start + 1) mc)))) + end. + + Lemma cost_set_vars_to_reg_range_commutes: + forall args start mc, + cost_set_vars_to_reg_range args start + (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc))) = + addMetricInstructions 1 + (addMetricLoads 1 (addMetricStores 1 (cost_set_vars_to_reg_range args start mc))). + Proof. + induction args; intros; cbn; try trivial. + destr (isRegZ a); cbn; rewrite IHargs; trivial. + Qed. + + Lemma cost_set_vars_to_reg_range_commutes': + forall args start mc, + cost_set_vars_to_reg_range args start + (addMetricInstructions 1 (addMetricLoads 1 mc)) = + addMetricInstructions 1 + (addMetricLoads 1 (cost_set_vars_to_reg_range args start mc)). + Proof. + induction args; intros; cbn; try trivial. + destr (isRegZ a); cbn; rewrite IHargs; trivial. + Qed. + Lemma set_vars_to_reg_range_correct: forall args start argvs e t1 t2 m1 m2 l1 l1' l2 mc2 maxvar frame post fpval, related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> @@ -855,6 +894,7 @@ Section Spilling. Forall (fun x => fp < x <= maxvar /\ (x < a0 \/ a7 < x)) args -> (forall m2' l2' mc2', related maxvar frame fpval t1 m1 l1' t2 m2' l2' -> + mc2' = (cost_set_vars_to_reg_range args start mc2) -> post t2 m2' l2' mc2') -> exec e (set_vars_to_reg_range args start) t2 m2 l2 mc2 post. Proof. @@ -874,50 +914,64 @@ Section Spilling. { eassumption. } eapply IHargs; try eassumption; try blia. (* establish related for IH: *) - unfold related. - eexists (map.put lStack a v), lRegs, _. - ssplit. - { reflexivity. } - { ecancel_assumption. } - { eassumption. } - { intros. rewrite map.get_put_dec in H. destr (a =? x0). 1: blia. eauto. } - { apply sep_comm. eapply sep_eq_put. 1: apply sep_comm; assumption. - intros lRegs' w ? G. subst lRegs'. - match goal with H: _ |- _ => specialize H with (1 := G) end. blia. } - { eassumption. } - { intros b A0 w B0. + * unfold related. + eexists (map.put lStack a v), lRegs, _. + ssplit. + { reflexivity. } + { ecancel_assumption. } + { eassumption. } + { intros. rewrite map.get_put_dec in H. destr (a =? x0). 1: blia. eauto. } + { apply sep_comm. eapply sep_eq_put. 1: apply sep_comm; assumption. + intros lRegs' w ? G. subst lRegs'. + match goal with H: _ |- _ => specialize H with (1 := G) end. blia. } + { eassumption. } + { intros b A0 w B0. rewrite map.get_put_dec in B0. destr (a =? b). 1: congruence. match goal with H: _ |- _ => eapply H end. 1: blia. match goal with H: _ |- _ => eapply H end. 1: blia. assumption. } - { blia. } + { blia. } + * intros. apply H6; auto. + cbn in *. destr (isRegZ start); try blia; destr (isRegZ a); try blia. + rewrite H0. rewrite cost_set_vars_to_reg_range_commutes. trivial. + eapply exec.set. { eassumption. } eapply IHargs; try eassumption; try blia. 2: { eapply map.getmany_of_list_put_diff. 2: eassumption. eapply List.not_In_Z_seq. blia. } - unfold related. eexists lStack, (map.put lRegs a v), _. - ssplit. - { reflexivity. } - { ecancel_assumption. } - { intros. rewrite map.get_put_dec in H. destr (a =? x). 1: blia. eauto. } - { eassumption. } - { eapply sep_eq_put. 1: assumption. - intros lStack' w ? G. subst lStack'. - match goal with H: _ |- _ => specialize H with (1 := G) end. blia. } - { apply sep_assoc. eapply sep_eq_put. 1: ecancel_assumption. - unfold ptsto, arg_regs. - intros l w (l_arg_regs & l_fpval & (? & ?) & ? & ?) G. subst. + * unfold related. eexists lStack, (map.put lRegs a v), _. + ssplit. + { reflexivity. } + { ecancel_assumption. } + { intros. rewrite map.get_put_dec in H. destr (a =? x). 1: blia. eauto. } + { eassumption. } + { eapply sep_eq_put. 1: assumption. + intros lStack' w ? G. subst lStack'. + match goal with H: _ |- _ => specialize H with (1 := G) end. blia. } + { apply sep_assoc. eapply sep_eq_put. 1: ecancel_assumption. + unfold ptsto, arg_regs. + intros l w (l_arg_regs & l_fpval & (? & ?) & ? & ?) G. subst. rewrite map.get_putmany_dec, map.get_put_dec, map.get_empty in G. destr (fp =? a). 1: unfold fp; blia. match goal with H: _ |- _ => specialize H with (1 := G) end. unfold a0, a7 in *. blia. } - { assumption. } - { assumption. } + { assumption. } + { assumption. } + * intros. apply H6; auto. + cbn in *. destr (isRegZ start); try blia; destr (isRegZ a); try blia. + rewrite H0. rewrite cost_set_vars_to_reg_range_commutes'. trivial. Qed. + Fixpoint cost_set_reg_range_to_vars (start : Z) (args: list Z) (mc : MetricLog) : MetricLog := + match args with + | [] => mc + | x :: xs => if isRegZ x then + addMetricInstructions 1 (addMetricLoads 1 (cost_set_reg_range_to_vars (start + 1) xs mc)) + else (addMetricInstructions 1 (addMetricLoads 2 (cost_set_reg_range_to_vars (start + 1) xs mc))) + end. + Lemma set_reg_range_to_vars_correct: forall args argvs start e t1 t2 m1 m2 l1 l2 mc2 maxvar frame post fpval, related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> @@ -929,12 +983,14 @@ Section Spilling. (forall l2' mc2', related maxvar frame fpval t1 m1 l1 t2 m2 l2' -> map.getmany_of_list l2' (List.unfoldn (Z.add 1) (List.length args) start) = Some argvs -> + mc2' = (cost_set_reg_range_to_vars start args mc2) -> post t2 m2 l2' mc2') -> exec e (set_reg_range_to_vars start args) t2 m2 l2 mc2 post. Proof. induction args; intros. - - simpl. eapply exec.skip. eapply H5. 1: eassumption. simpl. - destruct argvs. 1: reflexivity. discriminate. + - simpl. eapply exec.skip. eapply H5. 1: eassumption. + + simpl. destruct argvs. 1: reflexivity. discriminate. + + trivial. - simpl. unfold set_reg_to_var, stack_loc. destruct argvs as [|v vs]. { unfold map.getmany_of_list in H4. cbn in H4. simp. @@ -969,6 +1025,8 @@ Section Spilling. ++ rewrite Z.add_comm. eapply map.getmany_of_list_put_diff. 2: eassumption. eauto using List.not_In_Z_seq with zarith. + -- cbn. destr (isRegZ start); destr (isRegZ a); cbn in *; try blia. + rewrite H6; trivial. + eapply exec.seq_cps. eapply IHargs; try eassumption; try blia. intros. @@ -988,13 +1046,15 @@ Section Spilling. - rewrite Z.add_comm. eapply map.getmany_of_list_put_diff. 2: eassumption. eauto using List.not_In_Z_seq with zarith. } - unfold related. - repeat match goal with - | |- exists _, _ => eexists - | |- _ /\ _ => split - | |- _ => eassumption || reflexivity - end. - eapply put_arg_reg; try eassumption. blia. + -- unfold related. + repeat match goal with + | |- exists _, _ => eexists + | |- _ /\ _ => split + | |- _ => eassumption || reflexivity + end. + eapply put_arg_reg; try eassumption. blia. + -- cbn. destr (isRegZ start); destr (isRegZ a); cbn in *; try blia. + rewrite H6; trivial. Qed. Lemma grow_related_mem: forall maxvar frame t1 mSmall1 l1 t2 mSmall2 l2 mStack mCombined2 fpval, @@ -1285,7 +1345,7 @@ Section Spilling. 2: eapply Forall_le_max. cbv beta. subst maxvar'. clear. blia. } - intros mL4 lFL4 mcL4 R. + intros mL4 lFL4 mcL4 R Hcost. eapply exec.seq_cps. eapply exec.weaken. { eapply IHexec. 1: apply Ex. 2: exact R. @@ -1321,7 +1381,7 @@ Section Spilling. subst maxvar'. clear. blia. } { eassumption. } rename R into R0. - intros lFL6 mcL6 R GM. + intros lFL6 mcL6 R GM HCost. (* prove that if we remove the additional stack provided by exec.stackalloc and store the result vars back into the arg registers, the postcondition holds *) unfold related in R. fwd. rename lStack into lStack5, lRegs into lRegs5. @@ -1357,7 +1417,7 @@ Section Spilling. Lemma iarg_reg_isReg: forall i a, - (i <= 10) -> + (i <= 20) -> (isRegZ (iarg_reg i a) = true). Proof. intros. unfold isRegZ, iarg_reg. destr (32 <=? a); unfold spill_tmp; blia. @@ -1402,7 +1462,7 @@ Section Spilling. - (* exec.interact *) eapply exec.seq_cps. eapply set_reg_range_to_vars_correct; try eassumption; try (unfold a0, a7; blia). - intros *. intros R GM. clear l2 H4. + intros *. intros R GM CSet. clear l2 H4. unfold related in R. fwd. spec (subst_split (ok := mem_ok) m) as A. 1: eassumption. 1: ecancel_assumption. @@ -1447,9 +1507,8 @@ Section Spilling. unfold sep at 1 in C. destruct C as (mKeepL' & mRest & SC & ? & _). subst mKeepL'. move H2 at bottom. unfold map.split in H2. fwd. eapply map.shrink_disjoint_l; eassumption. } - - admit. - } + cbn in *. rewrite H5; rewrite CSet. admit. (* currently wrong *) + } (* related for set_vars_to_reg_range_correct: *) unfold related. eexists _, _, _. ssplit. @@ -1495,7 +1554,7 @@ Section Spilling. apply_in_hyps @map.getmany_of_list_length. apply_in_hyps @map.putmany_of_list_zip_sameLength. eapply set_reg_range_to_vars_correct; try eassumption || (unfold a0, a7 in *; blia). - intros lCL2 ? ? ?. + intros lCL2 ? ? ? ?. assert (bytes_per_word = 4 \/ bytes_per_word = 8) as B48. { unfold bytes_per_word. destruct width_cases as [E' | E']; rewrite E'; cbv; auto. } @@ -1564,7 +1623,7 @@ Section Spilling. 2: eapply Forall_le_max. cbv beta. subst maxvar'. clear. blia. } - intros mL4 lFL4 mcL4 R. + intros mL4 lFL4 mcL4 R CSet. eapply exec.seq_cps. eapply exec.weaken. { eapply IHexec. 2: exact R. @@ -1604,7 +1663,7 @@ Section Spilling. subst maxvar'. clear. blia. } { eassumption. } rename R into R0. - intros lFL6 mcL6 R GM. + intros lFL6 mcL6 R GM ?. (* prove that if we remove the additional stack provided by exec.stackalloc and store the result vars back into the caller's registers, states are still related and postcondition holds *) @@ -1740,7 +1799,7 @@ Section Spilling. end. 1,4,3,2: eassumption. unfold exec.cost_SStackalloc in *; repeat isReg_helper. - destr (isRegZ x); try solve_MetricLog. all: admit. + destr (isRegZ x); solve_MetricLog. - (* exec.lit *) eapply exec.seq_cps. eapply exec.lit. eapply save_ires_reg_correct. From 6bd3c058249d0a3625e06f698e262cdd6b06fce7 Mon Sep 17 00:00:00 2001 From: nathan-sheffield Date: Mon, 9 Oct 2023 14:29:02 -0400 Subject: [PATCH 17/62] cleaned up as much of the mess as i could for now --- bedrock2/src/bedrock2/MetricLogging.v | 67 +- bedrock2/src/bedrock2/MetricLoops.v | 732 ++++++++++++++++++ .../src/bedrock2/MetricWeakestPrecondition.v | 339 ++++++++ .../MetricWeakestPreconditionProperties.v | 324 ++++++++ .../NathanDidntNoticeMetricSemantics.v | 385 +++++++++ bedrock2/src/bedrock2/Notations.v | 21 + bedrock2/src/bedrock2/ProgramLogic.v | 12 +- bedrock2/src/bedrock2Examples/tckmnipow.v | 469 +++++++++++ 8 files changed, 2337 insertions(+), 12 deletions(-) create mode 100644 bedrock2/src/bedrock2/MetricLoops.v create mode 100644 bedrock2/src/bedrock2/MetricWeakestPrecondition.v create mode 100644 bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v create mode 100644 bedrock2/src/bedrock2/NathanDidntNoticeMetricSemantics.v create mode 100644 bedrock2/src/bedrock2/Notations.v create mode 100644 bedrock2/src/bedrock2Examples/tckmnipow.v diff --git a/bedrock2/src/bedrock2/MetricLogging.v b/bedrock2/src/bedrock2/MetricLogging.v index 451dacc70..29644b06b 100644 --- a/bedrock2/src/bedrock2/MetricLogging.v +++ b/bedrock2/src/bedrock2/MetricLogging.v @@ -29,6 +29,8 @@ Section Riscv. Definition subMetricLoads n log := withLoads (loads log - n) log. Definition subMetricJumps n log := withJumps (jumps log - n) log. + Definition metricAdd(metric: MetricLog -> Z) finalM initialM : Z := + Z.add (metric finalM) (metric initialM). Definition metricSub(metric: MetricLog -> Z) finalM initialM : Z := Z.sub (metric finalM) (metric initialM). @@ -40,8 +42,16 @@ Section Riscv. (op loads initialM finalM) (op jumps initialM finalM). + Definition metricsAdd := metricsOp metricAdd. Definition metricsSub := metricsOp metricSub. + Definition metricsMul (n : Z) (m : MetricLog) := + mkMetricLog + (n * instructions m) + (n * stores m) + (n * loads m) + (n * jumps m). + Definition metricLeq(metric: MetricLog -> Z) m1 m2: Prop := (metric m1) <= (metric m2). @@ -51,20 +61,17 @@ Section Riscv. metricLeq loads m1 m2 /\ metricLeq jumps m1 m2. - Definition metricMax(metric: MetricLog -> Z) m1 m2: Z := - Z.max (metric m1) (metric m2). - - Definition metricsMax := metricsOp metricMax. End Riscv. -Declare Scope MetricH_scope. Bind Scope MetricH_scope with MetricLog. Delimit Scope MetricH_scope with metricsH. Infix "<=" := metricsLeq : MetricH_scope. +Infix "+" := metricsAdd : MetricH_scope. Infix "-" := metricsSub : MetricH_scope. +Infix "*" := metricsMul : MetricH_scope. -#[export] Hint Unfold +#[global] Hint Unfold withInstructions withLoads withStores @@ -78,8 +85,11 @@ Infix "-" := metricsSub : MetricH_scope. subMetricStores subMetricJumps metricsOp + metricAdd + metricsAdd metricSub metricsSub + metricsMul metricLeq metricsLeq : unf_metric_log. @@ -107,3 +117,48 @@ Ltac solve_MetricLog := repeat unfold_MetricLog; repeat simpl_MetricLog; blia. + +Module MetricArith. + + Open Scope MetricH_scope. + + Lemma mul_sub_distr_r : forall n m p, (n - m) * p = n * p - m * p. + Proof. intros. unfold_MetricLog. f_equal; apply Z.mul_sub_distr_r. Qed. + + Lemma add_sub_swap : forall n m p, n + m - p = n - p + m. + Proof. intros. unfold_MetricLog. f_equal; apply Z.add_sub_swap. Qed. + + Lemma le_add_le_sub_r : forall n m p, n + p <= m <-> n <= m - p. + Proof. solve_MetricLog. Qed. + + Lemma le_trans : forall n m p, n <= m -> m <= p -> n <= p. + Proof. solve_MetricLog. Qed. + +End MetricArith. + +Lemma applyAddInstructions n a b c d : addMetricInstructions n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a+n; stores := b; loads := c; jumps := d |}. Proof. auto. Qed. +Lemma applyAddStores n a b c d : addMetricStores n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a; stores := b+n; loads := c; jumps := d |}. Proof. auto. Qed. +Lemma applyAddLoads n a b c d : addMetricLoads n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a; stores := b; loads := c+n; jumps := d |}. Proof. auto. Qed. +Lemma applyAddJumps n a b c d : addMetricJumps n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a; stores := b; loads := c; jumps := d+n |}. Proof. auto. Qed. + +Ltac applyAddMetricsGoal := ( + repeat ( + try rewrite applyAddInstructions; + try rewrite applyAddStores; + try rewrite applyAddLoads; + try rewrite applyAddJumps + ); + repeat rewrite <- Z.add_assoc; + cbn [Z.add Pos.add Pos.succ] + ). + +Ltac applyAddMetrics H := ( + repeat ( + try rewrite applyAddInstructions in H; + try rewrite applyAddStores in H; + try rewrite applyAddLoads in H; + try rewrite applyAddJumps in H + ); + repeat rewrite <- Z.add_assoc in H; + cbn [Z.add Pos.add Pos.succ] in H +). diff --git a/bedrock2/src/bedrock2/MetricLoops.v b/bedrock2/src/bedrock2/MetricLoops.v new file mode 100644 index 000000000..d8333a8ce --- /dev/null +++ b/bedrock2/src/bedrock2/MetricLoops.v @@ -0,0 +1,732 @@ +Require Import coqutil.Datatypes.PrimitivePair coqutil.Datatypes.HList coqutil.dlet. +Require Import Coq.Classes.Morphisms BinIntDef. +Require Import coqutil.Macros.unique coqutil.Map.Interface coqutil.Word.Interface. Import map. +Require Import coqutil.Word.Bitwidth. +Require Import coqutil.Map.Properties. +Require Import coqutil.Tactics.destr. +From bedrock2 Require Import Map.Separation Map.SeparationLogic. +From bedrock2 Require Import Syntax MetricSemantics Markers. +From bedrock2 Require Import MetricWeakestPrecondition MetricWeakestPreconditionProperties. + +Require Import bedrock2.MetricLogging. + +Section Loops. + Context {width: Z} {BW: Bitwidth width} {word: word.word width} {mem: map.map word Byte.byte}. + Context {locals: map.map String.string word}. + Context {ext_spec: ExtSpec}. + Context {word_ok : word.ok word} {mem_ok : map.ok mem}. + Context {locals_ok : map.ok locals}. + Context {ext_spec_ok : Semantics.ext_spec.ok ext_spec}. + + Context {fs : env}. + Let call := fs. + + Lemma wp_while: forall e c t m l (post: _ -> _ -> _ -> Prop), + (exists measure (lt:measure->measure->Prop) (inv:measure->trace->mem->locals->Prop), + Coq.Init.Wf.well_founded lt /\ + (exists v, inv v t m l) /\ + (forall v t m l, inv v t m l -> + exists b, dexpr m l e b /\ + (word.unsigned b <> 0%Z -> cmd call c t m l (fun t' m l => + exists v', inv v' t' m l /\ lt v' v)) /\ + (word.unsigned b = 0%Z -> post t m l))) -> + cmd call (cmd.while e c) t m l post. + Proof. + intros. destruct H as (measure & lt & inv & Hwf & HInit & Hbody). + destruct HInit as (v0 & HInit). + revert t m l HInit. pattern v0. revert v0. + eapply (well_founded_ind Hwf). intros. + specialize Hbody with (1 := HInit). destruct Hbody as (b & Hb & Ht & Hf). + eapply expr_sound in Hb. destruct Hb as (b' & Hb & ?). subst b'. + destr.destr (Z.eqb (word.unsigned b) 0). + - specialize Hf with (1 := E). eapply exec.while_false; eassumption. + - specialize Ht with (1 := E). eapply sound_cmd in Ht. + eapply exec.while_true; eauto. + cbv beta. intros * (v' & HInv & HLt). eapply sound_cmd. eauto. + Qed. + + Lemma tailrec_localsmap_1ghost + {e c t} {m: mem} {l} {post : trace -> mem -> locals -> Prop} + {measure: Type} {Ghost: Type} + (P Q: measure -> Ghost -> trace -> mem -> locals -> Prop) + (lt: measure -> measure -> Prop) + (Hwf: well_founded lt) + (v0: measure) (g0: Ghost) + (Hpre: P v0 g0 t m l) + (Hbody: forall v g t m l, + P v g t m l -> + exists br, expr m l e (eq br) /\ + (word.unsigned br <> 0%Z -> cmd call c t m l + (fun t' m' l' => exists v' g', + P v' g' t' m' l' /\ + lt v' v /\ + (forall t'' m'' l'', Q v' g' t'' m'' l'' -> Q v g t'' m'' l''))) /\ + (word.unsigned br = 0%Z -> Q v g t m l)) + (Hpost: forall t m l, Q v0 g0 t m l -> post t m l) + : cmd call (cmd.while e c) t m l post. + Proof. + eapply wp_while. + eexists measure, lt, (fun v t m l => + exists g, P v g t m l /\ forall t'' m'' l'', Q v g t'' m'' l'' -> Q v0 g0 t'' m'' l''). + split; [assumption|]. + split; [solve[eauto]|]. + intros vi ti mi li (gi & HPi & HQimpl). + specialize (Hbody vi gi ti mi li HPi). + destruct Hbody as (br & ? & Hbody). exists br; split; [assumption|]. + destruct Hbody as (Htrue & Hfalse). split; intros Hbr; + [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. + { eapply Proper_cmd; [ |eapply Hpc]. + intros tj mj lj (vj& gj & HPj & Hlt & Qji); eauto 9. } + { eauto. } + Qed. + + Lemma tailrec_localsmap_1ghost_parameterized_finalpost + {e c rest t} {m: mem} {l} + {measure: Type} {Ghost: Type} + (P Q: measure -> Ghost -> trace -> mem -> locals -> Prop) + (lt: measure -> measure -> Prop) + (Hwf: well_founded lt) + (v0: measure) (g0: Ghost) + (Hpre: P v0 g0 t m l) + (Hbody: forall v g t m l, + P v g t m l -> + exists br, expr m l e (eq br) /\ + (word.unsigned br <> 0%Z -> cmd call c t m l + (fun t' m' l' => exists v' g', + P v' g' t' m' l' /\ + lt v' v /\ + (forall t'' m'' l'', Q v' g' t'' m'' l'' -> Q v g t'' m'' l''))) /\ + (word.unsigned br = 0%Z -> cmd call rest t m l (Q v g))) + : cmd call (cmd.seq (cmd.while e c) rest) t m l (Q v0 g0). + Proof. + cbn. eapply tailrec_localsmap_1ghost with + (Q := fun v g t m l => cmd call rest t m l (Q v g)). + 1: eassumption. + 1: exact Hpre. + 2: intros *; exact id. + intros vi gi ti mi li HPi. + specialize (Hbody vi gi ti mi li HPi). + destruct Hbody as (br & ? & Hbody). exists br; split; [assumption|]. + destruct Hbody as (Htrue & Hfalse). split; intros Hbr; + [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. + { eapply Proper_cmd; [ |eapply Hpc]. + intros tj mj lj (vj& gj & HPj & Hlt & Qji). do 2 eexists. + split. 1: eassumption. split. 1: assumption. + intros. + eapply Proper_cmd. + 2: eassumption. + intros tk mk lk HH. eapply Qji. assumption. } + eapply Hpc. + Qed. + + (* marking logical connectives with the source file they were used in for limiting unfolding *) + Local Notation "A /\ B" := (Markers.split (A /\ B)). + Local Notation "A /\ B" := (Markers.split (A /\ B)) : type_scope. + + (* shallow reflection for resolving map accesses during symbolic execution *) + (* each lemma below is duplicated for various levels of use of this trick *) + Definition reconstruct (variables:list String.string) (values:tuple word (length variables)) : locals := + map.putmany_of_tuple (tuple.of_list variables) values map.empty. + Fixpoint gather (variables : list String.string) (l : locals) : option (locals * tuple word (length variables)) := + match variables with + | nil => Some (l, tt) + | cons x xs' => + match map.get l x with + | None => None + | Some v => + match gather xs' (map.remove l x) with + | None => None + | Some (l, vs') => Some (l, (pair.mk v vs')) + end + end + end. + + Lemma putmany_gather ks vs m me (H : gather ks m = Some (me, vs)) : + map.putmany_of_tuple (tuple.of_list ks) vs me = m. + Proof. + revert H; revert me; revert m; revert vs; induction ks; cbn [gather map.putmany_of_list]; intros. + { inversion H; subst. exact eq_refl. } + repeat match type of H with context[match ?x with _ => _ end] => destruct x eqn:? end; + repeat (match goal with H : _ |- _ => eapply IHks in H end); inversion H; subst; clear H. + cbn [map.putmany_of_tuple tuple.of_list length]. + match goal with H : _ |- _ => rewrite H; clear H end. + assert (map.get m a = Some r -> put (remove m a) a r = m). { + intro A. + apply map_ext. + intro k. + erewrite map.get_put_dec. + destr (String.eqb a k); try congruence. + rewrite map.get_remove_diff; congruence. + } + auto. + Qed. + + Definition enforce (variables : list String.string) (values:tuple word (length variables)) (l:locals) : Prop := + match gather variables l with + | None => False + | Some (remaining, r) => values = r /\ remaining = map.empty + end. + Lemma reconstruct_enforce variables ll lm (H : enforce variables ll lm) : lm = reconstruct variables ll. + progress cbv [enforce] in H. + repeat match type of H with context[match ?x with _ => _ end] => destruct x eqn:? end; + destruct H; subst. + symmetry. eapply putmany_gather. assumption. + Qed. + + Lemma hlist_forall_foralls: forall (argts : polymorphic_list.list Type) (P : hlist argts -> Prop), (forall x : hlist argts, P x) -> hlist.foralls P. + Proof. induction argts; cbn; auto. Qed. + + Import pair. + + Lemma while_localsmap + {e c t l mc} {m : mem} + {measure : Type} (invariant:_->_->_-> _ -> _ -> Prop) + {lt} (Hwf : well_founded lt) (v0 : measure) + {post : _->_->_-> _ -> Prop} + (Hpre : invariant v0 t m l mc) + (Hbody : forall v t m l mc, + invariant v t m l mc -> + exists br mcnew, expr m l e mc (eq (Datatypes.pair br mcnew)) /\ + (word.unsigned br <> 0 -> +<<<<<<< HEAD + cmd fs c t m l (fun t m l => exists v', invariant v' t m l /\ lt v' v)) /\ + (word.unsigned br = 0 -> post t m l)) + : cmd fs (cmd.while e c) t m l post. + Proof. + eapply wp_while. +======= + cmd call c t m l mcnew (fun t m l mcnew' => exists v', invariant v' t m l (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mcnew'))) /\ lt v' v)) /\ + (word.unsigned br = 0 -> post t m l (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mcnew))))) + : cmd call (cmd.while e c) t m l mc post. + Proof. + +>>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) + eexists measure, lt, invariant. + split. 1: exact Hwf. + split. 1: eauto. + + exact Hbody. + Qed. + + Lemma while + {e c t l mc} {m : mem} + (variables : list String.string) + {localstuple : tuple word (length variables)} + {measure : Type} (invariant:_->_->_->_-> ufunc word (length variables) Prop) + {lt} (Hwf : well_founded lt) (v0 : measure) + {post : _->_->_-> _ -> Prop} + (Pl : enforce variables localstuple l) + (Hpre : tuple.apply (invariant v0 t m mc) localstuple) + (Hbody : forall v t m mc, tuple.foralls (fun localstuple => + tuple.apply (invariant v t m mc) localstuple -> + let l := reconstruct variables localstuple in + exists br mcnew, expr m l e mc (eq (Datatypes.pair br mcnew)) /\ + (word.unsigned br <> 0 -> + cmd call c t m l mcnew (fun t m l mcnew' => + Markers.unique (Markers.left (tuple.existss (fun localstuple => + enforce variables localstuple l /\ + Markers.right (Markers.unique (exists v', + tuple.apply (invariant v' t m (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mcnew')))) localstuple /\ lt v' v))))))) /\ + (word.unsigned br = 0 -> post t m l (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mcnew)))))) + : cmd call (cmd.while e c) t m l mc post. + Proof. + eapply (while_localsmap (fun v t m l mc => + exists localstuple, enforce variables localstuple l /\ + tuple.apply (invariant v t m mc) localstuple)); + unfold Markers.split; eauto. + intros vi ti mi li mci (?&X&Y). + specialize (Hbody vi ti mi mci). + eapply hlist.foralls_forall in Hbody. + specialize (Hbody Y). + rewrite <-(reconstruct_enforce _ _ _ X) in Hbody. + destruct Hbody as (br & mcnew & Cond & Again & Done). + exists br. exists mcnew. split; [exact Cond|]. + + split; [|exact Done]. + intro NE. specialize (Again NE). + eapply Proper_cmd; [ |eapply Again]. + cbv [Morphisms.pointwise_relation Basics.impl Markers.right Markers.unique Markers.left] in *. + intros t' m' l' mc' Ex. + eapply hlist.existss_exists in Ex. cbv beta in Ex. destruct Ex as (ls & E & v' & Inv' & LT). + eauto. + Qed. + + Lemma tailrec + {e c t localsmap mc} {m : mem} + (ghosttypes : polymorphic_list.list Type) + (variables : list String.string) + {l0 : tuple word (length variables)} + {Pl : enforce variables l0 localsmap} + {post : _->_->_->_-> Prop} + {measure : Type} (spec:_->HList.arrows ghosttypes (_->_->ufunc word (length variables) (MetricLog -> Prop*(_->_->ufunc word (length variables) (MetricLog -> Prop))))) lt + (Hwf : well_founded lt) + (v0 : measure) + : hlist.foralls (fun (g0 : hlist ghosttypes) => forall + (Hpre : (tuple.apply (hlist.apply (spec v0) g0 t m) l0 mc).(1)) + + (Hbody : forall v, hlist.foralls (fun g => forall t m mc, tuple.foralls (fun l => + @dlet _ (fun _ => Prop) (reconstruct variables l) (fun localsmap : locals => + match tuple.apply (hlist.apply (spec v) g t m) l mc with S_ => + S_.(1) -> + Markers.unique (Markers.left (exists br mc', expr m localsmap e mc (eq (Datatypes.pair br mc')) /\ Markers.right ( + (word.unsigned br <> 0%Z -> cmd call c t m localsmap mc' + (fun t' m' localsmap' mc'' => + Markers.unique (Markers.left (hlist.existss (fun l' => enforce variables l' localsmap' /\ Markers.right ( + Markers.unique (Markers.left (hlist.existss (fun g' => exists v', + match tuple.apply (hlist.apply (spec v') g' t' m') l' (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc''))) with S' => + S'.(1) /\ Markers.right ( + lt v' v /\ + forall T M, hlist.foralls (fun L => forall MC, tuple.apply (S'.(2) T M) L MC -> tuple.apply (S_.(2) T M) L MC)) end))))))))) /\ + (word.unsigned br = 0%Z -> tuple.apply (S_.(2) t m) l (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mc')))))))end)))) + (Hpost : match (tuple.apply (hlist.apply (spec v0) g0 t m) l0 mc).(2) with Q0 => forall t m mc, hlist.foralls (fun l => tuple.apply (Q0 t m) l mc -> post t m (reconstruct variables l) mc)end) + , cmd call (cmd.while e c) t m localsmap mc post ). + Proof. + eapply hlist_forall_foralls; intros g0 **. +<<<<<<< HEAD + eapply wp_while. + eexists measure, lt, (fun vi ti mi localsmapi => +======= + eexists measure, lt, (fun vi ti mi localsmapi mci => +>>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) + exists gi li, localsmapi = reconstruct variables li /\ + match tuple.apply (hlist.apply (spec vi) gi ti mi) li mci with S_ => + S_.(1) /\ forall T M L MC, tuple.apply (S_.(2) T M) L MC-> + tuple.apply ((tuple.apply (hlist.apply (spec v0) g0 t m) l0 mc).(2) T M) L MC end). + cbv [Markers.split Markers.left Markers.right] in *. + split; [assumption|]. + split. { exists v0, g0, l0. split. 1: eapply reconstruct_enforce; eassumption. split; eauto. } + intros vi ti mi lmapi mci (gi&?&?&?&Qi); subst. + destruct (hlist.foralls_forall (hlist.foralls_forall (Hbody vi) gi ti mi mci) _ ltac:(eassumption)) as (br&?&?&X). + + + exists br; exists x0; split; [assumption|]. destruct X as (Htrue&Hfalse). split; intros Hbr; + [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. +<<<<<<< HEAD + { eapply Proper_cmd; [ |eapply Hpc]. + intros tj mj lmapj Hlj; eapply hlist.existss_exists in Hlj. +======= + { eapply Proper_cmd; [reflexivity..| | |eapply Hpc]. + { eapply Proper_call; firstorder idtac. } + intros tj mj lmapj mcj Hlj; eapply hlist.existss_exists in Hlj. +>>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) + destruct Hlj as (lj&Elj&HE); eapply reconstruct_enforce in Elj; subst lmapj. + eapply hlist.existss_exists in HE. destruct HE as (l&?&?&?&HR). + pose proof fun T M => hlist.foralls_forall (HR T M); clear HR. + eauto 9. } + { pose proof fun t m mc => hlist.foralls_forall (Hpost t m mc); clear Hpost; eauto. } + Qed. + + Lemma tailrec_localsmap + {e c t} {m : mem} {l mc} {post : _ -> _->_->_-> Prop} + {measure : Type} (spec:_->_->_->_->_->(Prop*(_->_->_->_-> Prop))) lt + (Hwf : well_founded lt) + (v0 : measure) (P0 := (spec v0 t m l mc).(1)) (Hpre : P0) + (Q0 := (spec v0 t m l mc).(2)) + (Hbody : forall v t m l mc, + let S := spec v t m l mc in let (P, Q) := S in + P -> + exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ + (word.unsigned br <> 0%Z -> cmd call c t m l mc' + (fun t' m' l' mc''=> exists v', + let S' := spec v' t' m' l' (addMetricInstructions 2 + (addMetricLoads 2 + (addMetricJumps 1 mc''))) in let '(P', Q') := S' in + P' /\ + lt v' v /\ + forall T M L MC, Q' T M L MC -> Q T M L MC)) /\ + (word.unsigned br = 0%Z -> Q t m l (addMetricInstructions 1 + (addMetricLoads 1 + (addMetricJumps 1 mc'))))) + (Hpost : forall t m l mc, Q0 t m l mc -> post t m l mc) + : cmd call (cmd.while e c) t m l mc post. + Proof. +<<<<<<< HEAD + eapply wp_while. + eexists measure, lt, (fun v t m l => + let S := spec v t m l in let '(P, Q) := S in + P /\ forall T M L, Q T M L -> Q0 T M L). +======= + eexists measure, lt, (fun v t m l mc => + let S := spec v t m l mc in let '(P, Q) := S in + P /\ forall T M L MC, Q T M L MC -> Q0 T M L MC). +>>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) + split; [assumption|]. + cbv [Markers.split] in *. + split; [solve[eauto]|]. + intros vi ti mi li mci (?&Qi). + destruct (Hbody _ _ _ _ _ ltac:(eassumption)) as (br&?&?&X); exists br; eexists; split; [eassumption|]. + destruct X as (Htrue&Hfalse). split; intros Hbr; + [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. +<<<<<<< HEAD + { eapply Proper_cmd; [ |eapply Hpc]. + intros tj mj lj (vj&dP&?&dQ); eauto 9. } +======= + { eapply Proper_cmd; [reflexivity..| | |eapply Hpc]. + { eapply Proper_call; firstorder idtac. } + intros tj mj lj mcj (vj&dP&?&dQ); eauto 9. } +>>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) + { eauto. } + Qed. + + Definition with_bottom {T} R (x y : option T) := + match x, y with + | None, Some _ => True + | Some x, Some y => R x y + | _, _ => False + end. + Lemma well_founded_with_bottom {T} R (H : @well_founded T R) : well_founded (with_bottom R). + Proof. + intros [x|]; cycle 1. + { constructor; intros [] HX; cbv [with_bottom] in HX; contradiction. } + pattern x. revert x. eapply (@well_founded_ind _ _ H). intros. + constructor. intros [y|] pf; eauto. + constructor. intros [] []. + Qed. + + Lemma atleastonce_localsmap + {e c t} {m : mem} {l mc} {post : _->_->_->_-> Prop} + {measure : Type} (invariant:_->_->_->_->_->Prop) lt + (Hwf : well_founded lt) + (Henter : exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ (word.unsigned br = 0%Z -> post t m l (addMetricInstructions 1 + (addMetricLoads 1 + (addMetricJumps 1 mc')))) + /\ (word.unsigned br <> 0 -> exists v', invariant v' t m l (addMetricInstructions 2 + (addMetricLoads 2 + (addMetricJumps 1 mc')))) + ) + (v0 : measure) (Hpre : invariant v0 t m l mc) + (Hbody : forall v t m l mc, invariant v t m l mc -> + cmd call c t m l mc (fun t m l mc => + exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ + (word.unsigned br <> 0 -> exists v', invariant v' t m l (addMetricInstructions 2 + (addMetricLoads 2 + (addMetricJumps 1 mc'))) /\ lt v' v) /\ + (word.unsigned br = 0 -> post t m l (addMetricInstructions 1 + (addMetricLoads 1 + (addMetricJumps 1 mc')))))) + : cmd call (cmd.while e c) t m l mc post. + Proof. +<<<<<<< HEAD + eapply wp_while. + eexists (option measure), (with_bottom lt), (fun ov t m l => + exists br, expr m l e (eq br) /\ + ((word.unsigned br <> 0 -> exists v, ov = Some v /\ invariant v t m l) /\ + (word.unsigned br = 0 -> ov = None /\ post t m l))). +======= + eexists (option measure), (with_bottom lt), (fun ov t m l mc => + exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ + ((word.unsigned br <> 0 -> exists v, ov = Some v /\ invariant v t m l (addMetricInstructions 2 + (addMetricLoads 2 + (addMetricJumps 1 mc')))) /\ + (word.unsigned br = 0 -> ov = None /\ post t m l (addMetricInstructions 1 + (addMetricLoads 1 + (addMetricJumps 1 mc')))))). +>>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) + split; auto using well_founded_with_bottom; []. split. + { destruct Henter as [br [mc' [He [Henter0 Henterm]]]]. + destruct (BinInt.Z.eq_dec (word.unsigned br) 0). + { exists None, br, mc'; split; trivial. + split; intros; try contradiction; split; eauto. } + { + assert (Hnonzero := n). + apply Henterm in Hnonzero. + destruct Hnonzero as [v Hinv]. + exists (Some v), br, mc'. + + split; trivial; []; split; try contradiction. +<<<<<<< HEAD + exists v0; split; trivial. } } + intros vi ti mi li (br&Ebr&Hcontinue&Hexit). + eexists; split; [eassumption|]; split. + { intros Hc; destruct (Hcontinue Hc) as (v&?&Hinv); subst. + eapply Proper_cmd; [ |eapply Hbody; eassumption]. +======= + + exists v; split; trivial. + + + } } + intros vi ti mi li mci (br&mcnew&Ebr&Hcontinue&Hexit). + eexists; eexists; split. + - eassumption. + - split. + + + (*eexists; eexists; split; [eassumption|]; split.*) + + { + intros Hc; destruct (Hcontinue Hc) as (v&?&Hinv); subst. + eapply Proper_cmd. + 1: { eapply Proper_call. } + 2: { (* eapply Hbody. eassumption. } + + [ | | eapply Hbody; eassumption]. [eapply Proper_call|]. +>>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) + intros t' m' l' (br'&Ebr'&Hinv'&Hpost'). + destruct (BinInt.Z.eq_dec (word.unsigned br') 0). + { exists None; split; try constructor. + exists br'; split; trivial; []. + split; intros; try contradiction. + split; eauto. } + { destruct (Hinv' ltac:(trivial)) as (v'&inv'<v'v). + exists (Some v'); split; trivial. (* NOTE: this [trivial] simpl-reduces [with_bottom] *) + exists br'; split; trivial. + split; intros; try contradiction. + eexists; split; eauto. } } + + eapply Hexit. + Qed. + *) Admitted. + + Lemma atleastonce + {e c t l mc} {m : mem} + (variables : list String.string) + {localstuple : tuple word (length variables)} + {Pl : enforce variables localstuple l} + {measure : Type} (invariant:_->_->_-> ufunc word (length variables) (MetricLog -> Prop)) + lt (Hwf : well_founded lt) + {post : _->_->_->_->Prop} + (Henter : exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ (word.unsigned br = 0%Z -> post t m l (addMetricInstructions 1 + (addMetricLoads 1 + (addMetricJumps 1 mc')))) /\ + (word.unsigned br <> 0 -> exists v', tuple.apply (invariant v' t m) localstuple (addMetricInstructions 2 + (addMetricLoads 2 + (addMetricJumps 1 mc')))) + + ) + (v0 : measure) (Hpre : tuple.apply (invariant v0 t m) localstuple mc) + (Hbody : forall v t m mc, tuple.foralls (fun localstuple => + tuple.apply (invariant v t m) localstuple mc -> + cmd call c t m (reconstruct variables localstuple) mc (fun t m l mc => + exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ + (word.unsigned br <> 0 -> Markers.unique (Markers.left (tuple.existss (fun localstuple => enforce variables localstuple l /\ Markers.right (Markers.unique (exists v', tuple.apply (invariant v' t m) localstuple (addMetricInstructions 2 + (addMetricLoads 2 + (addMetricJumps 1 mc'))) /\ lt v' v)))))) /\ + (word.unsigned br = 0 -> post t m l (addMetricInstructions 1 + (addMetricLoads 1 + (addMetricJumps 1 mc'))))))) + : cmd call (cmd.while e c) t m l mc post. + Proof. + eapply (atleastonce_localsmap (fun v t m l mc => exists localstuple, Logic.and (enforce variables localstuple l) (tuple.apply (invariant v t m) localstuple mc))); eauto. + { + destruct Henter as [br0 [mc0 [Henter [Henter0 Henternonzero]]]]. + exists br0. exists mc0. repeat (split; eauto). + intro Hnonzero. apply Henternonzero in Hnonzero. destruct Hnonzero as [v Hnonzero]. + exists v. exists localstuple. repeat (split; eauto). + } + { + intros vi ti mi li mci (?&X&Y). + specialize (Hbody vi ti mi). + eapply hlist.foralls_forall in Hbody. + specialize (Hbody Y). + rewrite <-(reconstruct_enforce _ _ _ X) in Hbody. + eapply Proper_cmd; [eapply Proper_call| |eapply Hbody]. + intros t' m' l' mc' (?&?&?&HH). + eexists; eexists; split; eauto. destruct HH as [HH1 HH2]. + split; intros; eauto. + specialize (HH1 ltac:(eauto)). + eapply hlist.existss_exists in HH1; destruct HH1 as (?&?&?&?&?). + eexists; split; eauto. + } + Qed. + (* + Lemma tailrec_earlyout_localsmap + {e c t} {m : mem} {l mc} {post : _-> _->_->_-> Prop} + {measure : Type} (spec:_->_->_->_->_->(Prop*(_->_->_->_-> Prop))) lt + (Hwf : well_founded lt) + (v0 : measure) (P0 := (spec v0 t m l mc).(1)) (Hpre : P0) + (Q0 := (spec v0 t m l mc).(2)) + (Hbody : forall v t m l mc, + let S := spec v t m l mc in let (P, Q) := S in + P -> + exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ + (word.unsigned br <> 0%Z -> cmd call c t m l mc' + (fun t' m' l' mc'' => + (exists br mc', expr m' l' e mc'' (eq (Datatypes.pair br mc')) /\ word.unsigned br = 0 /\ Q t' m' l' mc') \/ + exists v', let S' := spec v' t' m' l' in let '(P', Q') := S' in + P' /\ + lt v' v /\ + forall T M L MC, Q' T M L -> Q T M L)) /\ + (word.unsigned br = 0%Z -> Q t m l)) + (Hpost : forall t m l, Q0 t m l -> post t m l) + : cmd call (cmd.while e c) t m l post. + Proof. + eapply wp_while. + eexists (option measure), (with_bottom lt), (fun v t m l => + match v with + | None => exists br, expr m l e (eq br) /\ word.unsigned br = 0 /\ Q0 t m l + | Some v => + let S := spec v t m l in let '(P, Q) := S in + P /\ forall T M L, Q T M L -> Q0 T M L + end). + split; auto using well_founded_with_bottom; []; cbv [Markers.split] in *. + split. + { exists (Some v0); eauto. } + intros [vi|] ti mi li inv_i; [destruct inv_i as (?&Qi)|destruct inv_i as (br&Hebr&Hbr0&HQ)]. + { destruct (Hbody _ _ _ _ ltac:(eassumption)) as (br&?&X); exists br; split; [assumption|]. + destruct X as (Htrue&Hfalse). split; intros Hbr; + [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; eauto. + eapply Proper_cmd; [ |eapply Hpc]. + intros tj mj lj [(br'&Hbr'&Hz&HQ)|(vj&dP&?&dQ)]; + [exists None | exists (Some vj)]; cbn [with_bottom]; eauto 9. } + repeat esplit || eauto || intros; contradiction. + Qed. + + Lemma tailrec_earlyout + {e c t localsmap} {m : mem} + (ghosttypes : polymorphic_list.list Type) + (variables : list String.string) + {l0 : tuple word (length variables)} + {Pl : enforce variables l0 localsmap} + {post : _->_->_-> Prop} + {measure : Type} (spec:_->HList.arrows ghosttypes (_->_->ufunc word (length variables) (Prop*(_->_->ufunc word (length variables) Prop)))) lt + (Hwf : well_founded lt) + (v0 : measure) + : hlist.foralls (fun (g0 : hlist ghosttypes) => forall + (Hpre : (tuple.apply (hlist.apply (spec v0) g0 t m) l0).(1)) + (Hbody : forall v, hlist.foralls (fun g => forall t m, tuple.foralls (fun l => + @dlet _ (fun _ => Prop) (reconstruct variables l) (fun localsmap : locals => + match tuple.apply (hlist.apply (spec v) g t m) l with S_ => + S_.(1) -> + Markers.unique (Markers.left (exists br, expr m localsmap e (eq br) /\ Markers.right ( + (word.unsigned br <> 0%Z -> cmd call c t m localsmap + (fun t' m' localsmap' => + Markers.unique (Markers.left (hlist.existss (fun l' => enforce variables l' localsmap' /\ Markers.right ( + Markers.unique (Markers.left (exists br, expr m' localsmap' e (eq br) /\ Markers.right ( word.unsigned br = 0 /\ tuple.apply (S_.(2) t' m') l') ) ) \/ + Markers.unique (Markers.left (hlist.existss (fun g' => exists v', + match tuple.apply (hlist.apply (spec v') g' t' m') l' with S' => + S'.(1) /\ Markers.right ( + lt v' v /\ + forall T M, hlist.foralls (fun L => tuple.apply (S'.(2) T M) L -> tuple.apply (S_.(2) T M) L)) end))))))))) /\ + (word.unsigned br = 0%Z -> tuple.apply (S_.(2) t m) l))))end)))) + (Hpost : match (tuple.apply (hlist.apply (spec v0) g0 t m) l0).(2) with Q0 => forall t m, hlist.foralls (fun l => tuple.apply (Q0 t m) l -> post t m (reconstruct variables l))end) + , cmd call (cmd.while e c) t m localsmap post ). + Proof. + eapply hlist_forall_foralls; intros g0 **. + eapply wp_while. + eexists (option measure), (with_bottom lt), (fun vi ti mi localsmapi => + exists li, localsmapi = reconstruct variables li /\ + match vi with None => exists br, expr mi localsmapi e (eq br) /\ word.unsigned br = 0 /\ tuple.apply ((tuple.apply (hlist.apply (spec v0) g0 t m) l0).(2) ti mi) li | Some vi => + exists gi, match tuple.apply (hlist.apply (spec vi) gi ti mi) li with S_ => + S_.(1) /\ forall T M L, tuple.apply (S_.(2) T M) L -> + tuple.apply ((tuple.apply (hlist.apply (spec v0) g0 t m) l0).(2) T M) L end end). + cbv [Markers.unique Markers.split Markers.left Markers.right] in *. + split; eauto using well_founded_with_bottom. + split. { exists (Some v0), l0. split. 1: eapply reconstruct_enforce; eassumption. exists g0; split; eauto. } + intros [vi|] ti mi lmapi. + 2: { intros (ld&Hd&br&Hbr&Hz&Hdone). + eexists; split; eauto. + split; intros; try contradiction. + subst; eapply (hlist.foralls_forall (Hpost ti mi) _ Hdone). } + intros (?&?&gi&?&Qi); subst. + destruct (hlist.foralls_forall (hlist.foralls_forall (Hbody vi) gi ti mi) _ ltac:(eassumption)) as (br&?&X). + exists br; split; [assumption|]. destruct X as (Htrue&Hfalse). split; intros Hbr; + [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. + { eapply Proper_cmd; [ |eapply Hpc]. + intros tj mj lmapj Hlj; eapply hlist.existss_exists in Hlj. + destruct Hlj as (lj&Elj&HE); eapply reconstruct_enforce in Elj; subst lmapj. + destruct HE as [(br'&Hevalr'&Hz'&Hdone)|HE]. + { exists None; cbn. eauto 9. } + { eapply hlist.existss_exists in HE. destruct HE as (l&?&?&?&HR). + pose proof fun T M => hlist.foralls_forall (HR T M); clear HR. + eexists (Some _); eauto 9. } } + { pose proof fun t m => hlist.foralls_forall (Hpost t m); clear Hpost; eauto. } + Qed. + +<<<<<<< HEAD + + Lemma atleastonce + {e c t l} {m : mem} + (variables : list String.string) + {localstuple : tuple word (length variables)} + {Pl : enforce variables localstuple l} + {measure : Type} (invariant:_->_->_->ufunc word (length variables) Prop) + lt (Hwf : well_founded lt) + {post : _->_->_-> Prop} + (Henter : exists br, expr m l e (eq br) /\ (word.unsigned br = 0%Z -> post t m l)) + (v0 : measure) (Hpre : tuple.apply (invariant v0 t m) localstuple) + (Hbody : forall v t m, tuple.foralls (fun localstuple => + tuple.apply (invariant v t m) localstuple -> + cmd call c t m (reconstruct variables localstuple) (fun t m l => + exists br, expr m l e (eq br) /\ + (word.unsigned br <> 0 -> Markers.unique (Markers.left (tuple.existss (fun localstuple => enforce variables localstuple l /\ Markers.right (Markers.unique (exists v', tuple.apply (invariant v' t m) localstuple /\ lt v' v)))))) /\ + (word.unsigned br = 0 -> post t m l)))) + : cmd call (cmd.while e c) t m l post. + Proof. + eapply (atleastonce_localsmap (fun v t m l => exists localstuple, Logic.and (enforce variables localstuple l) (tuple.apply (invariant v t m) localstuple))); eauto. + intros vi ti mi li (?&X&Y). + specialize (Hbody vi ti mi). + eapply hlist.foralls_forall in Hbody. + specialize (Hbody Y). + rewrite <-(reconstruct_enforce _ _ _ X) in Hbody. + eapply Proper_cmd; [ |eapply Hbody]. + intros t' m' l' (?&?&HH&?). + eexists; split; eauto. + split; intros; eauto. + specialize (HH ltac:(eauto)). + eapply hlist.existss_exists in HH; destruct HH as (?&?&?&?&?). + eexists; split; eauto. + Qed. + +======= +>>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) + Lemma while_zero_iterations {e c t l} {m : mem} {post : _->_->_-> Prop} + (HCond: expr m l e (eq (word.of_Z 0))) + (HPost: post t m l) + : cmd call (cmd.while e c) t m l post. + Proof. + eapply (while_localsmap (fun n t' m' l' => t' = t /\ m' = m /\ l' = l) (PeanoNat.Nat.lt_wf 0) 0%nat). + 1: unfold split; auto. intros *. intros (? & ? & ?). subst. + eexists. split. 1: exact HCond. + rewrite Properties.word.unsigned_of_Z_0. + split; intros; congruence. + Qed. + + + (* Bedrock-style loop rule *) + Local Open Scope sep_scope. + Local Infix "*" := Separation.sep : type_scope. + Local Infix "==>" := Lift1Prop.impl1. + + Lemma tailrec_sep + e c t (m : mem) l (post : _->_->_-> Prop) + {measure : Type} P Q lt (Hwf : well_founded lt) (v0 : measure) R0 + (Hpre : (P v0 t l * R0) m) + (Hbody : forall v t m l R, (P v t l * R) m -> + exists br, expr m l e (eq br) /\ + (word.unsigned br <> 0%Z -> cmd call c t m l + (fun t' m' l' => exists v' dR, (P v' t' l' * (R * dR)) m' /\ + lt v' v /\ + forall T L, Q v' T L * dR ==> Q v T L)) /\ + (word.unsigned br = 0%Z -> (Q v t l * R) m)) + (Hpost : forall t m l, (Q v0 t l * R0) m -> post t m l) + : cmd call (cmd.while e c) t m l post. + Proof. + eapply wp_while. + eexists measure, lt, (fun v t m l => exists R, (P v t l * R) m /\ + forall T L, Q v T L * R ==> Q v0 T L * R0). + split; [assumption|]. + split. { exists v0, R0. split; [assumption|]. intros. reflexivity. } + intros vi ti mi li (Ri&?&Qi). + destruct (Hbody _ _ _ _ _ ltac:(eassumption)) as (br&?&X); exists br; split; [assumption|]. + destruct X as (Htrue&Hfalse). split; intros Hbr; + [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. + { eapply Proper_cmd; [ |eapply Hpc]. + intros tj mj lj (vj&dR&dP&?&dQ). + exists vj; split; [|assumption]. + exists (Ri * dR); split; [assumption|]. + intros. rewrite (sep_comm _ dR), <-(sep_assoc _ dR), dQ; trivial. } + { eapply Hpost, Qi, Hpc. } + Qed. +*) +End Loops. + +Ltac loop_simpl := + cbn [reconstruct map.putmany_of_list HList.tuple.to_list + HList.hlist.foralls HList.tuple.foralls + HList.hlist.existss HList.tuple.existss + HList.hlist.apply HList.tuple.apply HList.hlist + List.repeat Datatypes.length + HList.polymorphic_list.repeat HList.polymorphic_list.length + PrimitivePair.pair._1 PrimitivePair.pair._2] in *. + diff --git a/bedrock2/src/bedrock2/MetricWeakestPrecondition.v b/bedrock2/src/bedrock2/MetricWeakestPrecondition.v new file mode 100644 index 000000000..63570bd97 --- /dev/null +++ b/bedrock2/src/bedrock2/MetricWeakestPrecondition.v @@ -0,0 +1,339 @@ +Require Import coqutil.Macros.subst coqutil.Macros.unique coqutil.Map.Interface coqutil.Map.OfListWord. +Require Import Coq.ZArith.BinIntDef coqutil.Word.Interface coqutil.Word.Bitwidth. +Require Import coqutil.dlet bedrock2.Syntax bedrock2.Semantics. +Require Import bedrock2.MetricLogging. +Require Import bedrock2.MetricSemantics. + +Section WeakestPrecondition. + Context {width: Z} {BW: Bitwidth width} {word: word.word width} {mem: map.map word Byte.byte}. + Context {locals: map.map String.string word}. + Context {ext_spec: ExtSpec}. + Implicit Types (t : trace) (m : mem) (l : locals). + + Local Notation metrics := MetricLog. + + Definition literal v mc (post : (word * metrics) -> Prop) : Prop := + dlet! v := word.of_Z v in post (v, addMetricInstructions 8 (addMetricLoads 8 mc)). + Definition get (l : locals) (x : String.string) mc (post : (word * metrics) -> Prop) : Prop := + exists v, map.get l x = Some v /\ post (v, addMetricInstructions 1 (addMetricLoads 2 mc)). + Definition load s m a mc (post: (word * metrics) -> Prop) : Prop := + exists v, load s m a = Some v /\ post(v, addMetricInstructions 1 (addMetricLoads 2 mc)). + Definition store sz m a v post := + exists m', store sz m a v = Some m' /\ post m'. + + Section WithMemAndLocals. + Context (m : mem) (l : locals). + Definition expr_body (rec : _->_->(word*metrics->Prop)->Prop) (e : Syntax.expr) (mc : metrics) (post : word * metrics -> Prop) : Prop := + match e with + | expr.literal v => + literal v mc post + | expr.var x => + get l x mc post + | expr.op op e1 e2 => + rec e1 mc (fun '(v1, mc') => + rec e2 mc' (fun '(v2, mc'') => + post (interp_binop op v1 v2, addMetricInstructions 2 (addMetricLoads 2 mc'')))) + | expr.load s e => + rec e mc (fun '(a, mc') => + load s m a mc' post) + | expr.inlinetable s t e => + rec e mc (fun '(a, mc') => + load s (map.of_list_word t) a (addMetricInstructions 2 + (addMetricLoads 2 + (addMetricJumps 1 mc'))) post) + | expr.ite c e1 e2 => + rec c mc (fun '(b, mc') => rec (if word.eqb b (word.of_Z 0) then e2 else e1) (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc'))) post) + end. + Fixpoint expr e := expr_body expr e. + End WithMemAndLocals. + + Section WithF. + Context {A B} (f: A -> metrics -> (B * metrics -> Prop) -> Prop). + Definition list_map_body rec (xs : list A) (mc : metrics) (post : list B * metrics -> Prop) : Prop := + match xs with + | nil => post (nil, mc) + | cons x xs' => + f x mc (fun '(y, mc') => + rec xs' mc' (fun '(ys', mc'') => + post (cons y ys', mc''))) + end. + Fixpoint list_map xs := list_map_body list_map xs. + End WithF. + + Section WithFunctions. + + Context (e: env). + Context (call : String.string -> trace -> mem -> list word -> metrics -> (trace -> mem -> list word -> metrics -> Prop) -> Prop). + Definition dexpr m l e mc v := expr m l e mc (eq v). + Definition dexprs m l es mc vs := list_map (expr m l) es mc (eq vs). +(* All cases except cmd.while and cmd.call can be denoted by structural recursion + over the syntax. + For cmd.while and cmd.call, we fall back to the operational semantics *) + Definition cmd_body (rec:_->_->_->_->_->_->Prop) (c : cmd) (t : trace) (m : mem) (l : locals) (mc : metrics) + (post : trace -> mem -> locals -> metrics -> Prop) : Prop := + (* give value of each pure expression when stating its subproof *) + match c with + | cmd.skip => post t m l mc + | cmd.set x ev => + exists v mc', dexpr m l ev mc (v, mc') /\ + dlet! l := map.put l x v in + post t m l (addMetricInstructions 1 (addMetricLoads 1 mc')) + | cmd.unset x => + dlet! l := map.remove l x in + post t m l mc + | cmd.store sz ea ev => + exists a mc', dexpr m l ea mc (a, mc') /\ + exists v mc'', dexpr m l ev mc (v, mc'') /\ + store sz m a v (fun m => + post t m l (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc'')))) + | cmd.stackalloc x n c => + Z.modulo n (bytes_per_word width) = 0 /\ + forall a mStack mCombined, + anybytes a n mStack -> map.split mCombined m mStack -> + dlet! l := map.put l x a in + rec c t mCombined l (addMetricInstructions 1 (addMetricLoads 1 mc)) + (fun t' mCombined' l' mc' => + exists m' mStack', + anybytes a n mStack' /\ map.split mCombined' m' mStack' /\ + post t' m' l' mc') + | cmd.cond br ct cf => + exists v mc', dexpr m l br mc (v, mc') /\ + dlet! mc'' := addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc')) in + (word.unsigned v <> 0%Z -> rec ct t m l mc'' post) /\ + (word.unsigned v = 0%Z -> rec cf t m l mc'' post) + | cmd.seq c1 c2 => + rec c1 t m l mc (fun t m l mc => rec c2 t m l mc post) + | cmd.while _ _ => MetricSemantics.exec e c t m l mc post + | cmd.call binds fname arges => + exists args mc', dexprs m l arges mc (args, mc') /\ + MetricSemantics.call e fname t m args mc (fun t m rets mc'' => + exists l', map.putmany_of_list_zip binds rets l = Some l' /\ post t m l' (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc''))))) + | cmd.interact binds action arges => + exists args mc', dexprs m l arges mc (args, mc') /\ + exists mKeep mGive, map.split m mKeep mGive /\ + ext_spec t mGive action args (fun mReceive rets => + exists l', map.putmany_of_list_zip binds rets l = Some l' /\ + forall m', map.split m' mKeep mReceive -> + post (cons ((mGive, action, args), (mReceive, rets)) t) m' l' (addMetricInstructions 1 + (addMetricStores 1 + (addMetricLoads 2 mc')))) + end. + + Fixpoint cmd c := cmd_body cmd c. + End WithFunctions. + + Definition func call '(innames, outnames, c) (t : trace) (m : mem) (args : list word) (mc : metrics) (post : trace -> mem -> list word -> metrics -> Prop) := +exists l, map.of_list_zip innames args = Some l /\ + cmd call c t m l mc (fun t m l mc => + list_map (get l) outnames mc (fun '(rets, _) => + post t m rets mc)). + + Definition program := cmd. + + (* + Definition call_body rec (functions : list (String.string * (list String.string * list String.string * cmd.cmd))) + (fname : String.string) (t : trace) (m : mem) (args : list word) (mc: metrics) + (post : trace -> mem -> list word -> metrics -> Prop) : Prop := + match functions with + | nil => False + | cons (f, decl) functions => + if String.eqb f fname + then func (rec functions) decl t m args mc post + else rec functions fname t m args mc post + end. + Fixpoint call functions := call_body call functions. + + Definition program funcs main t m l mc post : Prop := cmd (call funcs) main t m l mc post. *) + +End WeakestPrecondition. +Notation call := MetricSemantics.call (only parsing). + +Ltac unfold1_cmd e := + lazymatch e with + @cmd ?width ?BW ?word ?mem ?locals ?ext_spec ?CA ?c ?t ?m ?l ?mc ?post => + let c := eval hnf in c in + constr:(@cmd_body width BW word mem locals ext_spec CA + (@cmd width BW word mem locals ext_spec CA) c t m l mc post) + end. +Ltac unfold1_cmd_goal := + let G := lazymatch goal with |- ?G => G end in + let G := unfold1_cmd G in + change G. + +Ltac unfold1_expr e := + lazymatch e with + @expr ?width ?word ?mem ?locals ?m ?l ?arg ?mc ?post => + let arg := eval hnf in arg in + constr:(@expr_body width word mem locals m l (@expr width word mem locals m l) arg mc post) + end. +Ltac unfold1_expr_goal := + let G := lazymatch goal with |- ?G => G end in + let G := unfold1_expr G in + change G. + +Ltac unfold1_list_map e := + lazymatch e with + @list_map ?A ?B ?P ?arg ?mc ?post => + let arg := eval hnf in arg in + constr:(@list_map_body A B P (@list_map A B P) arg mc post) + end. +Ltac unfold1_list_map_goal := + let G := lazymatch goal with |- ?G => G end in + let G := unfold1_list_map G in + change G. + +(* +Ltac unfold1_call e := + lazymatch e with + @call ?width ?BW ?word ?mem ?locals ?ext_spec ?fs ?fname ?t ?m ?l ?mc ?post => + let fs := eval hnf in fs in + constr:(@call_body width BW word mem locals ext_spec + (@call width BW word mem locals ext_spec) fs fname t m l mc post) + end. +Ltac unfold1_call_goal := + let G := lazymatch goal with |- ?G => G end in + let G := unfold1_call G in + change G. + + *) + +Import Coq.ZArith.ZArith. + +Notation "'fnspec!' name a0 .. an '/' g0 .. gn '~>' r0 .. rn ',' '{' 'requires' tr mem mc := pre ';' 'ensures' tr' mem' mc' ':=' post '}'" := + (fun functions => + (forall a0, + .. (forall an, + (forall g0, + .. (forall gn, + (forall tr mem mc, + pre -> + MetricWeakestPrecondition.call + functions name tr mem (cons a0 .. (cons an nil) ..) mc + (fun tr' mem' rets mc' => + (exists r0, + .. (exists rn, + rets = (cons r0 .. (cons rn nil) ..) /\ + post) ..)))) ..)) ..)) + (at level 200, + name at level 0, + a0 binder, an binder, + g0 binder, gn binder, + r0 closed binder, rn closed binder, + tr name, tr' name, mem name, mem' name, mc name, mc' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name a0 .. an '/' g0 .. gn ',' '{' 'requires' tr mem mc := pre ';' 'ensures' tr' mem' mc' ':=' post '}'" := + (fun functions => + (forall a0, + .. (forall an, + (forall g0, + .. (forall gn, + (forall tr mem mc, + pre -> + MetricWeakestPrecondition.call + functions name tr mem (cons a0 .. (cons an nil) ..) mc + (fun tr' mem' rets mc' => + rets = nil /\ post))) ..)) ..)) + (at level 200, + name at level 0, + a0 binder, an binder, + g0 binder, gn binder, + tr name, tr' name, mem name, mem' name, mc name, mc' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name a0 .. an '~>' r0 .. rn ',' '{' 'requires' tr mem mc := pre ';' 'ensures' tr' mem' mc' ':=' post '}'" := + (fun functions => + (forall a0, + .. (forall an, + (forall tr mem mc, + pre -> + MetricWeakestPrecondition.call + functions name tr mem (cons a0 .. (cons an nil) ..) mc + (fun tr' mem' rets mc' => + (exists r0, + .. (exists rn, + rets = (cons r0 .. (cons rn nil) ..) /\ + post) ..)))) ..)) + (at level 200, + name at level 0, + a0 binder, an binder, + r0 closed binder, rn closed binder, + tr name, tr' name, mem name, mem' name, mc name, mc' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name '/' g0 .. gn '~>' r0 .. rn ',' '{' 'requires' tr mem mc := pre ';' 'ensures' tr' mem' mc' ':=' post '}'" := + (fun functions => + (forall an, + (forall g0, + .. (forall gn, + (forall tr mem mc, + pre -> + MetricWeakestPrecondition.call + functions name tr mem nil mc + (fun tr' mem' rets mc' => + (exists r0, + .. (exists rn, + rets = (cons r0 .. (cons rn nil) ..) /\ + post) ..)))) ..))) + (at level 200, + name at level 0, + g0 binder, gn binder, + r0 closed binder, rn closed binder, + tr name, tr' name, mem name, mem' name, mc name, mc' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name a0 .. an ',' '{' 'requires' tr mem mc := pre ';' 'ensures' tr' mem' mc' ':=' post '}'" := + (fun functions => + (forall a0, + .. (forall an, + (forall tr mem mc, + pre -> + MetricWeakestPrecondition.call + functions name tr mem (cons a0 .. (cons an nil) ..) mc + (fun tr' mem' rets mc' => + rets = nil /\ post))) ..)) + (at level 200, + name at level 0, + a0 binder, an binder, + tr name, tr' name, mem name, mem' name, mc name, mc' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name '/' g0 .. gn ',' '{' 'requires' tr mem mc := pre ';' 'ensures' tr' mem' mc' ':=' post '}'" := + (fun functions => + (forall g0, + .. (forall gn, + (forall tr mem mc, + pre -> + MetricWeakestPrecondition.call + functions name tr mem nil mc + (fun tr' mem' rets mc' => + rets = nil /\ post))) ..)) + (at level 200, + name at level 0, + g0 binder, gn binder, + tr name, tr' name, mem name, mem' name, mc name, mc' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name '~>' r0 .. rn ',' '{' 'requires' tr mem mc := pre ';' 'ensures' tr' mem' mc' ':=' post '}'" := + (fun functions => + (forall tr mem mc, + pre -> + MetricWeakestPrecondition.call + functions name tr mem nil mc + (fun tr' mem' rets mc' => + (exists r0, + .. (exists rn, + rets = (cons r0 .. (cons rn nil) ..) /\ + post) ..)))) + (at level 200, + name at level 0, + r0 closed binder, rn closed binder, + tr name, tr' name, mem name, mem' name, mc name, mc' name, + pre at level 200, + post at level 200). diff --git a/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v b/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v new file mode 100644 index 000000000..4a7fdcefb --- /dev/null +++ b/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v @@ -0,0 +1,324 @@ +Require Import coqutil.Macros.subst coqutil.Macros.unique coqutil.Map.Interface coqutil.Word.Properties. +Require Import coqutil.Word.Bitwidth. +Require Import bedrock2.MetricLogging. +Require bedrock2.MetricWeakestPrecondition. + +Require Import Coq.Classes.Morphisms. + +Section MetricWeakestPrecondition. + Context {width} {BW: Bitwidth width} {word: word.word width} {mem: map.map word Byte.byte}. + Context {locals: map.map String.string word}. + Context {ext_spec: Semantics.ExtSpec}. + + Ltac ind_on X := + intros; + (* Note: Comment below dates from when we were using a parameter record p *) + (* Note: "before p" means actually "after p" when reading from top to bottom, because, + as the manual points out, "before" and "after" are with respect to the direction of + the move, and we're moving hypotheses upwards here. + We need to make sure not to revert/clear p, because the other lemmas depend on it. + If we still reverted/cleared p, we'd get errors like + "Error: Proper_load depends on the variable p which is not declared in the context." + when trying to use Proper_load, or, due to COQBUG https://github.com/coq/coq/issues/11487, + we'd get a typechecking failure at Qed time. *) + repeat match goal with x : ?T |- _ => first + [ constr_eq T X; move x before ext_spec + | constr_eq T X; move x before locals + | constr_eq T X; move x at top + | revert x ] end; + match goal with x : X |- _ => induction x end; + intros. + + Local Hint Mode word.word - : typeclass_instances. + + (* we prove weakening lemmas for all WP definitions in a syntax-directed fashion, + * moving from postcondition towards precondition one logical connective at a time. *) + Global Instance Proper_literal : Proper (pointwise_relation _ (pointwise_relation _ ((pointwise_relation _ Basics.impl) ==> Basics.impl))) MetricWeakestPrecondition.literal. + Proof using. clear. cbv [MetricWeakestPrecondition.literal]; cbv [Proper respectful pointwise_relation Basics.impl dlet.dlet]. eauto. Qed. + + Global Instance Proper_get : Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ ((pointwise_relation _ Basics.impl) ==> Basics.impl)))) MetricWeakestPrecondition.get. + Proof using. clear. cbv [MetricWeakestPrecondition.get]; cbv [Proper respectful pointwise_relation Basics.impl]; intros * ? (?&?&?); eauto. Qed. + + Global Instance Proper_load : Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ ((pointwise_relation _ Basics.impl) ==> Basics.impl))))) MetricWeakestPrecondition.load. + Proof using. clear. cbv [MetricWeakestPrecondition.load]; cbv [Proper respectful pointwise_relation Basics.impl]; intros * ? (?&?&?); eauto. Qed. + + Global Instance Proper_store : Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ ((pointwise_relation _ Basics.impl) ==> Basics.impl))))) MetricWeakestPrecondition.store. + Proof using. clear. cbv [MetricWeakestPrecondition.store]; cbv [Proper respectful pointwise_relation Basics.impl]; intros * ? (?&?&?); eauto. Qed. + + Global Instance Proper_expr : Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ ((pointwise_relation _ Basics.impl) ==> Basics.impl))))) MetricWeakestPrecondition.expr. + Proof using. + clear. + cbv [Proper respectful pointwise_relation Basics.impl]; ind_on Syntax.expr.expr; + cbn in *; intuition (try typeclasses eauto with core). + { eapply Proper_literal; eauto. } + { eapply Proper_get; eauto. } + { eapply IHa1; eauto; intuition idtac. destruct a4. eapply Proper_load; eauto using Proper_load. } + { eapply IHa1; eauto; intuition idtac. destruct a4. eapply Proper_load; eauto using Proper_load. } + { eapply IHa1_1; eauto. destruct a1. eapply IHa1_2; eauto. destruct a1. eauto. } + {eapply IHa1_1; eauto; intuition idtac. destruct a1. Tactics.destruct_one_match; eauto using Proper_load . } + Qed. + + Global Instance Proper_list_map {A B} : + Proper ((pointwise_relation _ (pointwise_relation _ (pointwise_relation _ Basics.impl ==> Basics.impl))) ==> pointwise_relation _ (pointwise_relation _ (pointwise_relation _ Basics.impl ==> Basics.impl))) (MetricWeakestPrecondition.list_map (A:=A) (B:=B)). + Proof using. + clear. + cbv [Proper respectful pointwise_relation Basics.impl]; ind_on (list A); + cbn in *; intuition (try typeclasses eauto with core). + eapply H; eauto. destruct a2. eapply IHa; eauto. destruct a2; eauto. + Qed. + + Context {word_ok : word.ok word} {mem_ok : map.ok mem}. + Context {locals_ok : map.ok locals}. + Context {ext_spec_ok : Semantics.ext_spec.ok ext_spec}. + + Global Instance Proper_cmd : + Proper + (pointwise_relation _ ( + pointwise_relation _ ( + pointwise_relation _ ( + pointwise_relation _ ( + pointwise_relation _ ( + pointwise_relation _ ( + (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ Basics.impl)))) ==> + Basics.impl))))))) MetricWeakestPrecondition.cmd. + Proof using ext_spec_ok locals_ok mem_ok word_ok. + pose proof I. (* to keep naming *) + cbv [Proper respectful pointwise_relation Basics.flip Basics.impl]; ind_on Syntax.cmd.cmd; + cbn in *; cbv [dlet.dlet] in *; intuition (try typeclasses eauto with core). + { destruct H1 as (?&?&?&?). eexists. eexists. split. + 1: eapply Proper_expr. + 1: cbv [pointwise_relation Basics.impl]; intuition eauto 2. + all: eauto. } + { destruct H1 as (?&?&?&?). eexists. eexists. split. + { eapply Proper_expr. + { cbv [pointwise_relation Basics.impl]; intuition eauto 2. } + { eauto. } } + { destruct H2 as (?&?&?&?). eexists. eexists. split. + { eapply Proper_expr. + { cbv [pointwise_relation Basics.impl]; intuition eauto 2. } + { eauto. } } + { eapply Proper_store; eauto; cbv [pointwise_relation Basics.impl]; eauto. } } } + + { eapply H1. 2: eapply H3; eassumption. intros ? ? ? ? (?&?&?&?&?). eauto 7. } + { destruct H1 as (?&?&?&?). eexists. eexists. split. + { eapply Proper_expr. + { cbv [pointwise_relation Basics.impl]; intuition eauto 2. } + { eauto. } } + { intuition eauto 6. } } + { eapply MetricSemantics.exec.weaken; eassumption. } + { destruct H1 as (?&?&?&?). eexists. eexists. split. + { eapply Proper_list_map; eauto; try exact H4; cbv [respectful pointwise_relation Basics.impl]; intuition eauto 2. + eapply Proper_expr; eauto. } + { eapply MetricSemantics.weaken_call. 1: eassumption. cbv beta. + (* COQBUG (performance), measured in Coq 8.9: + "firstorder eauto" works, but takes ~100s and increases memory usage by 1.8GB. + On the other hand, the line below takes just 5ms *) + cbv beta; intros ? ? ? ? (?&?&?); eauto. } } + { destruct H1 as (?&?&?&?). eexists. eexists. split. + { eapply Proper_list_map; eauto; try exact H4; cbv [respectful pointwise_relation Basics.impl]. + { eapply Proper_expr; eauto. } + { eauto. } } + { destruct H2 as (mKeep & mGive & ? & ?). + exists mKeep. exists mGive. + split; [assumption|]. + eapply Semantics.ext_spec.weaken; [|solve[eassumption]]. + intros ? ? (?&?&?); eauto 10. } } + Qed. + + Global Instance Proper_call : + Proper ( + (pointwise_relation _ ( + (pointwise_relation _ ( + (pointwise_relation _ ( + pointwise_relation _ ( + pointwise_relation _ ( + pointwise_relation _ ( + (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ Basics.impl)))) ==> + Basics.impl)))))))))) MetricWeakestPrecondition.call. + Proof using word_ok mem_ok locals_ok ext_spec_ok. + cbv [Proper respectful pointwise_relation Basics.impl]. + intros. eapply MetricSemantics.weaken_call; eassumption. + Qed. + +Global Instance Proper_program : + Proper ( + pointwise_relation _ ( + pointwise_relation _ ( + pointwise_relation _ ( + pointwise_relation _ ( + pointwise_relation _ ( + pointwise_relation _ ( + (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ Basics.impl)))) ==> + Basics.impl))))))) MetricWeakestPrecondition.program. + Proof using word_ok mem_ok locals_ok ext_spec_ok. + cbv [Proper respectful pointwise_relation Basics.impl MetricWeakestPrecondition.program]; intros. + eapply Proper_cmd; + cbv [Proper respectful pointwise_relation Basics.flip Basics.impl MetricWeakestPrecondition.func]; + try solve [typeclasses eauto with core]. + Qed. + + Ltac t := + repeat match goal with + | |- forall _, _ => progress intros + | H: exists _, _ |- _ => destruct H + | H: and _ _ |- _ => destruct H + | H: eq _ ?y |- _ => subst y + | H: False |- _ => destruct H + | _ => progress cbn in * + | _ => progress cbv [dlet.dlet MetricWeakestPrecondition.dexpr MetricWeakestPrecondition.dexprs MetricWeakestPrecondition.store] in * + end; eauto. + +Lemma expr_sound m l e mc post (H : MetricWeakestPrecondition.expr m l e mc post) + : exists v mc', MetricSemantics.eval_expr m l e mc = Some (v, mc') /\ post (v, mc'). +Proof using BW ext_spec ext_spec_ok locals +locals_ok mem mem_ok width word word_ok. + ind_on Syntax.expr; t. { destruct H. destruct H. eexists. eexists. rewrite H. eauto. } + { eapply IHe in H; t. cbv [MetricWeakestPrecondition.load] in H0; t. rewrite H. rewrite H0. eauto. } + { eapply IHe in H; t. cbv [MetricWeakestPrecondition.load] in H0; t. rewrite H. rewrite H0. + (* now we prove that 1+2=3 and 2+2=4 *) + eexists. eexists. split; eauto. + MetricLogging.unfold_MetricLog. MetricLogging.simpl_MetricLog. + repeat rewrite <- BinInt.Z.add_assoc in H1. + simpl in H1. + eauto. } + { eapply IHe1 in H; t. eapply IHe2 in H0; t. rewrite H, H0; eauto. } + { eapply IHe1 in H; t. rewrite H. Tactics.destruct_one_match. + { apply IHe3; t. } + { eapply IHe2 in H0; t. } } + Qed. + + Import ZArith coqutil.Tactics.Tactics. + + Lemma expr_complete: forall m l e mc v mc', + MetricSemantics.eval_expr m l e mc = Some (v, mc') -> + MetricWeakestPrecondition.dexpr m l e mc (v, mc'). + Proof using word_ok. + induction e; cbn; intros. + - inversion_clear H. reflexivity. + - eexists; eexists; destruct (map.get l x); try inversion H; try reflexivity. + - repeat (destruct_one_match_hyp; try discriminate; []). + eapply Proper_expr. + 2: { eapply IHe. rewrite E. reflexivity. } + intros (addr, oldmc) ?. apply pair_equal_spec in H0; destruct H0. + subst r m0. unfold MetricWeakestPrecondition.load. eexists; split; eauto. + apply Option.eq_of_eq_Some in H. auto. + - repeat (destruct_one_match_hyp; try discriminate; []). + eapply Proper_expr. + 2: { eapply IHe. rewrite E. reflexivity. } + intros (addr, oldmc) ?. apply pair_equal_spec in H0; destruct H0. + subst r m0. unfold MetricWeakestPrecondition.load. eexists; split; eauto. + apply Option.eq_of_eq_Some in H; auto. rewrite <- H. + apply pair_equal_spec; split; auto. + destruct oldmc. applyAddMetricsGoal. reflexivity. + - repeat (destruct_one_match_hyp; try discriminate; []). + eapply Proper_expr. + 2: { eapply IHe1. rewrite E. reflexivity. } + intros (v1, oldmc1) ?. apply pair_equal_spec in H0; destruct H0. + subst r m0. + eapply Proper_expr. + 2: { eapply IHe2. rewrite E0. reflexivity. } + intros (v2, oldmc2) ?. apply pair_equal_spec in H0; destruct H0. + subst r0 m1. congruence. + - repeat (destruct_one_match_hyp; try discriminate; []). + eapply Proper_expr. + 2: { eapply IHe1. rewrite E. reflexivity. } + intros (vc, oldmc) ?. apply pair_equal_spec in H0; destruct H0. + subst r m0. + destr (word.eqb vc (word.of_Z 0)). + + eapply IHe3. eassumption. + + eapply IHe2. eassumption. + Qed. + + +Lemma sound_args : forall m l args mc P, + MetricWeakestPrecondition.list_map (MetricWeakestPrecondition.expr m l) args mc P -> + exists x mc', MetricSemantics.eval_call_args m l args mc = Some (x, mc') /\ P (x, mc'). +Proof using BW ext_spec ext_spec_ok locals locals_ok mem mem_ok +width word word_ok. + induction args; cbn; repeat (subst; t). + eapply expr_sound in H; t; rewrite H. + eapply IHargs in H0; t; rewrite H0. + eauto. + Qed. + + Lemma sound_getmany l a mc P : + MetricWeakestPrecondition.list_map (MetricWeakestPrecondition.get l) a mc P + -> exists vs mc', map.getmany_of_list l a = Some vs /\ P (vs, mc'). + Proof. + cbv [map.getmany_of_list] in *. + revert P l mc; induction a; cbn; repeat (subst; t). + cbv [MetricWeakestPrecondition.get] in H; t. + epose proof (IHa _ l _ _); clear IHa; t. + rewrite H. erewrite H1. eexists; eexists; split; eauto. + Unshelve. + 3: exact H0. + all: cbv [respectful pointwise_relation Basics.impl MetricWeakestPrecondition.get]; intros; cbv beta; t. + Qed. + + Local Hint Constructors MetricSemantics.exec : core. + Lemma sound_cmd e c t m l mc post (H: MetricWeakestPrecondition.cmd e c t m l mc post) : MetricSemantics.exec e c t m l mc post. + Proof. + ind_on Syntax.cmd; repeat (t; try match reverse goal with H: MetricWeakestPrecondition.expr _ _ _ _ _ |- _ => eapply expr_sound in H end). + { destruct (BinInt.Z.eq_dec (Interface.word.unsigned x) (BinNums.Z0)) as [Hb|Hb]; cycle 1; apply pair_equal_spec in H1; destruct H1; econstructor; t. } + { inversion H0. inversion H. + apply pair_equal_spec in H3; destruct H3. + apply pair_equal_spec in H4; destruct H4. subst x0. + econstructor; t. + Admitted. + + Lemma weaken_cmd: forall e c t m l mc (post1 post2: _->_->_->_->Prop), + MetricWeakestPrecondition.cmd e c t m l mc post1 -> + (forall t m l mc, post1 t m l mc -> post2 t m l mc) -> + MetricWeakestPrecondition.cmd e c t m l mc post2. + Proof. Admitted. + + Lemma complete_args : forall m l args mc vs, + MetricSemantics.eval_call_args m l args mc = Some vs -> + MetricWeakestPrecondition.dexprs m l args mc vs. + Proof using word_ok. Admitted. + + Lemma complete_cmd: forall e c t m l mc post, + MetricSemantics.exec e c t m l mc post -> + MetricWeakestPrecondition.cmd e c t m l mc post. + Proof. Admitted. + + Lemma start_func: forall e fname fimpl t m args mc post, + map.get e fname = Some fimpl -> + MetricWeakestPrecondition.func e fimpl t m args mc post -> + MetricWeakestPrecondition.call e fname t m args mc post. + Proof. Admitted. + + (** Ad-hoc lemmas here? *) + + Import bedrock2.Syntax bedrock2.MetricSemantics bedrock2.MetricWeakestPrecondition. + + (* + Lemma interact_nomem call action binds arges t m l post + args (Hargs : dexprs m l arges args) + (Hext : ext_spec t map.empty binds args (fun mReceive (rets : list word) => + mReceive = map.empty /\ + exists l0 : locals, map.putmany_of_list_zip action rets l = Some l0 /\ + post (cons (map.empty, binds, args, (map.empty, rets)) t) m l0)) + : WeakestPrecondition.cmd call (cmd.interact action binds arges) t m l post. + Proof using word_ok mem_ok ext_spec_ok. + exists args; split; [exact Hargs|]. + exists m. + exists map.empty. + split; [eapply Properties.map.split_empty_r; exact eq_refl|]. + eapply ext_spec.weaken; [|eapply Hext]; intros ? ? [? [? []]]. subst a; subst. + eexists; split; [eassumption|]. + intros. eapply Properties.map.split_empty_r in H. subst. assumption. + Qed.*) + + Lemma intersect_expr: forall m l e mc (post1 post2: word * MetricLog -> Prop), + MetricWeakestPrecondition.expr m l e mc post1 -> + MetricWeakestPrecondition.expr m l e mc post2 -> + MetricWeakestPrecondition.expr m l e mc (fun v => post1 v /\ post2 v). + Proof using word_ok. Admitted. + + Lemma dexpr_expr (m : mem) l e mc P + (H : MetricWeakestPrecondition.expr m l e mc P) + : exists v, MetricWeakestPrecondition.dexpr m l e mc v /\ P v. + Proof using word_ok. Admitted. + +End MetricWeakestPrecondition. diff --git a/bedrock2/src/bedrock2/NathanDidntNoticeMetricSemantics.v b/bedrock2/src/bedrock2/NathanDidntNoticeMetricSemantics.v new file mode 100644 index 000000000..597817120 --- /dev/null +++ b/bedrock2/src/bedrock2/NathanDidntNoticeMetricSemantics.v @@ -0,0 +1,385 @@ +Require Import coqutil.sanity coqutil.Macros.subst coqutil.Macros.unique coqutil.Byte. +Require Import coqutil.Datatypes.PrimitivePair coqutil.Datatypes.HList. +Require Import coqutil.Decidable. +Require Import coqutil.Tactics.fwd. +Require Import coqutil.Map.Properties. +Require coqutil.Map.SortedListString. +Require Import bedrock2.Notations bedrock2.Syntax coqutil.Map.Interface coqutil.Map.OfListWord. +Require Import BinIntDef coqutil.Word.Interface coqutil.Word.Bitwidth. +Require Import bedrock2.MetricLogging. +Require Export bedrock2.Memory. +Require Import Coq.Lists.List. + +(* BW is not needed on the rhs, but helps infer width *) +Definition trace{width: Z}{BW: Bitwidth width}{word: word.word width}{mem: map.map word byte} := + list ((mem * String.string * list word) * (mem * list word)). + +Definition ExtSpec{width: Z}{BW: Bitwidth width}{word: word.word width}{mem: map.map word byte} := + (* Given a trace of what happened so far, + the given-away memory, an action label and a list of function call arguments, *) + trace -> mem -> String.string -> list word -> + (* and a postcondition on the received memory and function call results, *) + (mem -> list word -> Prop) -> + (* tells if this postcondition will hold *) + Prop. + +Existing Class ExtSpec. + +Module ext_spec. + Class ok{width: Z}{BW: Bitwidth width}{word: word.word width}{mem: map.map word byte} + {ext_spec: ExtSpec}: Prop := + { + (* The action name and arguments uniquely determine the footprint of the given-away memory. *) + unique_mGive_footprint: forall t1 t2 mGive1 mGive2 a args + (post1 post2: mem -> list word -> Prop), + ext_spec t1 mGive1 a args post1 -> + ext_spec t2 mGive2 a args post2 -> + map.same_domain mGive1 mGive2; + + weaken :> forall t mGive act args, + Morphisms.Proper + (Morphisms.respectful + (Morphisms.pointwise_relation Interface.map.rep + (Morphisms.pointwise_relation (list word) Basics.impl)) Basics.impl) + (ext_spec t mGive act args); + + intersect: forall t mGive a args + (post1 post2: mem -> list word -> Prop), + ext_spec t mGive a args post1 -> + ext_spec t mGive a args post2 -> + ext_spec t mGive a args (fun mReceive resvals => + post1 mReceive resvals /\ post2 mReceive resvals); + }. +End ext_spec. +Arguments ext_spec.ok {_ _ _ _} _. + +Section binops. + Context {width : Z} {word : Word.Interface.word width}. + Definition interp_binop (bop : bopname) : word -> word -> word := + match bop with + | bopname.add => word.add + | bopname.sub => word.sub + | bopname.mul => word.mul + | bopname.mulhuu => word.mulhuu + | bopname.divu => word.divu + | bopname.remu => word.modu + | bopname.and => word.and + | bopname.or => word.or + | bopname.xor => word.xor + | bopname.sru => word.sru + | bopname.slu => word.slu + | bopname.srs => word.srs + | bopname.lts => fun a b => + if word.lts a b then word.of_Z 1 else word.of_Z 0 + | bopname.ltu => fun a b => + if word.ltu a b then word.of_Z 1 else word.of_Z 0 + | bopname.eq => fun a b => + if word.eqb a b then word.of_Z 1 else word.of_Z 0 + end. +End binops. + +Definition env: map.map String.string Syntax.func := SortedListString.map _. +#[export] Instance env_ok: map.ok env := SortedListString.ok _. + +Section semantics. + Context {width: Z} {BW: Bitwidth width} {word: word.word width} {mem: map.map word byte}. + Context {locals: map.map String.string word}. + Context {ext_spec: ExtSpec}. + Local Notation metrics := MetricLog. + (* this is the expr evaluator that is used to verify execution time, the just-correctness-oriented version is below *) + Section WithMemAndLocals. + Context (m : mem) (l : locals). + + Fixpoint eval_expr (e : expr) (mc : metrics) : option (word * metrics) := + match e with + | expr.literal v => Some (word.of_Z v, addMetricInstructions 8 + (addMetricLoads 8 mc)) + | expr.var x => match map.get l x with + | Some v => Some (v, addMetricInstructions 1 + (addMetricLoads 2 mc)) + | None => None + end + | expr.inlinetable aSize t index => + 'Some (index', mc') <- eval_expr index mc | None; + 'Some v <- load aSize (map.of_list_word t) index' | None; + Some (v, (addMetricInstructions 3 + (addMetricLoads 4 + (addMetricJumps 1 mc')))) + | expr.load aSize a => + 'Some (a', mc') <- eval_expr a mc | None; + 'Some v <- load aSize m a' | None; + Some (v, addMetricInstructions 1 + (addMetricLoads 2 mc')) + | expr.op op e1 e2 => + 'Some (v1, mc') <- eval_expr e1 mc | None; + 'Some (v2, mc'') <- eval_expr e2 mc' | None; + Some (interp_binop op v1 v2, addMetricInstructions 2 + (addMetricLoads 2 mc'')) + | expr.ite c e1 e2 => + 'Some (vc, mc') <- eval_expr c mc | None; + eval_expr (if word.eqb vc (word.of_Z 0) then e2 else e1) + (addMetricInstructions 2 + (addMetricLoads 2 + (addMetricJumps 1 mc'))) + end. + + Fixpoint evaluate_call_args_log (arges : list expr) (mc : metrics) := + match arges with + | e :: tl => + 'Some (v, mc') <- eval_expr e mc | None; + 'Some (args, mc'') <- evaluate_call_args_log tl mc' | None; + Some (v :: args, mc'') + | _ => Some (nil, mc) + end. + + + End WithMemAndLocals. +End semantics. + +Module exec. Section WithParams. + Context {width: Z} {BW: Bitwidth width} {word: word.word width} {mem: map.map word byte}. + Context {locals: map.map String.string word}. + Context {ext_spec: ExtSpec}. + Section WithEnv. + Context (e: env). + + Local Notation metrics := MetricLog. + + Implicit Types post : trace -> mem -> locals -> metrics -> Prop. (* COQBUG: unification finds Type instead of Prop and fails to downgrade *) + Inductive exec: cmd -> trace -> mem -> locals -> metrics -> + (trace -> mem -> locals -> metrics -> Prop) -> Prop := + | skip: forall t m l mc post, + post t m l mc -> + exec cmd.skip t m l mc post + | set: forall x e t m l post mc v mc', + eval_expr m l e mc = Some (v, mc') -> + post t m (map.put l x v) (addMetricInstructions 1 + (addMetricLoads 1 mc')) -> + exec (cmd.set x e) t m l mc post + | unset: forall x t m l mc post, + post t m (map.remove l x) mc -> + exec (cmd.unset x) t m l mc post + | store: forall sz ea ev t m l mc post a mc' v mc'' m', + eval_expr m l ea mc = Some (a, mc') -> + eval_expr m l ev mc' = Some (v, mc'') -> + store sz m a v = Some m' -> + post t m' l (addMetricInstructions 1 + (addMetricLoads 1 + (addMetricStores 1 mc''))) -> + exec (cmd.store sz ea ev) t m l mc post + | stackalloc: forall x n body t mSmall l mc post, + Z.modulo n (bytes_per_word width) = 0 -> + (forall a mStack mCombined, + anybytes a n mStack -> + map.split mCombined mSmall mStack -> + exec body t mCombined (map.put l x a) (addMetricInstructions 1 (addMetricLoads 1 mc)) + (fun t' mCombined' l' mc' => + exists mSmall' mStack', + anybytes a n mStack' /\ + map.split mCombined' mSmall' mStack' /\ + post t' mSmall' l' mc')) -> + exec (cmd.stackalloc x n body) t mSmall l mc post + | if_true: forall t m l mc e c1 c2 post v mc', + eval_expr m l e mc = Some (v, mc') -> + word.unsigned v <> 0 -> + exec c1 t m l (addMetricInstructions 2 + (addMetricLoads 2 + (addMetricJumps 1 mc'))) post -> + exec (cmd.cond e c1 c2) t m l mc post + | if_false: forall e c1 c2 t m l mc post v mc', + eval_expr m l e mc = Some (v, mc') -> + word.unsigned v = 0 -> + exec c2 t m l (addMetricInstructions 2 + (addMetricLoads 2 + (addMetricJumps 1 mc'))) post -> + exec (cmd.cond e c1 c2) t m l mc post + | seq: forall c1 c2 t m l mc post mid, + exec c1 t m l mc mid -> + (forall t' m' l' mc', mid t' m' l' mc' -> exec c2 t' m' l' mc' post) -> + exec (cmd.seq c1 c2) t m l mc post + | while_false: forall e c t m l mc post v mc', + eval_expr m l e mc = Some (v, mc') -> + word.unsigned v = 0 -> + post t m l (addMetricInstructions 1 + (addMetricLoads 1 + (addMetricJumps 1 mc'))) -> + exec (cmd.while e c) t m l mc post + | while_true: forall e c t m l mc post v mc' mid, + eval_expr m l e mc = Some (v, mc') -> + word.unsigned v <> 0 -> + exec c t m l mc' mid -> + (forall t' m' l' mc'', mid t' m' l' mc'' -> exec (cmd.while e c) t' m' l' (addMetricInstructions 2 + (addMetricLoads 2 + (addMetricJumps 1 mc''))) post) -> + exec (cmd.while e c) t m l mc post + | call: forall binds fname arges t m l mc post params rets fbody args mc' lf mid, + map.get e fname = Some (params, rets, fbody) -> + evaluate_call_args_log m l arges mc = Some (args, mc') -> + map.of_list_zip params args = Some lf -> + exec fbody t m lf (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc')))) mid -> + (forall t' m' st1 mc'', mid t' m' st1 mc''-> + exists retvs, map.getmany_of_list st1 rets = Some retvs /\ + exists l', map.putmany_of_list_zip binds retvs l = Some l' /\ + post t' m' l' (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc''))))) -> + exec (cmd.call binds fname arges) t m l mc post + | interact: forall binds action arges args t m l mc post mKeep mGive mc' mid, + map.split m mKeep mGive -> + evaluate_call_args_log m l arges mc = Some (args, mc') -> + ext_spec t mGive action args mid -> + (forall mReceive resvals, mid mReceive resvals -> + exists l', map.putmany_of_list_zip binds resvals l = Some l' /\ + forall m', map.split m' mKeep mReceive -> + post (cons ((mGive, action, args), (mReceive, resvals)) t) m' l' (addMetricInstructions 1 + (addMetricStores 1 + (addMetricLoads 2 mc')))) -> + exec (cmd.interact binds action arges) t m l mc post. + + Context {word_ok: word.ok word} {mem_ok: map.ok mem} {ext_spec_ok: ext_spec.ok ext_spec}. + + Lemma weaken: forall t l m mc s post1, + exec s t m l mc post1 -> + forall post2, + (forall t' m' l' mc', post1 t' m' l' mc' -> post2 t' m' l' mc') -> + exec s t m l mc post2. + Proof. + induction 1; intros; try solve [econstructor; eauto]. + - eapply stackalloc. 1: assumption. + intros. + eapply H1; eauto. + intros. fwd. eauto 10. + - eapply call. + 4: eapply IHexec. + all: eauto. + intros. + edestruct H3 as (? & ? & ? & ? & ?); [eassumption|]. + eauto 10. + - eapply interact; try eassumption. + intros. + edestruct H2 as (? & ? & ?); [eassumption|]. + eauto 10. + Qed. + + Lemma intersect: forall t l m mc s post1, + exec s t m l mc post1 -> + forall post2, + exec s t m l mc post2 -> + exec s t m l mc (fun t' m' l' mc' => post1 t' m' l' mc' /\ post2 t' m' l' mc'). + Proof. + induction 1; + intros; + match goal with + | H: exec _ _ _ _ _ _ |- _ => inversion H; subst; clear H + end; + try match goal with + | H1: ?e = Some (?x1, ?y1, ?z1), H2: ?e = Some (?x2, ?y2, ?z2) |- _ => + replace x2 with x1 in * by congruence; + replace y2 with y1 in * by congruence; + replace z2 with z1 in * by congruence; + clear x2 y2 z2 H2 + end; + repeat match goal with + | H1: ?e = Some (?v1, ?mc1), H2: ?e = Some (?v2, ?mc2) |- _ => + replace v2 with v1 in * by congruence; + replace mc2 with mc1 in * by congruence; clear H2 + end; + repeat match goal with + | H1: ?e = Some ?v1, H2: ?e = Some ?v2 |- _ => + replace v2 with v1 in * by congruence; clear H2 + end; + try solve [econstructor; eauto | exfalso; congruence]. + + - econstructor. 1: eassumption. + intros. + rename H0 into Ex1, H12 into Ex2. + eapply weaken. 1: eapply H1. 1,2: eassumption. + 1: eapply Ex2. 1,2: eassumption. + cbv beta. + intros. fwd. + lazymatch goal with + | A: map.split _ _ _, B: map.split _ _ _ |- _ => + specialize @map.split_diff with (4 := A) (5 := B) as P + end. + edestruct P; try typeclasses eauto. 2: subst; eauto 10. + eapply anybytes_unique_domain; eassumption. + - econstructor. + + eapply IHexec. exact H5. (* not H *) + + simpl. intros *. intros [? ?]. eauto. + - eapply while_true. 1, 2: eassumption. + + eapply IHexec. exact H9. (* not H1 *) + + simpl. intros *. intros [? ?]. eauto. + - eapply call. 1, 2, 3: eassumption. + + eapply IHexec. exact H16. (* not H2 *) + + simpl. intros *. intros [? ?]. + edestruct H3 as (? & ? & ? & ? & ?); [eassumption|]. + edestruct H17 as (? & ? & ? & ? & ?); [eassumption|]. + repeat match goal with + | H1: ?e = Some ?v1, H2: ?e = Some ?v2 |- _ => + replace v2 with v1 in * by congruence; clear H2 + end. + eauto 10. + - pose proof ext_spec.unique_mGive_footprint as P. + specialize P with (1 := H1) (2 := H14). + destruct (map.split_diff P H H7). subst mKeep0 mGive0. clear H7. + eapply interact. 1,2: eassumption. + + eapply ext_spec.intersect; [ exact H1 | exact H14 ]. + + simpl. intros *. intros [? ?]. + edestruct H2 as (? & ? & ?); [eassumption|]. + edestruct H15 as (? & ? & ?); [eassumption|]. + repeat match goal with + | H1: ?e = Some ?v1, H2: ?e = Some ?v2 |- _ => + replace v2 with v1 in * by congruence; clear H2 + end. + eauto 10. + Qed. + + End WithEnv. + + Lemma extend_env: forall e1 e2, + map.extends e2 e1 -> + forall c t m l mc post, + exec e1 c t m l mc post -> + exec e2 c t m l mc post. + Proof. induction 2; try solve [econstructor; eauto]. Qed. + + End WithParams. +End exec. Notation exec := exec.exec. + +Section WithParams. + Context {width: Z} {BW: Bitwidth width} {word: word.word width} {mem: map.map word byte}. + Context {locals: map.map String.string word}. + Context {ext_spec: ExtSpec}. + Local Notation metrics := MetricLog. + + Implicit Types (l: locals) (m: mem) (post: trace -> mem -> list word -> metrics -> Prop). + + Definition call e fname t m args mc post := + exists argnames retnames body, + map.get e fname = Some (argnames, retnames, body) /\ + exists l, map.of_list_zip argnames args = Some l /\ + exec e body t m l mc (fun t' m' l' mc' => exists rets, + map.getmany_of_list l' retnames = Some rets /\ post t' m' rets mc'). + + Lemma weaken_call: forall e fname t m args mc post1, + call e fname t m args mc post1 -> + forall post2, (forall t' m' rets mc', post1 t' m' rets mc' -> post2 t' m' rets mc') -> + call e fname t m args mc post2. + Proof. + unfold call. intros. fwd. + do 4 eexists. 1: eassumption. + do 2 eexists. 1: eassumption. + eapply exec.weaken. 1: eassumption. + cbv beta. clear -H0. intros. fwd. eauto. + Qed. + + Lemma extend_env_call: forall e1 e2, + map.extends e2 e1 -> + forall f t m rets mc post, + call e1 f t m rets mc post -> + call e2 f t m rets mc post. + Proof. + unfold call. intros. fwd. repeat eexists. + - eapply H. eassumption. + - eassumption. + - eapply exec.extend_env; eassumption. + Qed. +End WithParams. diff --git a/bedrock2/src/bedrock2/Notations.v b/bedrock2/src/bedrock2/Notations.v new file mode 100644 index 000000000..5d1a500e9 --- /dev/null +++ b/bedrock2/src/bedrock2/Notations.v @@ -0,0 +1,21 @@ +Require Import coqutil.Macros.subst coqutil.Macros.unique. + +Notation "' x <- a | y ; f" := + (match a with + | x => f + | _ => y + end) + (right associativity, at level 70, x pattern). + +Notation "'bind_ex' x <- a ; f" := + (subst! a for a' in exists x, a' x /\ f) + (only parsing, right associativity, at level 60, f at level 200). +Notation "'bind_ex_Some' x <- a ; f" := + (subst! a for a' in exists x, a' = Some x /\ f) + (only parsing, right associativity, at level 60, f at level 200). +Notation "'bind_ex_pair' ( x , y ) <- a ; f" := + (subst! a for a' in exists x y, a' (x, y) /\ f) + (only parsing, right associativity, at level 60, f at level 200). +Notation "'bind_eq' x <- a ; f" := + (subst! a for a' in forall x, x = a' -> f) + (only parsing, right associativity, at level 60, f at level 200). diff --git a/bedrock2/src/bedrock2/ProgramLogic.v b/bedrock2/src/bedrock2/ProgramLogic.v index 5bc8e40b5..6bc3afdd6 100644 --- a/bedrock2/src/bedrock2/ProgramLogic.v +++ b/bedrock2/src/bedrock2/ProgramLogic.v @@ -235,14 +235,14 @@ Ltac straightline := eapply start_func; [exact H | clear H] end; cbv match beta delta [WeakestPrecondition.func] - | |- WeakestPrecondition.cmd _ (cmd.set ?s ?e) _ _ _ ?post => + | |- WeakestPrecondition.cmd _ (cmd.set ?s ?e) _ _ _ _ ?post => unfold1_cmd_goal; cbv beta match delta [cmd_body]; let __ := match s with String.String _ _ => idtac | String.EmptyString => idtac end in ident_of_constr_string_cps s ltac:(fun x => ensure_free x; (* NOTE: keep this consistent with the [exists _, _ /\ _] case far below *) letexists _ as x; split; [solve [repeat straightline]|]) - | |- cmd _ ?c _ _ _ ?post => + | |- cmd _ ?c _ _ _ _ ?post => let c := eval hnf in c in lazymatch c with | cmd.while _ _ => fail @@ -250,12 +250,12 @@ Ltac straightline := | cmd.interact _ _ _ => fail | _ => unfold1_cmd_goal; cbv beta match delta [cmd_body] end - | |- @list_map _ _ (get _) _ _ => unfold1_list_map_goal; cbv beta match delta [list_map_body] + | |- @list_map _ _ (get _ _) _ _ => unfold1_list_map_goal; cbv beta match delta [list_map_body] | |- @list_map _ _ (expr _ _) _ _ => unfold1_list_map_goal; cbv beta match delta [list_map_body] | |- @list_map _ _ _ nil _ => cbv beta match fix delta [list_map list_map_body] - | |- expr _ _ _ _ => unfold1_expr_goal; cbv beta match delta [expr_body] - | |- dexpr _ _ _ _ => cbv beta delta [dexpr] - | |- dexprs _ _ _ _ => cbv beta delta [dexprs] + | |- expr _ _ _ _ _ => unfold1_expr_goal; cbv beta match delta [expr_body] + | |- dexpr _ _ _ _ _ => cbv beta delta [dexpr] + | |- dexprs _ _ _ _ _ => cbv beta delta [dexprs] | |- literal _ _ => cbv beta delta [literal] | |- @get ?w ?W ?L ?l ?x ?P => let get' := eval cbv [get] in @get in diff --git a/bedrock2/src/bedrock2Examples/tckmnipow.v b/bedrock2/src/bedrock2Examples/tckmnipow.v new file mode 100644 index 000000000..b0a0780ab --- /dev/null +++ b/bedrock2/src/bedrock2Examples/tckmnipow.v @@ -0,0 +1,469 @@ +Require Import Coq.ZArith.ZArith coqutil.Z.div_mod_to_equations. +Require Import bedrock2.NotationsCustomEntry. +Require Import bedrock2.MetricLogging. +Import Syntax BinInt String List.ListNotations ZArith. +Require Import coqutil.Z.Lia. +Local Open Scope string_scope. Local Open Scope Z_scope. Local Open Scope list_scope. +Local Coercion literal (z : Z) : Syntax.expr := Syntax.expr.literal z. + +Require Import bedrock2.Syntax. +Require Import bedrock2.WeakestPrecondition. +Require Import bedrock2.WeakestPreconditionProperties. +Require Import bedrock2.Loops. +Require Import bedrock2.Map.SeparationLogic. +Module Import Coercions. + Import Map.Interface Word.Interface BinInt. + Coercion Z.of_nat : nat >-> Z. + Coercion word.unsigned : word.rep >-> Z. +Require Import coqutil.Tactics.eplace coqutil.Tactics.eabstract Coq.setoid_ring.Ring_tac. +From coqutil.Tactics Require Import Tactics letexists eabstract rdelta reference_to_string ident_of_string. + + +Definition ipow := func! (x, e) ~> ret { + ret = $1; + while (e) { + if (e & $1) { ret = ret * x }; + e = e >> $1; + x = x * x + } + }. + +From bedrock2 Require Import Semantics BasicC64Semantics WeakestPrecondition ProgramLogic. +From coqutil Require Import Word.Properties Word.Interface Tactics.letexists. +Import Interface.word. + +Definition initCost := {| instructions := 9; stores := 0; loads := 9; jumps := 0 |}. +Definition iterCost := {| instructions := 38; stores := 0; loads := 45; jumps := 2 |}. +Definition endCost := {| instructions := 2; stores := 0; loads := 3; jumps := 1 |}. + +Definition msb z := match z with + | Zpos _ => Z.log2 z + 1 + | _ => 0 + end. + +#[export] Instance spec_of_ipow : spec_of "ipow" := + fnspec! "ipow" x e ~> v, + { requires t m mc := True; + ensures t' m' mc' := unsigned v = unsigned x ^ unsigned e mod 2^64 /\ + (mc' - mc <= initCost + (msb (word.unsigned e)) * iterCost + endCost)%metricsH}. + +Module Z. + Lemma pow_mod x n m (Hnz: m <> 0) : (x mod m)^n mod m = x^n mod m. + Proof. + revert n. + eapply Z.order_induction_0; intros. + { intros ???; subst; split; auto. } + { rewrite 2Z.pow_0_r; trivial. } + { rewrite 2Z.pow_succ_r by trivial. + rewrite <-Z.mul_mod_idemp_r by trivial. + multimatch goal with H: _ |- _ => rewrite H end; + rewrite Z.mul_mod_idemp_l, Z.mul_mod_idemp_r; solve[trivial]. } + { rewrite 2Z.pow_neg_r; trivial. } + Qed. + + Lemma mod2_nonzero x : x mod 2 <> 0 -> x mod 2 = 1. + Proof. Z.div_mod_to_equations. blia. Qed. + + Lemma land_1_r x : Z.land x 1 = x mod 2. + Proof. + change 1 with (Z.ones 1) in *. + rewrite Z.land_ones in * by discriminate. + exact eq_refl. + Qed. +End Z. + +Require Import bedrock2.AbsintWordToZ coqutil.Z.Lia. + +Ltac t := + repeat match goal with x := _ |- _ => subst x end; + repeat match goal with |- context [word.unsigned ?e] => progress (idtac; let H := rbounded (word.unsigned e) in idtac) end; + repeat match goal with G: context [word.unsigned ?e] |- _ => progress (idtac; let H := rbounded (word.unsigned e) in idtac) end; + repeat match goal with |- context [word.unsigned ?e] => progress (idtac; let H := unsigned.zify_expr e in try rewrite H) end; + repeat match goal with G: context [word.unsigned ?e] |- _ => progress (idtac; let H := unsigned.zify_expr e in try rewrite H in G) end; + repeat match goal with H: absint_eq ?x ?x |- _ => clear H end; + cbv [absint_eq] in *. + +Lemma msb_shift z : 0 < z -> msb (z / 2) = msb z - 1. +Proof. + intro. + case (z / 2) eqn:Hdiv. + - enough (H1 : z = 1) by (rewrite H1; easy). + enough (z = z mod 2) by (Z.div_mod_to_equations; blia). + rewrite (Z.div_mod z 2) by blia. + rewrite Hdiv. + cbn. + rewrite Zmod_mod. + reflexivity. + - rewrite <- Z.div2_div in Hdiv. + rewrite (Zdiv2_odd_eqn z). + rewrite Hdiv. + rewrite <- Pos2Z.inj_mul. + case (Z.odd z); + [rewrite <- Pos2Z.inj_add | rewrite Z.add_0_r]; + unfold msb; + rewrite Z.add_simpl_r; + [rewrite Pos2Z.inj_add |]; rewrite Pos2Z.inj_mul; + [rewrite Z.log2_succ_double | rewrite Z.log2_double]; + blia. + - pose proof (Zlt_neg_0 p) as Hneg. + rewrite <- Hdiv in Hneg. + Z.div_mod_to_equations. + blia. +Qed. + +Lemma ipow_ok : program_logic_goal_for_function! ipow. +Proof. + + straightline. + straightline. + straightline. + straightline. + + + (* first one works, second one doesnt? *) + (* + match goal with + | |- WeakestPrecondition.cmd _ (cmd.set ?s ?e) _ _ _ _ ?post => + unfold1_cmd_goal; cbv beta match delta [cmd_body]; + let __ := match s with String.String _ _ => idtac | String.EmptyString => idtac end in + ident_of_constr_string_cps s ltac:(fun x => + ensure_free x; hnf; + lazymatch goal with + | |- exists x', ?P => + refine (let x := _ in ex_intro (fun x' => P) x _) + end) + end. + *) + (* + match goal with + | |- WeakestPrecondition.cmd _ (cmd.set ?s ?e) _ _ _ _ ?post => + unfold1_cmd_goal; cbv beta match delta [cmd_body]; + let __ := match s with String.String _ _ => idtac | String.EmptyString => idtac end in + ident_of_constr_string_cps s ltac:(fun x => + ensure_free x; + letexists_as _ x) +end. +*) + + match goal with + | |- WeakestPrecondition.cmd _ (cmd.set ?s ?e) _ _ _ _ ?post => + unfold1_cmd_goal; cbv beta match delta [cmd_body]; + let __ := match s with String.String _ _ => idtac | String.EmptyString => idtac end in + ident_of_constr_string_cps s ltac:(fun x => + ensure_free x; hnf; + lazymatch goal with + | |- exists x', ?P => + refine (let x := _ in ex_intro (fun x' => P) x _) + end); + let __ := match s with String.String _ _ => idtac | String.EmptyString => idtac end in + ident_of_constr_string_cps s ltac:(fun x => + ensure_free x; hnf; + lazymatch goal with + | |- exists x', ?P => + refine (let x := _ in ex_intro (fun x' => P) x _) + end) + end. + split; try(repeat straightline; subst ret ret'0; cbv [literal dlet.dlet]; trivial). + + refine ((Loops.tailrec + (* types of ghost variables*) HList.polymorphic_list.nil + (* program variables *) (["e";"ret";"x"] : list String.string)) + (fun v t m e ret x mc => PrimitivePair.pair.mk (v = word.unsigned e) (* precondition *) + (fun T M E RET X MC=> T = t /\ M = m /\ (* postcondition *) + word.unsigned RET = word.unsigned ret * word.unsigned x ^ word.unsigned e mod 2^64 /\ + (MC - mc <= msb (word.unsigned e) * iterCost + endCost)%metricsH)) + (fun n m => 0 <= n < m) (* well_founded relation *) + _ _ _ _ _) ; + (* TODO wrap this into a tactic with the previous refine *) + cbn [HList.hlist.foralls HList.tuple.foralls + HList.hlist.existss HList.tuple.existss + HList.hlist.apply HList.tuple.apply + HList.hlist + List.repeat Datatypes.length + HList.polymorphic_list.repeat HList.polymorphic_list.length + PrimitivePair.pair._1 PrimitivePair.pair._2] in *. + + { repeat straightline. } + { exact (Z.lt_wf _). } + { repeat straightline. } (* init precondition *) + { (* loop test *) + repeat straightline; try show_program. + { (* loop body *) + eexists. exists (addMetricInstructions 11 (addMetricLoads 12 mc')). split; [repeat straightline|]. (* if condition evaluation *) + (*this case shouldnt exist*) 1:{ cbv [literal dlet.dlet]. + cbv [addMetricLoads withLoads addMetricInstructions withInstructions instructions stores loads jumps]. + simpl. + repeat ( + try rewrite applyAddInstructions; + try rewrite applyAddStores; + try rewrite applyAddLoads; + try rewrite applyAddJumps + ); + repeat rewrite <- Z.add_assoc; + cbn [Z.add Pos.add Pos.succ] +. trivial. + } + split. (* if cases, path-blasting *) + { + (straightline || (split; trivial; [])). + (straightline || (split; trivial; [])). + (*once again i am forced to do things i would rather not*) + unfold1_cmd_goal. + cbv beta match delta [cmd_body]. + exists (mul x1 x2). exists (addMetricInstructions 17 (addMetricLoads 20 (addMetricJumps 1 mc'))). + split. 1: { repeat straightline. + cbv [addMetricLoads withLoads addMetricInstructions withInstructions instructions stores loads jumps]. + simpl. + repeat ( + try rewrite applyAddInstructions; + try rewrite applyAddStores; + try rewrite applyAddLoads; + try rewrite applyAddJumps + ); + repeat rewrite <- Z.add_assoc; + cbn [Z.add Pos.add Pos.succ] +. trivial. + } + (straightline || (split; trivial; [])). + unfold1_cmd_goal. + cbv beta match delta [cmd_body]. + exists (sru x0 (word.of_Z 1)). exists (addMetricInstructions 29 (addMetricLoads 33 (addMetricJumps 1 mc'))). + split. 1: { repeat straightline. cbv [literal dlet.dlet]. + cbv [addMetricLoads withLoads addMetricInstructions withInstructions instructions stores loads jumps]. + simpl. + repeat ( + try rewrite applyAddInstructions; + try rewrite applyAddStores; + try rewrite applyAddLoads; + try rewrite applyAddJumps + ); + repeat rewrite <- Z.add_assoc; + cbn [Z.add Pos.add Pos.succ] +. rewrite Pos.add_carry_spec. + (*pretty sure this adds up, but i think i'd like to sort out the automation problems before pressing ahead any farther*) Abort. + } + (* :( *) + 2: split. all:t. + { (* measure decreases *) + set (word.unsigned x0) in *. (* WHY does blia need this? *) + Z.div_mod_to_equations. blia. } + { (* invariant preserved *) + rewrite H3; clear H3. rename H0 into Hbit. + change (1+1) with 2 in *. + eapply Z.mod2_nonzero in Hbit. + epose proof (Z.div_mod _ 2 ltac:(discriminate)) as Heq; rewrite Hbit in Heq. + rewrite Heq at 2; clear Hbit Heq. + (* rewriting with equivalence modulo ... *) + rewrite !word.unsigned_mul. + unfold word.wrap. + rewrite ?Z.mul_mod_idemp_l by discriminate. + rewrite <-(Z.mul_mod_idemp_r _ (_^_)), Z.pow_mod by discriminate. + rewrite ?Z.pow_add_r by (pose proof word.unsigned_range x0; Z.div_mod_to_equations; blia). + rewrite ?Z.pow_twice_r, ?Z.pow_1_r, ?Z.pow_mul_l. + rewrite Z.mul_mod_idemp_r by discriminate. + f_equal; ring. } + { (* metrics correct *) + cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. + applyAddMetrics H4. + rewrite msb_shift in H4 by blia. + rewrite MetricArith.mul_sub_distr_r in H4. + rewrite <- MetricArith.add_sub_swap in H4. + rewrite <- MetricArith.le_add_le_sub_r in H4. + eapply MetricArith.le_trans. + 2: exact H4. + unfold iterCost. + solve_MetricLog. + } + } + { + repeat (straightline || (split; trivial; [])). 2: split. all: t. + { (* measure decreases *) + set (word.unsigned x0) in *. (* WHY does blia need this? *) + Z.div_mod_to_equations; blia. } + { (* invariant preserved *) + rewrite H3; clear H3. rename H0 into Hbit. + change (1+1) with 2 in *. + epose proof (Z.div_mod _ 2 ltac:(discriminate)) as Heq; rewrite Hbit in Heq. + rewrite Heq at 2; clear Hbit Heq. + (* rewriting with equivalence modulo ... *) + rewrite !word.unsigned_mul, ?Z.mul_mod_idemp_l by discriminate. + cbv [word.wrap]. + rewrite <-(Z.mul_mod_idemp_r _ (_^_)), Z.pow_mod by discriminate. + rewrite ?Z.add_0_r, Z.pow_twice_r, ?Z.pow_1_r, ?Z.pow_mul_l. + rewrite Z.mul_mod_idemp_r by discriminate. + f_equal; ring. } + { (* metrics correct *) + cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. + applyAddMetrics H4. + rewrite msb_shift in H4 by blia. + rewrite MetricArith.mul_sub_distr_r in H4. + rewrite <- MetricArith.add_sub_swap in H4. + rewrite <- MetricArith.le_add_le_sub_r in H4. + eapply MetricArith.le_trans. + 2: exact H4. + unfold iterCost. + solve_MetricLog. + } + } + } + { (* postcondition *) + rewrite H, Z.pow_0_r, Z.mul_1_r, word.wrap_unsigned. + repeat match goal with + | |- _ /\ _ => split + | |- _ = _ => reflexivity + end. + unfold msb, iterCost, endCost. + subst mc'. + solve_MetricLog. + } + } + + repeat straightline. + repeat match goal with + | |- _ /\ _ => split + | |- _ = _ => reflexivity + | |- exists _, _ => letexists + | _ => t + end. + 1: setoid_rewrite H1; setoid_rewrite Z.mul_1_l; trivial. + + unfold initCost, iterCost, endCost in *. + solve_MetricLog. + + (*end hacking*) + + + + (* + Print straightline. + (let __ := match s with + | String.String _ _ => idtac + | "" => idtac + end in + ident_of_string.ident_of_constr_string_cps s + ltac:(fun x => ensure_free x; letexists _ as x; split; [ solve [ repeat straightline ] | ])). + *) + + straightline. + straightline. + repeat straightline. + cbv beta delta [dexpr]. + eexists; eexists; split. + - repeat straightline. trivial. unfold1_expr_goal; cbv beta match delta [expr_body]. straightline. + + Print straightline. + + Check ((Loops.tailrec + (* types of ghost variables*) HList.polymorphic_list.nil + (* program variables *) (["e";"ret";"x"] : list String.string)) + (fun v t m e ret x mc => PrimitivePair.pair.mk (v = word.unsigned e) (* precondition *) + (fun T M E RET X MC=> T = t /\ M = m /\ (* postcondition *) + word.unsigned RET = word.unsigned ret * word.unsigned x ^ word.unsigned e mod 2^64 /\ + (MC - mc <= msb (word.unsigned e) * iterCost + endCost)%metricsH)) + (fun n m => 0 <= n < m) (* well_founded relation *) + _ _ _ _ _). ; + (* TODO wrap this into a tactic with the previous refine *) + cbn [HList.hlist.foralls HList.tuple.foralls + HList.hlist.existss HList.tuple.existss + HList.hlist.apply HList.tuple.apply + HList.hlist + List.repeat Datatypes.length + HList.polymorphic_list.repeat HList.polymorphic_list.length + PrimitivePair.pair._1 PrimitivePair.pair._2] in *. + + { repeat straightline. } + { exact (Z.lt_wf _). } + { repeat straightline. } (* init precondition *) + { (* loop test *) + repeat straightline; try show_program. + { (* loop body *) + eexists; eexists; split; [repeat straightline|]. (* if condition evaluation *) + split. (* if cases, path-blasting *) + { + repeat (straightline || (split; trivial; [])). 2: split. all:t. + { (* measure decreases *) + set (word.unsigned x0) in *. (* WHY does blia need this? *) + Z.div_mod_to_equations. blia. } + { (* invariant preserved *) + rewrite H3; clear H3. rename H0 into Hbit. + change (1+1) with 2 in *. + eapply Z.mod2_nonzero in Hbit. + epose proof (Z.div_mod _ 2 ltac:(discriminate)) as Heq; rewrite Hbit in Heq. + rewrite Heq at 2; clear Hbit Heq. + (* rewriting with equivalence modulo ... *) + rewrite !word.unsigned_mul. + unfold word.wrap. + rewrite ?Z.mul_mod_idemp_l by discriminate. + rewrite <-(Z.mul_mod_idemp_r _ (_^_)), Z.pow_mod by discriminate. + rewrite ?Z.pow_add_r by (pose proof word.unsigned_range x0; Z.div_mod_to_equations; blia). + rewrite ?Z.pow_twice_r, ?Z.pow_1_r, ?Z.pow_mul_l. + rewrite Z.mul_mod_idemp_r by discriminate. + f_equal; ring. } + { (* metrics correct *) + cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. + applyAddMetrics H4. + rewrite msb_shift in H4 by blia. + rewrite MetricArith.mul_sub_distr_r in H4. + rewrite <- MetricArith.add_sub_swap in H4. + rewrite <- MetricArith.le_add_le_sub_r in H4. + eapply MetricArith.le_trans. + 2: exact H4. + unfold iterCost. + solve_MetricLog. + } + } + { + repeat (straightline || (split; trivial; [])). 2: split. all: t. + { (* measure decreases *) + set (word.unsigned x0) in *. (* WHY does blia need this? *) + Z.div_mod_to_equations; blia. } + { (* invariant preserved *) + rewrite H3; clear H3. rename H0 into Hbit. + change (1+1) with 2 in *. + epose proof (Z.div_mod _ 2 ltac:(discriminate)) as Heq; rewrite Hbit in Heq. + rewrite Heq at 2; clear Hbit Heq. + (* rewriting with equivalence modulo ... *) + rewrite !word.unsigned_mul, ?Z.mul_mod_idemp_l by discriminate. + cbv [word.wrap]. + rewrite <-(Z.mul_mod_idemp_r _ (_^_)), Z.pow_mod by discriminate. + rewrite ?Z.add_0_r, Z.pow_twice_r, ?Z.pow_1_r, ?Z.pow_mul_l. + rewrite Z.mul_mod_idemp_r by discriminate. + f_equal; ring. } + { (* metrics correct *) + cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. + applyAddMetrics H4. + rewrite msb_shift in H4 by blia. + rewrite MetricArith.mul_sub_distr_r in H4. + rewrite <- MetricArith.add_sub_swap in H4. + rewrite <- MetricArith.le_add_le_sub_r in H4. + eapply MetricArith.le_trans. + 2: exact H4. + unfold iterCost. + solve_MetricLog. + } + } + } + { (* postcondition *) + rewrite H, Z.pow_0_r, Z.mul_1_r, word.wrap_unsigned. + repeat match goal with + | |- _ /\ _ => split + | |- _ = _ => reflexivity + end. + unfold msb, iterCost, endCost. + subst mc'. + solve_MetricLog. + } + } + + repeat straightline. + repeat match goal with + | |- _ /\ _ => split + | |- _ = _ => reflexivity + | |- exists _, _ => letexists + | _ => t + end. + 1: setoid_rewrite H1; setoid_rewrite Z.mul_1_l; trivial. + + unfold initCost, iterCost, endCost in *. + solve_MetricLog. +Qed. From 9a0f6fd895132ea0198322b40df2c433869a529b Mon Sep 17 00:00:00 2001 From: nathan-sheffield Date: Mon, 9 Oct 2023 14:30:32 -0400 Subject: [PATCH 18/62] sorry to leave you with this andy --- .../NathanDidntNoticeMetricSemantics.v | 385 ------------------ 1 file changed, 385 deletions(-) delete mode 100644 bedrock2/src/bedrock2/NathanDidntNoticeMetricSemantics.v diff --git a/bedrock2/src/bedrock2/NathanDidntNoticeMetricSemantics.v b/bedrock2/src/bedrock2/NathanDidntNoticeMetricSemantics.v deleted file mode 100644 index 597817120..000000000 --- a/bedrock2/src/bedrock2/NathanDidntNoticeMetricSemantics.v +++ /dev/null @@ -1,385 +0,0 @@ -Require Import coqutil.sanity coqutil.Macros.subst coqutil.Macros.unique coqutil.Byte. -Require Import coqutil.Datatypes.PrimitivePair coqutil.Datatypes.HList. -Require Import coqutil.Decidable. -Require Import coqutil.Tactics.fwd. -Require Import coqutil.Map.Properties. -Require coqutil.Map.SortedListString. -Require Import bedrock2.Notations bedrock2.Syntax coqutil.Map.Interface coqutil.Map.OfListWord. -Require Import BinIntDef coqutil.Word.Interface coqutil.Word.Bitwidth. -Require Import bedrock2.MetricLogging. -Require Export bedrock2.Memory. -Require Import Coq.Lists.List. - -(* BW is not needed on the rhs, but helps infer width *) -Definition trace{width: Z}{BW: Bitwidth width}{word: word.word width}{mem: map.map word byte} := - list ((mem * String.string * list word) * (mem * list word)). - -Definition ExtSpec{width: Z}{BW: Bitwidth width}{word: word.word width}{mem: map.map word byte} := - (* Given a trace of what happened so far, - the given-away memory, an action label and a list of function call arguments, *) - trace -> mem -> String.string -> list word -> - (* and a postcondition on the received memory and function call results, *) - (mem -> list word -> Prop) -> - (* tells if this postcondition will hold *) - Prop. - -Existing Class ExtSpec. - -Module ext_spec. - Class ok{width: Z}{BW: Bitwidth width}{word: word.word width}{mem: map.map word byte} - {ext_spec: ExtSpec}: Prop := - { - (* The action name and arguments uniquely determine the footprint of the given-away memory. *) - unique_mGive_footprint: forall t1 t2 mGive1 mGive2 a args - (post1 post2: mem -> list word -> Prop), - ext_spec t1 mGive1 a args post1 -> - ext_spec t2 mGive2 a args post2 -> - map.same_domain mGive1 mGive2; - - weaken :> forall t mGive act args, - Morphisms.Proper - (Morphisms.respectful - (Morphisms.pointwise_relation Interface.map.rep - (Morphisms.pointwise_relation (list word) Basics.impl)) Basics.impl) - (ext_spec t mGive act args); - - intersect: forall t mGive a args - (post1 post2: mem -> list word -> Prop), - ext_spec t mGive a args post1 -> - ext_spec t mGive a args post2 -> - ext_spec t mGive a args (fun mReceive resvals => - post1 mReceive resvals /\ post2 mReceive resvals); - }. -End ext_spec. -Arguments ext_spec.ok {_ _ _ _} _. - -Section binops. - Context {width : Z} {word : Word.Interface.word width}. - Definition interp_binop (bop : bopname) : word -> word -> word := - match bop with - | bopname.add => word.add - | bopname.sub => word.sub - | bopname.mul => word.mul - | bopname.mulhuu => word.mulhuu - | bopname.divu => word.divu - | bopname.remu => word.modu - | bopname.and => word.and - | bopname.or => word.or - | bopname.xor => word.xor - | bopname.sru => word.sru - | bopname.slu => word.slu - | bopname.srs => word.srs - | bopname.lts => fun a b => - if word.lts a b then word.of_Z 1 else word.of_Z 0 - | bopname.ltu => fun a b => - if word.ltu a b then word.of_Z 1 else word.of_Z 0 - | bopname.eq => fun a b => - if word.eqb a b then word.of_Z 1 else word.of_Z 0 - end. -End binops. - -Definition env: map.map String.string Syntax.func := SortedListString.map _. -#[export] Instance env_ok: map.ok env := SortedListString.ok _. - -Section semantics. - Context {width: Z} {BW: Bitwidth width} {word: word.word width} {mem: map.map word byte}. - Context {locals: map.map String.string word}. - Context {ext_spec: ExtSpec}. - Local Notation metrics := MetricLog. - (* this is the expr evaluator that is used to verify execution time, the just-correctness-oriented version is below *) - Section WithMemAndLocals. - Context (m : mem) (l : locals). - - Fixpoint eval_expr (e : expr) (mc : metrics) : option (word * metrics) := - match e with - | expr.literal v => Some (word.of_Z v, addMetricInstructions 8 - (addMetricLoads 8 mc)) - | expr.var x => match map.get l x with - | Some v => Some (v, addMetricInstructions 1 - (addMetricLoads 2 mc)) - | None => None - end - | expr.inlinetable aSize t index => - 'Some (index', mc') <- eval_expr index mc | None; - 'Some v <- load aSize (map.of_list_word t) index' | None; - Some (v, (addMetricInstructions 3 - (addMetricLoads 4 - (addMetricJumps 1 mc')))) - | expr.load aSize a => - 'Some (a', mc') <- eval_expr a mc | None; - 'Some v <- load aSize m a' | None; - Some (v, addMetricInstructions 1 - (addMetricLoads 2 mc')) - | expr.op op e1 e2 => - 'Some (v1, mc') <- eval_expr e1 mc | None; - 'Some (v2, mc'') <- eval_expr e2 mc' | None; - Some (interp_binop op v1 v2, addMetricInstructions 2 - (addMetricLoads 2 mc'')) - | expr.ite c e1 e2 => - 'Some (vc, mc') <- eval_expr c mc | None; - eval_expr (if word.eqb vc (word.of_Z 0) then e2 else e1) - (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc'))) - end. - - Fixpoint evaluate_call_args_log (arges : list expr) (mc : metrics) := - match arges with - | e :: tl => - 'Some (v, mc') <- eval_expr e mc | None; - 'Some (args, mc'') <- evaluate_call_args_log tl mc' | None; - Some (v :: args, mc'') - | _ => Some (nil, mc) - end. - - - End WithMemAndLocals. -End semantics. - -Module exec. Section WithParams. - Context {width: Z} {BW: Bitwidth width} {word: word.word width} {mem: map.map word byte}. - Context {locals: map.map String.string word}. - Context {ext_spec: ExtSpec}. - Section WithEnv. - Context (e: env). - - Local Notation metrics := MetricLog. - - Implicit Types post : trace -> mem -> locals -> metrics -> Prop. (* COQBUG: unification finds Type instead of Prop and fails to downgrade *) - Inductive exec: cmd -> trace -> mem -> locals -> metrics -> - (trace -> mem -> locals -> metrics -> Prop) -> Prop := - | skip: forall t m l mc post, - post t m l mc -> - exec cmd.skip t m l mc post - | set: forall x e t m l post mc v mc', - eval_expr m l e mc = Some (v, mc') -> - post t m (map.put l x v) (addMetricInstructions 1 - (addMetricLoads 1 mc')) -> - exec (cmd.set x e) t m l mc post - | unset: forall x t m l mc post, - post t m (map.remove l x) mc -> - exec (cmd.unset x) t m l mc post - | store: forall sz ea ev t m l mc post a mc' v mc'' m', - eval_expr m l ea mc = Some (a, mc') -> - eval_expr m l ev mc' = Some (v, mc'') -> - store sz m a v = Some m' -> - post t m' l (addMetricInstructions 1 - (addMetricLoads 1 - (addMetricStores 1 mc''))) -> - exec (cmd.store sz ea ev) t m l mc post - | stackalloc: forall x n body t mSmall l mc post, - Z.modulo n (bytes_per_word width) = 0 -> - (forall a mStack mCombined, - anybytes a n mStack -> - map.split mCombined mSmall mStack -> - exec body t mCombined (map.put l x a) (addMetricInstructions 1 (addMetricLoads 1 mc)) - (fun t' mCombined' l' mc' => - exists mSmall' mStack', - anybytes a n mStack' /\ - map.split mCombined' mSmall' mStack' /\ - post t' mSmall' l' mc')) -> - exec (cmd.stackalloc x n body) t mSmall l mc post - | if_true: forall t m l mc e c1 c2 post v mc', - eval_expr m l e mc = Some (v, mc') -> - word.unsigned v <> 0 -> - exec c1 t m l (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc'))) post -> - exec (cmd.cond e c1 c2) t m l mc post - | if_false: forall e c1 c2 t m l mc post v mc', - eval_expr m l e mc = Some (v, mc') -> - word.unsigned v = 0 -> - exec c2 t m l (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc'))) post -> - exec (cmd.cond e c1 c2) t m l mc post - | seq: forall c1 c2 t m l mc post mid, - exec c1 t m l mc mid -> - (forall t' m' l' mc', mid t' m' l' mc' -> exec c2 t' m' l' mc' post) -> - exec (cmd.seq c1 c2) t m l mc post - | while_false: forall e c t m l mc post v mc', - eval_expr m l e mc = Some (v, mc') -> - word.unsigned v = 0 -> - post t m l (addMetricInstructions 1 - (addMetricLoads 1 - (addMetricJumps 1 mc'))) -> - exec (cmd.while e c) t m l mc post - | while_true: forall e c t m l mc post v mc' mid, - eval_expr m l e mc = Some (v, mc') -> - word.unsigned v <> 0 -> - exec c t m l mc' mid -> - (forall t' m' l' mc'', mid t' m' l' mc'' -> exec (cmd.while e c) t' m' l' (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc''))) post) -> - exec (cmd.while e c) t m l mc post - | call: forall binds fname arges t m l mc post params rets fbody args mc' lf mid, - map.get e fname = Some (params, rets, fbody) -> - evaluate_call_args_log m l arges mc = Some (args, mc') -> - map.of_list_zip params args = Some lf -> - exec fbody t m lf (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc')))) mid -> - (forall t' m' st1 mc'', mid t' m' st1 mc''-> - exists retvs, map.getmany_of_list st1 rets = Some retvs /\ - exists l', map.putmany_of_list_zip binds retvs l = Some l' /\ - post t' m' l' (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc''))))) -> - exec (cmd.call binds fname arges) t m l mc post - | interact: forall binds action arges args t m l mc post mKeep mGive mc' mid, - map.split m mKeep mGive -> - evaluate_call_args_log m l arges mc = Some (args, mc') -> - ext_spec t mGive action args mid -> - (forall mReceive resvals, mid mReceive resvals -> - exists l', map.putmany_of_list_zip binds resvals l = Some l' /\ - forall m', map.split m' mKeep mReceive -> - post (cons ((mGive, action, args), (mReceive, resvals)) t) m' l' (addMetricInstructions 1 - (addMetricStores 1 - (addMetricLoads 2 mc')))) -> - exec (cmd.interact binds action arges) t m l mc post. - - Context {word_ok: word.ok word} {mem_ok: map.ok mem} {ext_spec_ok: ext_spec.ok ext_spec}. - - Lemma weaken: forall t l m mc s post1, - exec s t m l mc post1 -> - forall post2, - (forall t' m' l' mc', post1 t' m' l' mc' -> post2 t' m' l' mc') -> - exec s t m l mc post2. - Proof. - induction 1; intros; try solve [econstructor; eauto]. - - eapply stackalloc. 1: assumption. - intros. - eapply H1; eauto. - intros. fwd. eauto 10. - - eapply call. - 4: eapply IHexec. - all: eauto. - intros. - edestruct H3 as (? & ? & ? & ? & ?); [eassumption|]. - eauto 10. - - eapply interact; try eassumption. - intros. - edestruct H2 as (? & ? & ?); [eassumption|]. - eauto 10. - Qed. - - Lemma intersect: forall t l m mc s post1, - exec s t m l mc post1 -> - forall post2, - exec s t m l mc post2 -> - exec s t m l mc (fun t' m' l' mc' => post1 t' m' l' mc' /\ post2 t' m' l' mc'). - Proof. - induction 1; - intros; - match goal with - | H: exec _ _ _ _ _ _ |- _ => inversion H; subst; clear H - end; - try match goal with - | H1: ?e = Some (?x1, ?y1, ?z1), H2: ?e = Some (?x2, ?y2, ?z2) |- _ => - replace x2 with x1 in * by congruence; - replace y2 with y1 in * by congruence; - replace z2 with z1 in * by congruence; - clear x2 y2 z2 H2 - end; - repeat match goal with - | H1: ?e = Some (?v1, ?mc1), H2: ?e = Some (?v2, ?mc2) |- _ => - replace v2 with v1 in * by congruence; - replace mc2 with mc1 in * by congruence; clear H2 - end; - repeat match goal with - | H1: ?e = Some ?v1, H2: ?e = Some ?v2 |- _ => - replace v2 with v1 in * by congruence; clear H2 - end; - try solve [econstructor; eauto | exfalso; congruence]. - - - econstructor. 1: eassumption. - intros. - rename H0 into Ex1, H12 into Ex2. - eapply weaken. 1: eapply H1. 1,2: eassumption. - 1: eapply Ex2. 1,2: eassumption. - cbv beta. - intros. fwd. - lazymatch goal with - | A: map.split _ _ _, B: map.split _ _ _ |- _ => - specialize @map.split_diff with (4 := A) (5 := B) as P - end. - edestruct P; try typeclasses eauto. 2: subst; eauto 10. - eapply anybytes_unique_domain; eassumption. - - econstructor. - + eapply IHexec. exact H5. (* not H *) - + simpl. intros *. intros [? ?]. eauto. - - eapply while_true. 1, 2: eassumption. - + eapply IHexec. exact H9. (* not H1 *) - + simpl. intros *. intros [? ?]. eauto. - - eapply call. 1, 2, 3: eassumption. - + eapply IHexec. exact H16. (* not H2 *) - + simpl. intros *. intros [? ?]. - edestruct H3 as (? & ? & ? & ? & ?); [eassumption|]. - edestruct H17 as (? & ? & ? & ? & ?); [eassumption|]. - repeat match goal with - | H1: ?e = Some ?v1, H2: ?e = Some ?v2 |- _ => - replace v2 with v1 in * by congruence; clear H2 - end. - eauto 10. - - pose proof ext_spec.unique_mGive_footprint as P. - specialize P with (1 := H1) (2 := H14). - destruct (map.split_diff P H H7). subst mKeep0 mGive0. clear H7. - eapply interact. 1,2: eassumption. - + eapply ext_spec.intersect; [ exact H1 | exact H14 ]. - + simpl. intros *. intros [? ?]. - edestruct H2 as (? & ? & ?); [eassumption|]. - edestruct H15 as (? & ? & ?); [eassumption|]. - repeat match goal with - | H1: ?e = Some ?v1, H2: ?e = Some ?v2 |- _ => - replace v2 with v1 in * by congruence; clear H2 - end. - eauto 10. - Qed. - - End WithEnv. - - Lemma extend_env: forall e1 e2, - map.extends e2 e1 -> - forall c t m l mc post, - exec e1 c t m l mc post -> - exec e2 c t m l mc post. - Proof. induction 2; try solve [econstructor; eauto]. Qed. - - End WithParams. -End exec. Notation exec := exec.exec. - -Section WithParams. - Context {width: Z} {BW: Bitwidth width} {word: word.word width} {mem: map.map word byte}. - Context {locals: map.map String.string word}. - Context {ext_spec: ExtSpec}. - Local Notation metrics := MetricLog. - - Implicit Types (l: locals) (m: mem) (post: trace -> mem -> list word -> metrics -> Prop). - - Definition call e fname t m args mc post := - exists argnames retnames body, - map.get e fname = Some (argnames, retnames, body) /\ - exists l, map.of_list_zip argnames args = Some l /\ - exec e body t m l mc (fun t' m' l' mc' => exists rets, - map.getmany_of_list l' retnames = Some rets /\ post t' m' rets mc'). - - Lemma weaken_call: forall e fname t m args mc post1, - call e fname t m args mc post1 -> - forall post2, (forall t' m' rets mc', post1 t' m' rets mc' -> post2 t' m' rets mc') -> - call e fname t m args mc post2. - Proof. - unfold call. intros. fwd. - do 4 eexists. 1: eassumption. - do 2 eexists. 1: eassumption. - eapply exec.weaken. 1: eassumption. - cbv beta. clear -H0. intros. fwd. eauto. - Qed. - - Lemma extend_env_call: forall e1 e2, - map.extends e2 e1 -> - forall f t m rets mc post, - call e1 f t m rets mc post -> - call e2 f t m rets mc post. - Proof. - unfold call. intros. fwd. repeat eexists. - - eapply H. eassumption. - - eassumption. - - eapply exec.extend_env; eassumption. - Qed. -End WithParams. From 55d0f22b50fa00c579b09458406317ff66bfadd4 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Thu, 25 Jan 2024 02:08:36 -0500 Subject: [PATCH 19/62] Recover ipow metrics proof --- bedrock2/src/bedrock2/MetricLoops.v | 217 +++++---- bedrock2/src/bedrock2/MetricProgramLogic.v | 388 ++++++++++++++++ bedrock2/src/bedrock2/ProgramLogic.v | 10 +- bedrock2/src/bedrock2Examples/metric_ipow.v | 214 +++++++++ bedrock2/src/bedrock2Examples/tckmnipow.v | 469 -------------------- 5 files changed, 706 insertions(+), 592 deletions(-) create mode 100644 bedrock2/src/bedrock2/MetricProgramLogic.v create mode 100644 bedrock2/src/bedrock2Examples/metric_ipow.v delete mode 100644 bedrock2/src/bedrock2Examples/tckmnipow.v diff --git a/bedrock2/src/bedrock2/MetricLoops.v b/bedrock2/src/bedrock2/MetricLoops.v index d8333a8ce..723953561 100644 --- a/bedrock2/src/bedrock2/MetricLoops.v +++ b/bedrock2/src/bedrock2/MetricLoops.v @@ -6,6 +6,7 @@ Require Import coqutil.Map.Properties. Require Import coqutil.Tactics.destr. From bedrock2 Require Import Map.Separation Map.SeparationLogic. From bedrock2 Require Import Syntax MetricSemantics Markers. +From bedrock2 Require Semantics. From bedrock2 Require Import MetricWeakestPrecondition MetricWeakestPreconditionProperties. Require Import bedrock2.MetricLogging. @@ -13,109 +14,113 @@ Require Import bedrock2.MetricLogging. Section Loops. Context {width: Z} {BW: Bitwidth width} {word: word.word width} {mem: map.map word Byte.byte}. Context {locals: map.map String.string word}. - Context {ext_spec: ExtSpec}. + Context {ext_spec: Semantics.ExtSpec}. Context {word_ok : word.ok word} {mem_ok : map.ok mem}. Context {locals_ok : map.ok locals}. Context {ext_spec_ok : Semantics.ext_spec.ok ext_spec}. - Context {fs : env}. + Context {fs : Semantics.env}. Let call := fs. - Lemma wp_while: forall e c t m l (post: _ -> _ -> _ -> Prop), - (exists measure (lt:measure->measure->Prop) (inv:measure->trace->mem->locals->Prop), + Lemma wp_while: forall e c t m l mc (post: _ -> _ -> _ -> _ -> Prop), + (exists measure (lt:measure->measure->Prop) (inv:measure->Semantics.trace->mem->locals->MetricLog->Prop), Coq.Init.Wf.well_founded lt /\ - (exists v, inv v t m l) /\ - (forall v t m l, inv v t m l -> - exists b, dexpr m l e b /\ - (word.unsigned b <> 0%Z -> cmd call c t m l (fun t' m l => - exists v', inv v' t' m l /\ lt v' v)) /\ - (word.unsigned b = 0%Z -> post t m l))) -> - cmd call (cmd.while e c) t m l post. + (exists v, inv v t m l mc) /\ + (forall v t m l mc, inv v t m l mc -> + exists bv bmc, dexpr m l e mc (bv, bmc) /\ + (word.unsigned bv <> 0%Z -> cmd call c t m l bmc (fun t' m l mc => + exists v', inv v' t' m l (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) + /\ lt v' v)) /\ + (word.unsigned bv = 0%Z -> post t m l + (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 bmc)))))) -> + cmd call (cmd.while e c) t m l mc post. Proof. intros. destruct H as (measure & lt & inv & Hwf & HInit & Hbody). destruct HInit as (v0 & HInit). - revert t m l HInit. pattern v0. revert v0. + revert t m l mc HInit. pattern v0. revert v0. eapply (well_founded_ind Hwf). intros. - specialize Hbody with (1 := HInit). destruct Hbody as (b & Hb & Ht & Hf). - eapply expr_sound in Hb. destruct Hb as (b' & Hb & ?). subst b'. - destr.destr (Z.eqb (word.unsigned b) 0). - - specialize Hf with (1 := E). eapply exec.while_false; eassumption. + specialize Hbody with (1 := HInit) (mc := mc). destruct Hbody as (bv & bmc & Hb & Ht & Hf). + eapply expr_sound in Hb. destruct Hb as (bv' & bmc' & Hb & Heq). inversion Heq. subst bv' bmc'. + destr.destr (Z.eqb (word.unsigned bv) 0). + - specialize Hf with (1 := E). eapply exec.while_false; try eassumption. - specialize Ht with (1 := E). eapply sound_cmd in Ht. eapply exec.while_true; eauto. cbv beta. intros * (v' & HInv & HLt). eapply sound_cmd. eauto. Qed. Lemma tailrec_localsmap_1ghost - {e c t} {m: mem} {l} {post : trace -> mem -> locals -> Prop} + {e c t} {m: mem} {l} {mc} {post : Semantics.trace -> mem -> locals -> MetricLog -> Prop} {measure: Type} {Ghost: Type} - (P Q: measure -> Ghost -> trace -> mem -> locals -> Prop) + (P Q: measure -> Ghost -> Semantics.trace -> mem -> locals -> MetricLog -> Prop) (lt: measure -> measure -> Prop) (Hwf: well_founded lt) (v0: measure) (g0: Ghost) - (Hpre: P v0 g0 t m l) - (Hbody: forall v g t m l, - P v g t m l -> - exists br, expr m l e (eq br) /\ - (word.unsigned br <> 0%Z -> cmd call c t m l - (fun t' m' l' => exists v' g', - P v' g' t' m' l' /\ + (Hpre: P v0 g0 t m l mc) + (Hbody: forall v g t m l mc, + P v g t m l mc -> + exists brv brmc, expr m l e mc (eq (brv, brmc)) /\ + (word.unsigned brv <> 0%Z -> cmd call c t m l brmc + (fun t' m' l' mc' => exists v' g', + P v' g' t' m' l' (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc'))) /\ lt v' v /\ - (forall t'' m'' l'', Q v' g' t'' m'' l'' -> Q v g t'' m'' l''))) /\ - (word.unsigned br = 0%Z -> Q v g t m l)) - (Hpost: forall t m l, Q v0 g0 t m l -> post t m l) - : cmd call (cmd.while e c) t m l post. + (forall t'' m'' l'' mc'', Q v' g' t'' m'' l'' mc'' -> Q v g t'' m'' l'' mc''))) /\ + (word.unsigned brv = 0%Z -> Q v g t m l + (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 brmc))))) + (Hpost: forall t m l mc, Q v0 g0 t m l mc -> post t m l mc) + : cmd call (cmd.while e c) t m l mc post. Proof. eapply wp_while. - eexists measure, lt, (fun v t m l => - exists g, P v g t m l /\ forall t'' m'' l'', Q v g t'' m'' l'' -> Q v0 g0 t'' m'' l''). + eexists measure, lt, (fun v t m l mc => + exists g, P v g t m l mc /\ forall t'' m'' l'' mc'', Q v g t'' m'' l'' mc'' -> Q v0 g0 t'' m'' l'' mc''). split; [assumption|]. split; [solve[eauto]|]. - intros vi ti mi li (gi & HPi & HQimpl). - specialize (Hbody vi gi ti mi li HPi). - destruct Hbody as (br & ? & Hbody). exists br; split; [assumption|]. + intros vi ti mi li mci (gi & HPi & HQimpl). + specialize (Hbody vi gi ti mi li mci HPi). + destruct Hbody as (brv & brmc & ? & Hbody). exists brv, brmc; split; [assumption|]. destruct Hbody as (Htrue & Hfalse). split; intros Hbr; [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. { eapply Proper_cmd; [ |eapply Hpc]. - intros tj mj lj (vj& gj & HPj & Hlt & Qji); eauto 9. } + intros tj mj lj mcj (vj& gj & HPj & Hlt & Qji); eauto 9. } { eauto. } Qed. Lemma tailrec_localsmap_1ghost_parameterized_finalpost - {e c rest t} {m: mem} {l} + {e c rest t} {m: mem} {l mc} {measure: Type} {Ghost: Type} - (P Q: measure -> Ghost -> trace -> mem -> locals -> Prop) + (P Q: measure -> Ghost -> Semantics.trace -> mem -> locals -> MetricLog -> Prop) (lt: measure -> measure -> Prop) (Hwf: well_founded lt) (v0: measure) (g0: Ghost) - (Hpre: P v0 g0 t m l) - (Hbody: forall v g t m l, - P v g t m l -> - exists br, expr m l e (eq br) /\ - (word.unsigned br <> 0%Z -> cmd call c t m l - (fun t' m' l' => exists v' g', - P v' g' t' m' l' /\ + (Hpre: P v0 g0 t m l mc) + (Hbody: forall v g t m l mc, + P v g t m l mc -> + exists brv brmc, expr m l e mc (eq (brv, brmc)) /\ + (word.unsigned brv <> 0%Z -> cmd call c t m l brmc + (fun t' m' l' mc' => exists v' g', + P v' g' t' m' l' (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc'))) /\ lt v' v /\ - (forall t'' m'' l'', Q v' g' t'' m'' l'' -> Q v g t'' m'' l''))) /\ - (word.unsigned br = 0%Z -> cmd call rest t m l (Q v g))) - : cmd call (cmd.seq (cmd.while e c) rest) t m l (Q v0 g0). + (forall t'' m'' l'' mc'', Q v' g' t'' m'' l'' mc'' -> Q v g t'' m'' l'' mc''))) /\ + (word.unsigned brv = 0%Z -> cmd call rest t m l + (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 brmc))) (Q v g))) + : cmd call (cmd.seq (cmd.while e c) rest) t m l mc (Q v0 g0). Proof. cbn. eapply tailrec_localsmap_1ghost with - (Q := fun v g t m l => cmd call rest t m l (Q v g)). + (Q := fun v g t m l mc => cmd call rest t m l mc (Q v g)). 1: eassumption. 1: exact Hpre. 2: intros *; exact id. - intros vi gi ti mi li HPi. - specialize (Hbody vi gi ti mi li HPi). - destruct Hbody as (br & ? & Hbody). exists br; split; [assumption|]. + intros vi gi ti mi li mci HPi. + specialize (Hbody vi gi ti mi li mci HPi). + destruct Hbody as (brv & brmc & ? & Hbody). exists brv, brmc; split; [assumption|]. destruct Hbody as (Htrue & Hfalse). split; intros Hbr; [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. { eapply Proper_cmd; [ |eapply Hpc]. - intros tj mj lj (vj& gj & HPj & Hlt & Qji). do 2 eexists. + intros tj mj lj mcj (vj& gj & HPj & Hlt & Qji). do 2 eexists. split. 1: eassumption. split. 1: assumption. intros. eapply Proper_cmd. 2: eassumption. - intros tk mk lk HH. eapply Qji. assumption. } + intros tk mk lk mck HH. eapply Qji. assumption. } eapply Hpc. Qed. @@ -180,31 +185,25 @@ Section Loops. Lemma while_localsmap {e c t l mc} {m : mem} - {measure : Type} (invariant:_->_->_-> _ -> _ -> Prop) + {measure : Type} (invariant:_->_->_->_->_->Prop) {lt} (Hwf : well_founded lt) (v0 : measure) - {post : _->_->_-> _ -> Prop} + {post : _->_->_->_-> Prop} (Hpre : invariant v0 t m l mc) (Hbody : forall v t m l mc, invariant v t m l mc -> - exists br mcnew, expr m l e mc (eq (Datatypes.pair br mcnew)) /\ - (word.unsigned br <> 0 -> -<<<<<<< HEAD - cmd fs c t m l (fun t m l => exists v', invariant v' t m l /\ lt v' v)) /\ - (word.unsigned br = 0 -> post t m l)) - : cmd fs (cmd.while e c) t m l post. + exists brv brmc, expr m l e mc (eq (Datatypes.pair brv brmc)) /\ + (word.unsigned brv <> 0 -> + cmd fs c t m l brmc (fun t m l mc => exists v', + invariant v' t m l (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) + /\ lt v' v)) /\ + (word.unsigned brv = 0 -> post t m l + (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 brmc))))) + : cmd fs (cmd.while e c) t m l mc post. Proof. eapply wp_while. -======= - cmd call c t m l mcnew (fun t m l mcnew' => exists v', invariant v' t m l (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mcnew'))) /\ lt v' v)) /\ - (word.unsigned br = 0 -> post t m l (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mcnew))))) - : cmd call (cmd.while e c) t m l mc post. - Proof. - ->>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) eexists measure, lt, invariant. split. 1: exact Hwf. split. 1: eauto. - exact Hbody. Qed. @@ -220,14 +219,14 @@ Section Loops. (Hbody : forall v t m mc, tuple.foralls (fun localstuple => tuple.apply (invariant v t m mc) localstuple -> let l := reconstruct variables localstuple in - exists br mcnew, expr m l e mc (eq (Datatypes.pair br mcnew)) /\ - (word.unsigned br <> 0 -> - cmd call c t m l mcnew (fun t m l mcnew' => + exists brv brmc, expr m l e mc (eq (Datatypes.pair brv brmc)) /\ + (word.unsigned brv <> 0 -> + cmd call c t m l brmc (fun t m l mc => Markers.unique (Markers.left (tuple.existss (fun localstuple => enforce variables localstuple l /\ Markers.right (Markers.unique (exists v', - tuple.apply (invariant v' t m (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mcnew')))) localstuple /\ lt v' v))))))) /\ - (word.unsigned br = 0 -> post t m l (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mcnew)))))) + tuple.apply (invariant v' t m (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc)))) localstuple /\ lt v' v))))))) /\ + (word.unsigned brv = 0 -> post t m l (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 brmc)))))) : cmd call (cmd.while e c) t m l mc post. Proof. eapply (while_localsmap (fun v t m l mc => @@ -239,8 +238,8 @@ Section Loops. eapply hlist.foralls_forall in Hbody. specialize (Hbody Y). rewrite <-(reconstruct_enforce _ _ _ X) in Hbody. - destruct Hbody as (br & mcnew & Cond & Again & Done). - exists br. exists mcnew. split; [exact Cond|]. + destruct Hbody as (brv & brmc & Cond & Again & Done). + exists brv. exists brmc. split; [exact Cond|]. split; [|exact Done]. intro NE. specialize (Again NE). @@ -263,31 +262,27 @@ Section Loops. (v0 : measure) : hlist.foralls (fun (g0 : hlist ghosttypes) => forall (Hpre : (tuple.apply (hlist.apply (spec v0) g0 t m) l0 mc).(1)) - + (Hbody : forall v, hlist.foralls (fun g => forall t m mc, tuple.foralls (fun l => @dlet _ (fun _ => Prop) (reconstruct variables l) (fun localsmap : locals => match tuple.apply (hlist.apply (spec v) g t m) l mc with S_ => S_.(1) -> - Markers.unique (Markers.left (exists br mc', expr m localsmap e mc (eq (Datatypes.pair br mc')) /\ Markers.right ( - (word.unsigned br <> 0%Z -> cmd call c t m localsmap mc' - (fun t' m' localsmap' mc'' => + Markers.unique (Markers.left (exists brv brmc, expr m localsmap e mc (eq (Datatypes.pair brv brmc)) /\ Markers.right ( + (word.unsigned brv <> 0%Z -> cmd call c t m localsmap brmc + (fun t' m' localsmap' mc' => Markers.unique (Markers.left (hlist.existss (fun l' => enforce variables l' localsmap' /\ Markers.right ( Markers.unique (Markers.left (hlist.existss (fun g' => exists v', - match tuple.apply (hlist.apply (spec v') g' t' m') l' (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc''))) with S' => + match tuple.apply (hlist.apply (spec v') g' t' m') l' (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc'))) with S' => S'.(1) /\ Markers.right ( lt v' v /\ forall T M, hlist.foralls (fun L => forall MC, tuple.apply (S'.(2) T M) L MC -> tuple.apply (S_.(2) T M) L MC)) end))))))))) /\ - (word.unsigned br = 0%Z -> tuple.apply (S_.(2) t m) l (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mc')))))))end)))) + (word.unsigned brv = 0%Z -> tuple.apply (S_.(2) t m) l (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 brmc)))))))end)))) (Hpost : match (tuple.apply (hlist.apply (spec v0) g0 t m) l0 mc).(2) with Q0 => forall t m mc, hlist.foralls (fun l => tuple.apply (Q0 t m) l mc -> post t m (reconstruct variables l) mc)end) , cmd call (cmd.while e c) t m localsmap mc post ). Proof. eapply hlist_forall_foralls; intros g0 **. -<<<<<<< HEAD eapply wp_while. - eexists measure, lt, (fun vi ti mi localsmapi => -======= eexists measure, lt, (fun vi ti mi localsmapi mci => ->>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) exists gi li, localsmapi = reconstruct variables li /\ match tuple.apply (hlist.apply (spec vi) gi ti mi) li mci with S_ => S_.(1) /\ forall T M L MC, tuple.apply (S_.(2) T M) L MC-> @@ -296,19 +291,13 @@ Section Loops. split; [assumption|]. split. { exists v0, g0, l0. split. 1: eapply reconstruct_enforce; eassumption. split; eauto. } intros vi ti mi lmapi mci (gi&?&?&?&Qi); subst. - destruct (hlist.foralls_forall (hlist.foralls_forall (Hbody vi) gi ti mi mci) _ ltac:(eassumption)) as (br&?&?&X). + destruct (hlist.foralls_forall (hlist.foralls_forall (Hbody vi) gi ti mi mci) _ ltac:(eassumption)) as (brv&brmc&?&X). + - - exists br; exists x0; split; [assumption|]. destruct X as (Htrue&Hfalse). split; intros Hbr; + exists brv; exists brmc; split; [assumption|]. destruct X as (Htrue&Hfalse). split; intros Hbr; [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. -<<<<<<< HEAD { eapply Proper_cmd; [ |eapply Hpc]. - intros tj mj lmapj Hlj; eapply hlist.existss_exists in Hlj. -======= - { eapply Proper_cmd; [reflexivity..| | |eapply Hpc]. - { eapply Proper_call; firstorder idtac. } intros tj mj lmapj mcj Hlj; eapply hlist.existss_exists in Hlj. ->>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) destruct Hlj as (lj&Elj&HE); eapply reconstruct_enforce in Elj; subst lmapj. eapply hlist.existss_exists in HE. destruct HE as (l&?&?&?&HR). pose proof fun T M => hlist.foralls_forall (HR T M); clear HR. @@ -340,16 +329,10 @@ Section Loops. (Hpost : forall t m l mc, Q0 t m l mc -> post t m l mc) : cmd call (cmd.while e c) t m l mc post. Proof. -<<<<<<< HEAD eapply wp_while. - eexists measure, lt, (fun v t m l => - let S := spec v t m l in let '(P, Q) := S in - P /\ forall T M L, Q T M L -> Q0 T M L). -======= eexists measure, lt, (fun v t m l mc => let S := spec v t m l mc in let '(P, Q) := S in P /\ forall T M L MC, Q T M L MC -> Q0 T M L MC). ->>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) split; [assumption|]. cbv [Markers.split] in *. split; [solve[eauto]|]. @@ -357,14 +340,8 @@ Section Loops. destruct (Hbody _ _ _ _ _ ltac:(eassumption)) as (br&?&?&X); exists br; eexists; split; [eassumption|]. destruct X as (Htrue&Hfalse). split; intros Hbr; [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. -<<<<<<< HEAD { eapply Proper_cmd; [ |eapply Hpc]. - intros tj mj lj (vj&dP&?&dQ); eauto 9. } -======= - { eapply Proper_cmd; [reflexivity..| | |eapply Hpc]. - { eapply Proper_call; firstorder idtac. } intros tj mj lj mcj (vj&dP&?&dQ); eauto 9. } ->>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) { eauto. } Qed. @@ -382,7 +359,9 @@ Section Loops. constructor. intros [y|] pf; eauto. constructor. intros [] []. Qed. - + + (*METRICSTODO + Lemma atleastonce_localsmap {e c t} {m : mem} {l mc} {post : _->_->_->_-> Prop} {measure : Type} (invariant:_->_->_->_->_->Prop) lt @@ -429,7 +408,7 @@ Section Loops. split; intros; try contradiction; split; eauto. } { assert (Hnonzero := n). - apply Henterm in Hnonzero. + apply Henterm in Hnonzero. destruct Hnonzero as [v Hinv]. exists (Some v), br, mc'. @@ -442,7 +421,7 @@ Section Loops. eapply Proper_cmd; [ |eapply Hbody; eassumption]. ======= - exists v; split; trivial. + exists v; split; trivial. } } @@ -450,11 +429,11 @@ Section Loops. eexists; eexists; split. - eassumption. - split. - - + + (*eexists; eexists; split; [eassumption|]; split.*) - { + { intros Hc; destruct (Hcontinue Hc) as (v&?&Hinv); subst. eapply Proper_cmd. 1: { eapply Proper_call. } @@ -473,11 +452,11 @@ Section Loops. exists br'; split; trivial. split; intros; try contradiction. eexists; split; eauto. } } - + eapply Hexit. - Qed. + Qed. *) Admitted. - + Lemma atleastonce {e c t l mc} {m : mem} (variables : list String.string) @@ -514,7 +493,7 @@ Section Loops. intro Hnonzero. apply Henternonzero in Hnonzero. destruct Hnonzero as [v Hnonzero]. exists v. exists localstuple. repeat (split; eauto). } - { + { intros vi ti mi li mci (?&X&Y). specialize (Hbody vi ti mi). eapply hlist.foralls_forall in Hbody. @@ -529,6 +508,9 @@ Section Loops. eexists; split; eauto. } Qed. + + METRICSTODO*) + (* Lemma tailrec_earlyout_localsmap {e c t} {m : mem} {l mc} {post : _-> _->_->_-> Prop} @@ -729,4 +711,3 @@ Ltac loop_simpl := List.repeat Datatypes.length HList.polymorphic_list.repeat HList.polymorphic_list.length PrimitivePair.pair._1 PrimitivePair.pair._2] in *. - diff --git a/bedrock2/src/bedrock2/MetricProgramLogic.v b/bedrock2/src/bedrock2/MetricProgramLogic.v new file mode 100644 index 000000000..d04d937eb --- /dev/null +++ b/bedrock2/src/bedrock2/MetricProgramLogic.v @@ -0,0 +1,388 @@ +From coqutil.Tactics Require Import Tactics letexists eabstract rdelta reference_to_string ident_of_string. +Require Import coqutil.Map.Interface. +Require Import bedrock2.Syntax. +Require Import bedrock2.MetricWeakestPrecondition. +Require Import bedrock2.MetricWeakestPreconditionProperties. +Require Import bedrock2.MetricLoops. +Require Import bedrock2.Map.SeparationLogic bedrock2.Scalars. + +Definition spec_of (procname:String.string) := Semantics.env -> Prop. +Existing Class spec_of. + +Module Import Coercions. + Import Map.Interface Word.Interface BinInt. + Coercion Z.of_nat : nat >-> Z. + Coercion word.unsigned : word.rep >-> Z. + + Definition sepclause_of_map {key value map} (m : @map.rep key value map) + : map.rep -> Prop := Logic.eq m. + Coercion sepclause_of_map : Interface.map.rep >-> Funclass. +End Coercions. + +Goal True. + assert_succeeds epose (fun k v M (m : @Interface.map.rep k v M) => m _). +Abort. + +Section bindcmd. + Context {T : Type}. + Fixpoint bindcmd (c : Syntax.cmd) (k : Syntax.cmd -> T) {struct c} : T := + match c with + | cmd.cond e c1 c2 => bindcmd c1 (fun c1 => bindcmd c2 (fun c2 => let c := cmd.cond e c1 c2 in k c)) + | cmd.seq c1 c2 => bindcmd c1 (fun c1 => bindcmd c2 (fun c2 => let c := cmd.seq c1 c2 in k c)) + | cmd.while e c => bindcmd c (fun c => let c := cmd.while e c in k c) + | c => k c + end. +End bindcmd. + +(* TODO: use a deduplicating set instead of a list *) +Fixpoint callees (c : Syntax.cmd) : list String.string := + match c with + | cmd.cond _ c1 c2 | cmd.seq c1 c2 => callees c1 ++ callees c2 + | cmd.while _ c | cmd.stackalloc _ _ c => callees c + | cmd.call _ f _ => cons f nil + | _ => nil + end. + +Ltac assuming_correctness_of_in callees functions P := + lazymatch callees with + | nil => P + | cons ?f ?callees => + let f_spec := lazymatch constr:(_:spec_of f) with ?x => x end in + constr:(f_spec functions -> ltac:(let t := assuming_correctness_of_in callees functions P in exact t)) + end. +Require Import String List coqutil.Macros.ident_to_string. + +Ltac program_logic_goal_for_function proc := + let __ := constr:(proc : Syntax.func) in + constr_string_basename_of_constr_reference_cps ltac:(Tactics.head proc) ltac:(fun fname => + let spec := lazymatch constr:(_:spec_of fname) with ?s => s end in + exact (forall (functions : @map.rep _ _ Semantics.env) (EnvContains : map.get functions fname = Some proc), ltac:( + let callees := eval cbv in (callees (snd proc)) in + let s := assuming_correctness_of_in callees functions (spec functions) in + exact s))). +Definition program_logic_goal_for (_ : Syntax.func) (P : Prop) := P. + +Notation "program_logic_goal_for_function! proc" := (program_logic_goal_for proc ltac:( + program_logic_goal_for_function proc)) + (at level 10, only parsing). + +(* Users might want to override this with + Ltac normalize_body_of_function f ::= Tactics.rdelta.rdelta f. + in case cbv does more simplification than desired. *) +Ltac normalize_body_of_function f := eval cbv in f. + +Ltac bind_body_of_function f_ := + let f := normalize_body_of_function f_ in + let fargs := open_constr:(_) in + let frets := open_constr:(_) in + let fbody := open_constr:(_) in + let funif := open_constr:((fargs, frets, fbody)) in + unify f funif; + let G := lazymatch goal with |- ?G => G end in + let P := lazymatch eval pattern f_ in G with ?P _ => P end in + change (bindcmd fbody (fun c : Syntax.cmd => P (fargs, frets, c))); + cbv beta iota delta [bindcmd]; intros. + +(* note: f might have some implicit parameters (eg a record of constants) *) +Ltac enter f := + cbv beta delta [program_logic_goal_for]; + bind_body_of_function f; + lazymatch goal with |- ?s ?p => let s := rdelta s in change (s p); cbv beta end. + +Require coqutil.Map.SortedList. (* special-case eq_refl *) + +Ltac straightline_cleanup := + match goal with + (* TODO remove superfluous _ after .rep, but that will break some proofs that rely on + x not being cleared to instantiate evars with terms depending on x *) + | x : Word.Interface.word.rep _ |- _ => clear x + | x : Init.Byte.byte |- _ => clear x + | x : Semantics.trace |- _ => clear x + | x : Syntax.cmd |- _ => clear x + | x : Syntax.expr |- _ => clear x + | x : coqutil.Map.Interface.map.rep |- _ => clear x + | x : BinNums.Z |- _ => clear x + | x : unit |- _ => clear x + | x : bool |- _ => clear x + | x : list _ |- _ => clear x + | x : nat |- _ => clear x + (* same TODO as above *) + | x := _ : Word.Interface.word.rep _ |- _ => clear x + | x := _ : Init.Byte.byte |- _ => clear x + | x := _ : Semantics.trace |- _ => clear x + | x := _ : Syntax.cmd |- _ => clear x + | x := _ : Syntax.expr |- _ => clear x + | x := _ : coqutil.Map.Interface.map.rep |- _ => clear x + | x := _ : BinNums.Z |- _ => clear x + | x := _ : unit |- _ => clear x + | x := _ : bool |- _ => clear x + | x := _ : list _ |- _ => clear x + | x := _ : nat |- _ => clear x + | |- forall _, _ => intros + | |- let _ := _ in _ => intros + | |- dlet.dlet ?v (fun x => ?P) => change (let x := v in P); intros + | _ => progress (cbn [Semantics.interp_binop] in * ) + | H: exists _, _ |- _ => assert_succeeds progress destruct H as (_&_); destruct H + | H: _ /\ _ |- _ => destruct H + | x := ?y |- ?G => is_var y; subst x + | H: ?x = ?y |- _ => constr_eq x y; clear H + | H: ?x = ?y |- _ => is_var x; is_var y; assert_fails (idtac; let __ := eval cbv [x] in x in idtac); subst x + | H: ?x = ?y |- _ => is_var x; is_var y; assert_fails (idtac; let __ := eval cbv [y] in y in idtac); subst y + | H: ?x = ?v |- _ => + is_var x; + assert_fails (idtac; let __ := eval cbv delta [x] in x in idtac); + lazymatch v with context[x] => fail | _ => idtac end; + let x' := fresh x in + rename x into x'; + simple refine (let x := v in _); + change (x' = x) in H; + symmetry in H; + destruct H + end. + +Import MetricWeakestPrecondition. +Import coqutil.Map.Interface. + +Ltac straightline_stackalloc := + match goal with Hanybytes: Memory.anybytes ?a ?n ?mStack |- _ => + let m := match goal with H : map.split ?mCobined ?m mStack |- _ => m end in + let mCombined := match goal with H : map.split ?mCobined ?m mStack |- _ => mCobined end in + let Hsplit := match goal with H : map.split ?mCobined ?m mStack |- _ => H end in + let Hm := multimatch goal with H : _ m |- _ => H end in + let Hm' := fresh Hm in + let Htmp := fresh in + let Pm := match type of Hm with ?P m => P end in + assert_fails (assert (Separation.sep Pm (Array.array Separation.ptsto (Interface.word.of_Z (BinNums.Zpos BinNums.xH)) a _) mCombined) as _ by ecancel_assumption); + rename Hm into Hm'; + let stack := fresh "stack" in + let stack_length := fresh "length_" stack in (* MUST remain in context for deallocation *) + destruct (Array.anybytes_to_array_1 mStack a n Hanybytes) as (stack&Htmp&stack_length); + epose proof (ex_intro _ m (ex_intro _ mStack (conj Hsplit (conj Hm' Htmp))) + : Separation.sep _ (Array.array Separation.ptsto (Interface.word.of_Z (BinNums.Zpos BinNums.xH)) a _) mCombined) as Hm; + clear Htmp; (* note: we could clear more here if we assumed only one separation-logic description of each memory is present *) + try (let m' := fresh m in rename m into m'); rename mCombined into m; + ( assert (BinInt.Z.of_nat (Datatypes.length stack) = n) + by (rewrite stack_length; apply (ZifyInst.of_nat_to_nat_eq n)) + || fail 2 "negative stackalloc of size" n ) + end. + +Ltac straightline_stackdealloc := + lazymatch goal with |- exists _ _, Memory.anybytes ?a ?n _ /\ map.split ?m _ _ /\ _ => + let Hm := multimatch goal with Hm : _ m |- _ => Hm end in + let stack := match type of Hm with context [Array.array Separation.ptsto _ a ?stack] => stack end in + let length_stack := match goal with H : Datatypes.length stack = _ |- _ => H end in + let Hm' := fresh Hm in + pose proof Hm as Hm'; + let Psep := match type of Hm with ?P _ => P end in + let Htmp := fresh "Htmp" in + eassert (Lift1Prop.iff1 Psep (Separation.sep _ (Array.array Separation.ptsto (Interface.word.of_Z (BinNums.Zpos BinNums.xH)) a stack))) as Htmp + by ecancel || fail "failed to find stack frame in" Psep "using ecancel"; + eapply (fun m => proj1 (Htmp m)) in Hm; + let m' := fresh m in + rename m into m'; + let mStack := fresh in + destruct Hm as (m&mStack&Hsplit&Hm&Harray1); move Hm at bottom; + pose proof Array.array_1_to_anybytes _ _ _ Harray1 as Hanybytes; + rewrite length_stack in Hanybytes; + refine (ex_intro _ m (ex_intro _ mStack (conj Hanybytes (conj Hsplit _)))); + clear Htmp Hsplit mStack Harray1 Hanybytes + end. + +Ltac rename_to_different H := + idtac; + let G := fresh H "'0" in + rename H into G. +Ltac ensure_free H := + try rename_to_different H. + +Ltac eq_uniq_step := + match goal with + | |- ?x = ?y => + let x := rdelta x in + let y := rdelta y in + first [ is_evar x | is_evar y | constr_eq x y ]; exact eq_refl + | |- ?lhs = ?rhs => + let lh := head lhs in + is_constructor lh; + let rh := head rhs in + constr_eq lh rh; + f_equal (* NOTE: this is not sound, we really want just one f_equal application not a heuristic tactic *) + end. +Ltac eq_uniq := repeat eq_uniq_step. + +Ltac fwd_uniq_step := + match goal with + | |- exists x : ?T, _ => + let ev := open_constr:(match _ return T with x => x end) in + eexists ev; + let rec f := + tryif has_evar ev + then fwd_uniq_step + else idtac + in f + | |- _ /\ _ => split; [ solve [repeat fwd_uniq_step; eq_uniq] | ] + | _ => solve [ eq_uniq ] + end. +Ltac fwd_uniq := repeat fwd_uniq_step. + +Ltac straightline := + match goal with + | _ => straightline_cleanup + | |- program_logic_goal_for ?f _ => + enter f; intros; + match goal with + | H: map.get ?functions ?fname = Some _ |- _ => + eapply start_func; [exact H | clear H] + end; + cbv match beta delta [MetricWeakestPrecondition.func] + | |- MetricWeakestPrecondition.cmd _ (cmd.set ?s ?e) _ _ _ _ ?post => + unfold1_cmd_goal; cbv beta match delta [cmd_body]; + let __ := match s with String.String _ _ => idtac | String.EmptyString => idtac end in + ident_of_constr_string_cps s ltac:(fun x => + ensure_free x; + (* NOTE: keep this consistent with the [exists _, _ /\ _] case far below *) + letexists _ as x; split; [solve [repeat straightline]|]) + | |- cmd _ ?c _ _ _ _ ?post => + let c := eval hnf in c in + lazymatch c with + | cmd.while _ _ => fail + | cmd.cond _ _ _ => fail + | cmd.interact _ _ _ => fail + | _ => unfold1_cmd_goal; cbv beta match delta [cmd_body] + end + | |- @list_map _ _ (get _ _) _ _ _ => unfold1_list_map_goal; cbv beta match delta [list_map_body] + | |- @list_map _ _ (expr _ _) _ _ _ => unfold1_list_map_goal; cbv beta match delta [list_map_body] + | |- @list_map _ _ _ nil _ _ => cbv beta match fix delta [list_map list_map_body] + | |- expr _ _ _ _ _ => unfold1_expr_goal; cbv beta match delta [expr_body] + | |- dexpr _ _ _ _ _ => cbv beta delta [dexpr] + | |- dexprs _ _ _ _ _ => cbv beta delta [dexprs] + | |- literal _ _ _ => cbv beta delta [literal] + | |- @get ?w ?W ?L ?l ?x ?mc ?P => + let get' := eval cbv [get] in @get in + change (get' w W L l x mc P); cbv beta + | |- load _ _ _ _ _ => cbv beta delta [load] + | |- @MetricLoops.enforce ?width ?word ?locals ?names ?values ?map => + let values := eval cbv in values in + change (@MetricLoops.enforce width word locals names values map); + exact (conj (eq_refl values) eq_refl) + | |- @eq (@coqutil.Map.Interface.map.rep String.string Interface.word.rep _) _ _ => + eapply SortedList.eq_value; exact eq_refl + | |- @map.get String.string Interface.word.rep ?M ?m ?k = Some ?e' => + let e := rdelta e' in + is_evar e; + once (let v := multimatch goal with x := context[@map.put _ _ M _ k ?v] |- _ => v end in + (* cbv is slower than this, cbv with whitelist would have an enormous whitelist, cbv delta for map is slower than this, generalize unrelated then cbv is slower than this, generalize then vm_compute is slower than this, lazy is as slow as this: *) + unify e v; exact (eq_refl (Some v))) + | |- @coqutil.Map.Interface.map.get String.string Interface.word.rep _ _ _ = Some ?v => + let v' := rdelta v in is_evar v'; (change v with v'); exact eq_refl + | |- ?x = ?y => + let y := rdelta y in is_evar y; change (x=y); exact eq_refl + | |- ?x = ?y => + let x := rdelta x in is_evar x; change (x=y); exact eq_refl + | |- ?x = ?y => + let x := rdelta x in let y := rdelta y in constr_eq x y; exact eq_refl + | |- store Syntax.access_size.one _ _ _ _ => + eapply Scalars.store_one_of_sep; [solve[ecancel_assumption]|] + | |- store Syntax.access_size.two _ _ _ _ => + eapply Scalars.store_two_of_sep; [solve[ecancel_assumption]|] + | |- store Syntax.access_size.four _ _ _ _ => + eapply Scalars.store_four_of_sep; [solve[ecancel_assumption]|] + | |- store Syntax.access_size.word _ _ _ _ => + eapply Scalars.store_word_of_sep; [solve[ecancel_assumption]|] + | |- bedrock2.Memory.load Syntax.access_size.one ?m ?a = Some ?ev => + try subst ev; refine (@Scalars.load_one_of_sep _ _ _ _ _ _ _ _ _ _); ecancel_assumption + | |- @bedrock2.Memory.load _ ?word ?mem Syntax.access_size.two ?m ?a = Some ?ev => + try subst ev; refine (@Scalars.load_two_of_sep _ word _ mem _ a _ _ m _); ecancel_assumption + | |- @bedrock2.Memory.load _ ?word ?mem Syntax.access_size.four ?m ?a = Some ?ev => + try subst ev; refine (@Scalars.load_four_of_sep_32bit _ word _ mem _ eq_refl a _ _ m _); ecancel_assumption + | |- @bedrock2.Memory.load _ ?word ?mem Syntax.access_size.four ?m ?a = Some ?ev => + try subst ev; refine (@Scalars.load_four_of_sep _ word _ mem _ a _ _ m _); ecancel_assumption + | |- @bedrock2.Memory.load _ ?word ?mem Syntax.access_size.word ?m ?a = Some ?ev => + try subst ev; refine (@Scalars.load_word_of_sep _ word _ mem _ a _ _ m _); ecancel_assumption + | |- exists l', Interface.map.of_list_zip ?ks ?vs = Some l' /\ _ => + letexists; split; [exact eq_refl|] (* TODO: less unification here? *) + | |- exists l', Interface.map.putmany_of_list_zip ?ks ?vs ?l = Some l' /\ _ => + letexists; split; [exact eq_refl|] (* TODO: less unification here? *) + | _ => fwd_uniq_step + | |- exists x, ?P /\ ?Q => + let x := fresh x in refine (let x := _ in ex_intro (fun x => P /\ Q) x _); + split; [solve [repeat straightline]|] + (* NOTE: metrics only case; maybe try to unify with non-metrics? *) + | |- exists x y, ?P /\ ?Q => + eexists; eexists; split; [solve [repeat straightline]|] + (* eexists instead of letexists ensures unification of (?a,?b) = (const,const) + does not unfold the const aggressively (e.g. word to Naive) *) + | |- exists x, Markers.split (?P /\ ?Q) => + let x := fresh x in refine (let x := _ in ex_intro (fun x => P /\ Q) x _); + split; [solve [repeat straightline]|] + | |- Markers.unique (exists x, Markers.split (?P /\ ?Q)) => + let x := fresh x in refine (let x := _ in ex_intro (fun x => P /\ Q) x _); + split; [solve [repeat straightline]|] + | |- Markers.unique (Markers.left ?G) => + change G; + unshelve (idtac; repeat match goal with + | |- Markers.split (?P /\ Markers.right ?Q) => + split; [eabstract (repeat straightline) | change Q] + | |- exists _, _ => letexists + end); [] + | |- Markers.split ?G => change G; split + | |- True => exact I + | |- False \/ _ => right + | |- _ \/ False => left + (*| |- (_, _) = (_, _) => exact eq_refl (* NOTE: metrics-only case *)*) + | |- BinInt.Z.modulo ?z (Memory.bytes_per_word _) = BinInt.Z0 /\ _ => + lazymatch Coq.setoid_ring.InitialRing.isZcst z with + | true => split; [exact eq_refl|] + end + | |- _ => straightline_stackalloc + | |- _ => straightline_stackdealloc + | |- context[sep (sep _ _) _] => progress (flatten_seps_in_goal; cbn [seps]) + | H : context[sep (sep _ _) _] |- _ => progress (flatten_seps_in H; cbn [seps] in H) + end. + +(* TODO: once we can automatically prove some calls, include the success-only version of this in [straightline] *) +Ltac straightline_call := + lazymatch goal with + | |- MetricWeakestPrecondition.call ?functions ?callee _ _ _ _ => + let callee_spec := lazymatch constr:(_:spec_of callee) with ?s => s end in + let Hcall := lazymatch goal with H: callee_spec functions |- _ => H end in + eapply MetricWeakestPreconditionProperties.Proper_call; cycle -1; + [ eapply Hcall | try eabstract (solve [Morphisms.solve_proper]) .. ]; + [ .. | intros ? ? ? ?] + end. + +Ltac current_trace_mem_locals := + lazymatch goal with + | |- MetricWeakestPrecondition.cmd _ _ ?t ?m ?l _ => constr:((t, m, l)) + end. + +Ltac seprewrite Hrw := + let tml := current_trace_mem_locals in + let m := lazymatch tml with (_, ?m, _) => m end in + let H := multimatch goal with H: _ m |- _ => H end in + seprewrite_in Hrw H. +Ltac seprewrite_by Hrw tac := + let tml := current_trace_mem_locals in + let m := lazymatch tml with (_, ?m, _) => m end in + let H := multimatch goal with H: _ m |- _ => H end in + seprewrite_in_by Hrw H tac. + +Ltac show_program := + lazymatch goal with + | |- @cmd ?width ?BW ?word ?mem ?locals ?ext_spec ?E ?c ?F ?G ?H ?I => + let c' := eval cbv in c in + change (@cmd width BW word mem locals ext_spec E (fst (c, c')) F G H I) + end. + +Ltac subst_words := + repeat match goal with x := _ : coqutil.Word.Interface.word.rep |- _ => subst x end. + +Require Import coqutil.Tactics.eplace Coq.setoid_ring.Ring_tac. +Ltac ring_simplify_words := + subst_words; + repeat match goal with H : context [?w] |- _ => + let __ := constr:(w : Interface.word.rep) in + progress eplace w with _ in H by (ring_simplify; reflexivity) end; + repeat match goal with |- context [?w] => + let __ := constr:(w : Interface.word.rep) in + progress eplace w with _ by (ring_simplify; reflexivity) end. diff --git a/bedrock2/src/bedrock2/ProgramLogic.v b/bedrock2/src/bedrock2/ProgramLogic.v index 6bc3afdd6..2b6970016 100644 --- a/bedrock2/src/bedrock2/ProgramLogic.v +++ b/bedrock2/src/bedrock2/ProgramLogic.v @@ -235,14 +235,14 @@ Ltac straightline := eapply start_func; [exact H | clear H] end; cbv match beta delta [WeakestPrecondition.func] - | |- WeakestPrecondition.cmd _ (cmd.set ?s ?e) _ _ _ _ ?post => + | |- WeakestPrecondition.cmd _ (cmd.set ?s ?e) _ _ _ ?post => unfold1_cmd_goal; cbv beta match delta [cmd_body]; let __ := match s with String.String _ _ => idtac | String.EmptyString => idtac end in ident_of_constr_string_cps s ltac:(fun x => ensure_free x; (* NOTE: keep this consistent with the [exists _, _ /\ _] case far below *) letexists _ as x; split; [solve [repeat straightline]|]) - | |- cmd _ ?c _ _ _ _ ?post => + | |- cmd _ ?c _ _ _ ?post => let c := eval hnf in c in lazymatch c with | cmd.while _ _ => fail @@ -253,9 +253,9 @@ Ltac straightline := | |- @list_map _ _ (get _ _) _ _ => unfold1_list_map_goal; cbv beta match delta [list_map_body] | |- @list_map _ _ (expr _ _) _ _ => unfold1_list_map_goal; cbv beta match delta [list_map_body] | |- @list_map _ _ _ nil _ => cbv beta match fix delta [list_map list_map_body] - | |- expr _ _ _ _ _ => unfold1_expr_goal; cbv beta match delta [expr_body] - | |- dexpr _ _ _ _ _ => cbv beta delta [dexpr] - | |- dexprs _ _ _ _ _ => cbv beta delta [dexprs] + | |- expr _ _ _ _ => unfold1_expr_goal; cbv beta match delta [expr_body] + | |- dexpr _ _ _ _ => cbv beta delta [dexpr] + | |- dexprs _ _ _ _ => cbv beta delta [dexprs] | |- literal _ _ => cbv beta delta [literal] | |- @get ?w ?W ?L ?l ?x ?P => let get' := eval cbv [get] in @get in diff --git a/bedrock2/src/bedrock2Examples/metric_ipow.v b/bedrock2/src/bedrock2Examples/metric_ipow.v new file mode 100644 index 000000000..4b188ced5 --- /dev/null +++ b/bedrock2/src/bedrock2Examples/metric_ipow.v @@ -0,0 +1,214 @@ +Require Import Coq.ZArith.ZArith coqutil.Z.div_mod_to_equations. +Require Import bedrock2.NotationsCustomEntry. +Require Import bedrock2.MetricLogging. +Import Syntax BinInt String List.ListNotations ZArith. +Require Import coqutil.Z.Lia. +Local Open Scope string_scope. Local Open Scope Z_scope. Local Open Scope list_scope. + +Definition ipow := func! (x, e) ~> ret { + ret = $1; + while (e) { + if (e & $1) { ret = ret * x }; + e = e >> $1; + x = x * x + } +}. + +From bedrock2 Require Import BasicC64Semantics MetricWeakestPrecondition MetricProgramLogic. +From bedrock2 Require Import MetricLoops. +From coqutil Require Import Word.Properties Word.Interface Tactics.letexists. +Import Interface.word. + +Definition initCost := {| instructions := 9; stores := 0; loads := 9; jumps := 0 |}. +Definition iterCost := {| instructions := 38; stores := 0; loads := 45; jumps := 2 |}. +Definition endCost := {| instructions := 2; stores := 0; loads := 3; jumps := 1 |}. + +Definition msb z := match z with + | Zpos _ => Z.log2 z + 1 + | _ => 0 + end. + +#[export] Instance spec_of_ipow : spec_of "ipow" := + fnspec! "ipow" x e ~> v, + { requires t m mc := True; + ensures t' m' mc' := unsigned v = unsigned x ^ unsigned e mod 2^64 /\ + (mc' - mc <= initCost + (msb (word.unsigned e)) * iterCost + endCost)%metricsH + }. + +Module Z. + Lemma pow_mod x n m (Hnz: m <> 0) : (x mod m)^n mod m = x^n mod m. + Proof. + revert n. + eapply Z.order_induction_0; intros. + { intros ???; subst; split; auto. } + { rewrite 2Z.pow_0_r; trivial. } + { rewrite 2Z.pow_succ_r by trivial. + rewrite <-Z.mul_mod_idemp_r by trivial. + multimatch goal with H: _ |- _ => rewrite H end; + rewrite Z.mul_mod_idemp_l, Z.mul_mod_idemp_r; solve[trivial]. } + { rewrite 2Z.pow_neg_r; trivial. } + Qed. + + Lemma mod2_nonzero x : x mod 2 <> 0 -> x mod 2 = 1. + Proof. Z.div_mod_to_equations. blia. Qed. + + Lemma land_1_r x : Z.land x 1 = x mod 2. + Proof. + change 1 with (Z.ones 1) in *. + rewrite Z.land_ones in * by discriminate. + exact eq_refl. + Qed. +End Z. + +Require Import bedrock2.AbsintWordToZ coqutil.Z.Lia. + +Ltac t := + repeat match goal with x := _ |- _ => subst x end; + repeat match goal with |- context [word.unsigned ?e] => progress (idtac; let H := rbounded (word.unsigned e) in idtac) end; + repeat match goal with G: context [word.unsigned ?e] |- _ => progress (idtac; let H := rbounded (word.unsigned e) in idtac) end; + repeat match goal with |- context [word.unsigned ?e] => progress (idtac; let H := unsigned.zify_expr e in try rewrite H) end; + repeat match goal with G: context [word.unsigned ?e] |- _ => progress (idtac; let H := unsigned.zify_expr e in try rewrite H in G) end; + repeat match goal with H: absint_eq ?x ?x |- _ => clear H end; + cbv [absint_eq] in *. + +Lemma msb_shift z : 0 < z -> msb (z / 2) = msb z - 1. +Proof. + intro. + case (z / 2) eqn:Hdiv. + - enough (H1 : z = 1) by (rewrite H1; easy). + enough (z = z mod 2) by (Z.div_mod_to_equations; blia). + rewrite (Z.div_mod z 2) by blia. + rewrite Hdiv. + cbn. + rewrite Zmod_mod. + reflexivity. + - rewrite <- Z.div2_div in Hdiv. + rewrite (Zdiv2_odd_eqn z). + rewrite Hdiv. + rewrite <- Pos2Z.inj_mul. + case (Z.odd z); + [rewrite <- Pos2Z.inj_add | rewrite Z.add_0_r]; + unfold msb; + rewrite Z.add_simpl_r; + [rewrite Pos2Z.inj_add |]; rewrite Pos2Z.inj_mul; + [rewrite Z.log2_succ_double | rewrite Z.log2_double]; + blia. + - pose proof (Zlt_neg_0 p) as Hneg. + rewrite <- Hdiv in Hneg. + Z.div_mod_to_equations. + blia. +Qed. + +Lemma ipow_ok : program_logic_goal_for_function! ipow. +Proof. + repeat straightline. + match goal with H : True |- _ => clear H end. + + refine ((MetricLoops.tailrec + (* types of ghost variables*) HList.polymorphic_list.nil + (* program variables *) (["e";"ret";"x"] : list String.string)) + (fun v t m e ret x mc => PrimitivePair.pair.mk (v = word.unsigned e) (* precondition *) + (fun T M E RET X MC => T = t /\ M = m /\ (* postcondition *) + word.unsigned RET = word.unsigned ret * word.unsigned x ^ word.unsigned e mod 2^64 /\ + (MC - mc <= msb (word.unsigned e) * iterCost + endCost)%metricsH)) + (fun n m => 0 <= n < m) (* well_founded relation *) + _ _ _ _ _); + (* TODO wrap this into a tactic with the previous refine *) + cbn [HList.hlist.foralls HList.tuple.foralls + HList.hlist.existss HList.tuple.existss + HList.hlist.apply HList.tuple.apply + HList.hlist + List.repeat Datatypes.length + HList.polymorphic_list.repeat HList.polymorphic_list.length + PrimitivePair.pair._1 PrimitivePair.pair._2] in *. + + { repeat straightline. } + { exact (Z.lt_wf _). } + { repeat straightline. } (* init precondition *) + { (* loop test *) + repeat straightline; try show_program. + { (* loop body *) + eexists; eexists; split; [repeat straightline|]. (* if condition evaluation *) + split. (* if cases, path-blasting *) + { + repeat (straightline || (split; trivial; [])). 2: split. all:t. + { (* measure decreases *) + set (word.unsigned x0) in *. (* WHY does blia need this? *) + Z.div_mod_to_equations. blia. } + { (* invariant preserved *) + rewrite H3; clear H3. rename H0 into Hbit. + change (1+1) with 2 in *. + eapply Z.mod2_nonzero in Hbit. + epose proof (Z.div_mod _ 2 ltac:(discriminate)) as Heq; rewrite Hbit in Heq. + rewrite Heq at 2; clear Hbit Heq. + (* rewriting with equivalence modulo ... *) + rewrite !word.unsigned_mul. + unfold word.wrap. + rewrite ?Z.mul_mod_idemp_l by discriminate. + rewrite <-(Z.mul_mod_idemp_r _ (_^_)), Z.pow_mod by discriminate. + rewrite ?Z.pow_add_r by (pose proof word.unsigned_range x0; Z.div_mod_to_equations; blia). + rewrite ?Z.pow_twice_r, ?Z.pow_1_r, ?Z.pow_mul_l. + rewrite Z.mul_mod_idemp_r by discriminate. + f_equal; ring. } + { (* metrics correct *) + cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. + applyAddMetrics H4. + rewrite msb_shift in H4 by blia. + rewrite MetricArith.mul_sub_distr_r in H4. + rewrite <- MetricArith.add_sub_swap in H4. + rewrite <- MetricArith.le_add_le_sub_r in H4. + eapply MetricArith.le_trans. + 2: exact H4. + unfold iterCost. + solve_MetricLog. + } + } + { + repeat (straightline || (split; trivial; [])). 2: split. all: t. + { (* measure decreases *) + set (word.unsigned x0) in *. (* WHY does blia need this? *) + Z.div_mod_to_equations; blia. } + { (* invariant preserved *) + rewrite H3; clear H3. rename H0 into Hbit. + change (1+1) with 2 in *. + epose proof (Z.div_mod _ 2 ltac:(discriminate)) as Heq; rewrite Hbit in Heq. + rewrite Heq at 2; clear Hbit Heq. + (* rewriting with equivalence modulo ... *) + rewrite !word.unsigned_mul, ?Z.mul_mod_idemp_l by discriminate. + cbv [word.wrap]. + rewrite <-(Z.mul_mod_idemp_r _ (_^_)), Z.pow_mod by discriminate. + rewrite ?Z.add_0_r, Z.pow_twice_r, ?Z.pow_1_r, ?Z.pow_mul_l. + rewrite Z.mul_mod_idemp_r by discriminate. + f_equal; ring. } + { (* metrics correct *) + cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. + applyAddMetrics H4. + rewrite msb_shift in H4 by blia. + rewrite MetricArith.mul_sub_distr_r in H4. + rewrite <- MetricArith.add_sub_swap in H4. + rewrite <- MetricArith.le_add_le_sub_r in H4. + eapply MetricArith.le_trans. + 2: exact H4. + unfold iterCost. + solve_MetricLog. + } + } + } + { (* postcondition *) + rewrite H, Z.pow_0_r, Z.mul_1_r, word.wrap_unsigned. + repeat match goal with + | |- _ /\ _ => split + | |- _ = _ => reflexivity + end. + unfold msb, iterCost, endCost. + subst brmc. + solve_MetricLog. + } + } + + repeat straightline. + + repeat (split || letexists || t || trivial). + { setoid_rewrite H1; setoid_rewrite Z.mul_1_l; trivial. } + all: unfold initCost, iterCost, endCost in *; solve_MetricLog. +Qed. diff --git a/bedrock2/src/bedrock2Examples/tckmnipow.v b/bedrock2/src/bedrock2Examples/tckmnipow.v deleted file mode 100644 index b0a0780ab..000000000 --- a/bedrock2/src/bedrock2Examples/tckmnipow.v +++ /dev/null @@ -1,469 +0,0 @@ -Require Import Coq.ZArith.ZArith coqutil.Z.div_mod_to_equations. -Require Import bedrock2.NotationsCustomEntry. -Require Import bedrock2.MetricLogging. -Import Syntax BinInt String List.ListNotations ZArith. -Require Import coqutil.Z.Lia. -Local Open Scope string_scope. Local Open Scope Z_scope. Local Open Scope list_scope. -Local Coercion literal (z : Z) : Syntax.expr := Syntax.expr.literal z. - -Require Import bedrock2.Syntax. -Require Import bedrock2.WeakestPrecondition. -Require Import bedrock2.WeakestPreconditionProperties. -Require Import bedrock2.Loops. -Require Import bedrock2.Map.SeparationLogic. -Module Import Coercions. - Import Map.Interface Word.Interface BinInt. - Coercion Z.of_nat : nat >-> Z. - Coercion word.unsigned : word.rep >-> Z. -Require Import coqutil.Tactics.eplace coqutil.Tactics.eabstract Coq.setoid_ring.Ring_tac. -From coqutil.Tactics Require Import Tactics letexists eabstract rdelta reference_to_string ident_of_string. - - -Definition ipow := func! (x, e) ~> ret { - ret = $1; - while (e) { - if (e & $1) { ret = ret * x }; - e = e >> $1; - x = x * x - } - }. - -From bedrock2 Require Import Semantics BasicC64Semantics WeakestPrecondition ProgramLogic. -From coqutil Require Import Word.Properties Word.Interface Tactics.letexists. -Import Interface.word. - -Definition initCost := {| instructions := 9; stores := 0; loads := 9; jumps := 0 |}. -Definition iterCost := {| instructions := 38; stores := 0; loads := 45; jumps := 2 |}. -Definition endCost := {| instructions := 2; stores := 0; loads := 3; jumps := 1 |}. - -Definition msb z := match z with - | Zpos _ => Z.log2 z + 1 - | _ => 0 - end. - -#[export] Instance spec_of_ipow : spec_of "ipow" := - fnspec! "ipow" x e ~> v, - { requires t m mc := True; - ensures t' m' mc' := unsigned v = unsigned x ^ unsigned e mod 2^64 /\ - (mc' - mc <= initCost + (msb (word.unsigned e)) * iterCost + endCost)%metricsH}. - -Module Z. - Lemma pow_mod x n m (Hnz: m <> 0) : (x mod m)^n mod m = x^n mod m. - Proof. - revert n. - eapply Z.order_induction_0; intros. - { intros ???; subst; split; auto. } - { rewrite 2Z.pow_0_r; trivial. } - { rewrite 2Z.pow_succ_r by trivial. - rewrite <-Z.mul_mod_idemp_r by trivial. - multimatch goal with H: _ |- _ => rewrite H end; - rewrite Z.mul_mod_idemp_l, Z.mul_mod_idemp_r; solve[trivial]. } - { rewrite 2Z.pow_neg_r; trivial. } - Qed. - - Lemma mod2_nonzero x : x mod 2 <> 0 -> x mod 2 = 1. - Proof. Z.div_mod_to_equations. blia. Qed. - - Lemma land_1_r x : Z.land x 1 = x mod 2. - Proof. - change 1 with (Z.ones 1) in *. - rewrite Z.land_ones in * by discriminate. - exact eq_refl. - Qed. -End Z. - -Require Import bedrock2.AbsintWordToZ coqutil.Z.Lia. - -Ltac t := - repeat match goal with x := _ |- _ => subst x end; - repeat match goal with |- context [word.unsigned ?e] => progress (idtac; let H := rbounded (word.unsigned e) in idtac) end; - repeat match goal with G: context [word.unsigned ?e] |- _ => progress (idtac; let H := rbounded (word.unsigned e) in idtac) end; - repeat match goal with |- context [word.unsigned ?e] => progress (idtac; let H := unsigned.zify_expr e in try rewrite H) end; - repeat match goal with G: context [word.unsigned ?e] |- _ => progress (idtac; let H := unsigned.zify_expr e in try rewrite H in G) end; - repeat match goal with H: absint_eq ?x ?x |- _ => clear H end; - cbv [absint_eq] in *. - -Lemma msb_shift z : 0 < z -> msb (z / 2) = msb z - 1. -Proof. - intro. - case (z / 2) eqn:Hdiv. - - enough (H1 : z = 1) by (rewrite H1; easy). - enough (z = z mod 2) by (Z.div_mod_to_equations; blia). - rewrite (Z.div_mod z 2) by blia. - rewrite Hdiv. - cbn. - rewrite Zmod_mod. - reflexivity. - - rewrite <- Z.div2_div in Hdiv. - rewrite (Zdiv2_odd_eqn z). - rewrite Hdiv. - rewrite <- Pos2Z.inj_mul. - case (Z.odd z); - [rewrite <- Pos2Z.inj_add | rewrite Z.add_0_r]; - unfold msb; - rewrite Z.add_simpl_r; - [rewrite Pos2Z.inj_add |]; rewrite Pos2Z.inj_mul; - [rewrite Z.log2_succ_double | rewrite Z.log2_double]; - blia. - - pose proof (Zlt_neg_0 p) as Hneg. - rewrite <- Hdiv in Hneg. - Z.div_mod_to_equations. - blia. -Qed. - -Lemma ipow_ok : program_logic_goal_for_function! ipow. -Proof. - - straightline. - straightline. - straightline. - straightline. - - - (* first one works, second one doesnt? *) - (* - match goal with - | |- WeakestPrecondition.cmd _ (cmd.set ?s ?e) _ _ _ _ ?post => - unfold1_cmd_goal; cbv beta match delta [cmd_body]; - let __ := match s with String.String _ _ => idtac | String.EmptyString => idtac end in - ident_of_constr_string_cps s ltac:(fun x => - ensure_free x; hnf; - lazymatch goal with - | |- exists x', ?P => - refine (let x := _ in ex_intro (fun x' => P) x _) - end) - end. - *) - (* - match goal with - | |- WeakestPrecondition.cmd _ (cmd.set ?s ?e) _ _ _ _ ?post => - unfold1_cmd_goal; cbv beta match delta [cmd_body]; - let __ := match s with String.String _ _ => idtac | String.EmptyString => idtac end in - ident_of_constr_string_cps s ltac:(fun x => - ensure_free x; - letexists_as _ x) -end. -*) - - match goal with - | |- WeakestPrecondition.cmd _ (cmd.set ?s ?e) _ _ _ _ ?post => - unfold1_cmd_goal; cbv beta match delta [cmd_body]; - let __ := match s with String.String _ _ => idtac | String.EmptyString => idtac end in - ident_of_constr_string_cps s ltac:(fun x => - ensure_free x; hnf; - lazymatch goal with - | |- exists x', ?P => - refine (let x := _ in ex_intro (fun x' => P) x _) - end); - let __ := match s with String.String _ _ => idtac | String.EmptyString => idtac end in - ident_of_constr_string_cps s ltac:(fun x => - ensure_free x; hnf; - lazymatch goal with - | |- exists x', ?P => - refine (let x := _ in ex_intro (fun x' => P) x _) - end) - end. - split; try(repeat straightline; subst ret ret'0; cbv [literal dlet.dlet]; trivial). - - refine ((Loops.tailrec - (* types of ghost variables*) HList.polymorphic_list.nil - (* program variables *) (["e";"ret";"x"] : list String.string)) - (fun v t m e ret x mc => PrimitivePair.pair.mk (v = word.unsigned e) (* precondition *) - (fun T M E RET X MC=> T = t /\ M = m /\ (* postcondition *) - word.unsigned RET = word.unsigned ret * word.unsigned x ^ word.unsigned e mod 2^64 /\ - (MC - mc <= msb (word.unsigned e) * iterCost + endCost)%metricsH)) - (fun n m => 0 <= n < m) (* well_founded relation *) - _ _ _ _ _) ; - (* TODO wrap this into a tactic with the previous refine *) - cbn [HList.hlist.foralls HList.tuple.foralls - HList.hlist.existss HList.tuple.existss - HList.hlist.apply HList.tuple.apply - HList.hlist - List.repeat Datatypes.length - HList.polymorphic_list.repeat HList.polymorphic_list.length - PrimitivePair.pair._1 PrimitivePair.pair._2] in *. - - { repeat straightline. } - { exact (Z.lt_wf _). } - { repeat straightline. } (* init precondition *) - { (* loop test *) - repeat straightline; try show_program. - { (* loop body *) - eexists. exists (addMetricInstructions 11 (addMetricLoads 12 mc')). split; [repeat straightline|]. (* if condition evaluation *) - (*this case shouldnt exist*) 1:{ cbv [literal dlet.dlet]. - cbv [addMetricLoads withLoads addMetricInstructions withInstructions instructions stores loads jumps]. - simpl. - repeat ( - try rewrite applyAddInstructions; - try rewrite applyAddStores; - try rewrite applyAddLoads; - try rewrite applyAddJumps - ); - repeat rewrite <- Z.add_assoc; - cbn [Z.add Pos.add Pos.succ] -. trivial. - } - split. (* if cases, path-blasting *) - { - (straightline || (split; trivial; [])). - (straightline || (split; trivial; [])). - (*once again i am forced to do things i would rather not*) - unfold1_cmd_goal. - cbv beta match delta [cmd_body]. - exists (mul x1 x2). exists (addMetricInstructions 17 (addMetricLoads 20 (addMetricJumps 1 mc'))). - split. 1: { repeat straightline. - cbv [addMetricLoads withLoads addMetricInstructions withInstructions instructions stores loads jumps]. - simpl. - repeat ( - try rewrite applyAddInstructions; - try rewrite applyAddStores; - try rewrite applyAddLoads; - try rewrite applyAddJumps - ); - repeat rewrite <- Z.add_assoc; - cbn [Z.add Pos.add Pos.succ] -. trivial. - } - (straightline || (split; trivial; [])). - unfold1_cmd_goal. - cbv beta match delta [cmd_body]. - exists (sru x0 (word.of_Z 1)). exists (addMetricInstructions 29 (addMetricLoads 33 (addMetricJumps 1 mc'))). - split. 1: { repeat straightline. cbv [literal dlet.dlet]. - cbv [addMetricLoads withLoads addMetricInstructions withInstructions instructions stores loads jumps]. - simpl. - repeat ( - try rewrite applyAddInstructions; - try rewrite applyAddStores; - try rewrite applyAddLoads; - try rewrite applyAddJumps - ); - repeat rewrite <- Z.add_assoc; - cbn [Z.add Pos.add Pos.succ] -. rewrite Pos.add_carry_spec. - (*pretty sure this adds up, but i think i'd like to sort out the automation problems before pressing ahead any farther*) Abort. - } - (* :( *) - 2: split. all:t. - { (* measure decreases *) - set (word.unsigned x0) in *. (* WHY does blia need this? *) - Z.div_mod_to_equations. blia. } - { (* invariant preserved *) - rewrite H3; clear H3. rename H0 into Hbit. - change (1+1) with 2 in *. - eapply Z.mod2_nonzero in Hbit. - epose proof (Z.div_mod _ 2 ltac:(discriminate)) as Heq; rewrite Hbit in Heq. - rewrite Heq at 2; clear Hbit Heq. - (* rewriting with equivalence modulo ... *) - rewrite !word.unsigned_mul. - unfold word.wrap. - rewrite ?Z.mul_mod_idemp_l by discriminate. - rewrite <-(Z.mul_mod_idemp_r _ (_^_)), Z.pow_mod by discriminate. - rewrite ?Z.pow_add_r by (pose proof word.unsigned_range x0; Z.div_mod_to_equations; blia). - rewrite ?Z.pow_twice_r, ?Z.pow_1_r, ?Z.pow_mul_l. - rewrite Z.mul_mod_idemp_r by discriminate. - f_equal; ring. } - { (* metrics correct *) - cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. - applyAddMetrics H4. - rewrite msb_shift in H4 by blia. - rewrite MetricArith.mul_sub_distr_r in H4. - rewrite <- MetricArith.add_sub_swap in H4. - rewrite <- MetricArith.le_add_le_sub_r in H4. - eapply MetricArith.le_trans. - 2: exact H4. - unfold iterCost. - solve_MetricLog. - } - } - { - repeat (straightline || (split; trivial; [])). 2: split. all: t. - { (* measure decreases *) - set (word.unsigned x0) in *. (* WHY does blia need this? *) - Z.div_mod_to_equations; blia. } - { (* invariant preserved *) - rewrite H3; clear H3. rename H0 into Hbit. - change (1+1) with 2 in *. - epose proof (Z.div_mod _ 2 ltac:(discriminate)) as Heq; rewrite Hbit in Heq. - rewrite Heq at 2; clear Hbit Heq. - (* rewriting with equivalence modulo ... *) - rewrite !word.unsigned_mul, ?Z.mul_mod_idemp_l by discriminate. - cbv [word.wrap]. - rewrite <-(Z.mul_mod_idemp_r _ (_^_)), Z.pow_mod by discriminate. - rewrite ?Z.add_0_r, Z.pow_twice_r, ?Z.pow_1_r, ?Z.pow_mul_l. - rewrite Z.mul_mod_idemp_r by discriminate. - f_equal; ring. } - { (* metrics correct *) - cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. - applyAddMetrics H4. - rewrite msb_shift in H4 by blia. - rewrite MetricArith.mul_sub_distr_r in H4. - rewrite <- MetricArith.add_sub_swap in H4. - rewrite <- MetricArith.le_add_le_sub_r in H4. - eapply MetricArith.le_trans. - 2: exact H4. - unfold iterCost. - solve_MetricLog. - } - } - } - { (* postcondition *) - rewrite H, Z.pow_0_r, Z.mul_1_r, word.wrap_unsigned. - repeat match goal with - | |- _ /\ _ => split - | |- _ = _ => reflexivity - end. - unfold msb, iterCost, endCost. - subst mc'. - solve_MetricLog. - } - } - - repeat straightline. - repeat match goal with - | |- _ /\ _ => split - | |- _ = _ => reflexivity - | |- exists _, _ => letexists - | _ => t - end. - 1: setoid_rewrite H1; setoid_rewrite Z.mul_1_l; trivial. - - unfold initCost, iterCost, endCost in *. - solve_MetricLog. - - (*end hacking*) - - - - (* - Print straightline. - (let __ := match s with - | String.String _ _ => idtac - | "" => idtac - end in - ident_of_string.ident_of_constr_string_cps s - ltac:(fun x => ensure_free x; letexists _ as x; split; [ solve [ repeat straightline ] | ])). - *) - - straightline. - straightline. - repeat straightline. - cbv beta delta [dexpr]. - eexists; eexists; split. - - repeat straightline. trivial. unfold1_expr_goal; cbv beta match delta [expr_body]. straightline. - - Print straightline. - - Check ((Loops.tailrec - (* types of ghost variables*) HList.polymorphic_list.nil - (* program variables *) (["e";"ret";"x"] : list String.string)) - (fun v t m e ret x mc => PrimitivePair.pair.mk (v = word.unsigned e) (* precondition *) - (fun T M E RET X MC=> T = t /\ M = m /\ (* postcondition *) - word.unsigned RET = word.unsigned ret * word.unsigned x ^ word.unsigned e mod 2^64 /\ - (MC - mc <= msb (word.unsigned e) * iterCost + endCost)%metricsH)) - (fun n m => 0 <= n < m) (* well_founded relation *) - _ _ _ _ _). ; - (* TODO wrap this into a tactic with the previous refine *) - cbn [HList.hlist.foralls HList.tuple.foralls - HList.hlist.existss HList.tuple.existss - HList.hlist.apply HList.tuple.apply - HList.hlist - List.repeat Datatypes.length - HList.polymorphic_list.repeat HList.polymorphic_list.length - PrimitivePair.pair._1 PrimitivePair.pair._2] in *. - - { repeat straightline. } - { exact (Z.lt_wf _). } - { repeat straightline. } (* init precondition *) - { (* loop test *) - repeat straightline; try show_program. - { (* loop body *) - eexists; eexists; split; [repeat straightline|]. (* if condition evaluation *) - split. (* if cases, path-blasting *) - { - repeat (straightline || (split; trivial; [])). 2: split. all:t. - { (* measure decreases *) - set (word.unsigned x0) in *. (* WHY does blia need this? *) - Z.div_mod_to_equations. blia. } - { (* invariant preserved *) - rewrite H3; clear H3. rename H0 into Hbit. - change (1+1) with 2 in *. - eapply Z.mod2_nonzero in Hbit. - epose proof (Z.div_mod _ 2 ltac:(discriminate)) as Heq; rewrite Hbit in Heq. - rewrite Heq at 2; clear Hbit Heq. - (* rewriting with equivalence modulo ... *) - rewrite !word.unsigned_mul. - unfold word.wrap. - rewrite ?Z.mul_mod_idemp_l by discriminate. - rewrite <-(Z.mul_mod_idemp_r _ (_^_)), Z.pow_mod by discriminate. - rewrite ?Z.pow_add_r by (pose proof word.unsigned_range x0; Z.div_mod_to_equations; blia). - rewrite ?Z.pow_twice_r, ?Z.pow_1_r, ?Z.pow_mul_l. - rewrite Z.mul_mod_idemp_r by discriminate. - f_equal; ring. } - { (* metrics correct *) - cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. - applyAddMetrics H4. - rewrite msb_shift in H4 by blia. - rewrite MetricArith.mul_sub_distr_r in H4. - rewrite <- MetricArith.add_sub_swap in H4. - rewrite <- MetricArith.le_add_le_sub_r in H4. - eapply MetricArith.le_trans. - 2: exact H4. - unfold iterCost. - solve_MetricLog. - } - } - { - repeat (straightline || (split; trivial; [])). 2: split. all: t. - { (* measure decreases *) - set (word.unsigned x0) in *. (* WHY does blia need this? *) - Z.div_mod_to_equations; blia. } - { (* invariant preserved *) - rewrite H3; clear H3. rename H0 into Hbit. - change (1+1) with 2 in *. - epose proof (Z.div_mod _ 2 ltac:(discriminate)) as Heq; rewrite Hbit in Heq. - rewrite Heq at 2; clear Hbit Heq. - (* rewriting with equivalence modulo ... *) - rewrite !word.unsigned_mul, ?Z.mul_mod_idemp_l by discriminate. - cbv [word.wrap]. - rewrite <-(Z.mul_mod_idemp_r _ (_^_)), Z.pow_mod by discriminate. - rewrite ?Z.add_0_r, Z.pow_twice_r, ?Z.pow_1_r, ?Z.pow_mul_l. - rewrite Z.mul_mod_idemp_r by discriminate. - f_equal; ring. } - { (* metrics correct *) - cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. - applyAddMetrics H4. - rewrite msb_shift in H4 by blia. - rewrite MetricArith.mul_sub_distr_r in H4. - rewrite <- MetricArith.add_sub_swap in H4. - rewrite <- MetricArith.le_add_le_sub_r in H4. - eapply MetricArith.le_trans. - 2: exact H4. - unfold iterCost. - solve_MetricLog. - } - } - } - { (* postcondition *) - rewrite H, Z.pow_0_r, Z.mul_1_r, word.wrap_unsigned. - repeat match goal with - | |- _ /\ _ => split - | |- _ = _ => reflexivity - end. - unfold msb, iterCost, endCost. - subst mc'. - solve_MetricLog. - } - } - - repeat straightline. - repeat match goal with - | |- _ /\ _ => split - | |- _ = _ => reflexivity - | |- exists _, _ => letexists - | _ => t - end. - 1: setoid_rewrite H1; setoid_rewrite Z.mul_1_l; trivial. - - unfold initCost, iterCost, endCost in *. - solve_MetricLog. -Qed. From 6c4b51fe2563e9238b8bf01a39ab4f79de6d8c29 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Thu, 4 Apr 2024 05:16:25 -0400 Subject: [PATCH 20/62] extremely finnicky bugs in straightline --- bedrock2/src/bedrock2/MetricProgramLogic.v | 4 ++-- bedrock2/src/bedrock2/ProgramLogic.v | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/bedrock2/src/bedrock2/MetricProgramLogic.v b/bedrock2/src/bedrock2/MetricProgramLogic.v index d04d937eb..813ad255f 100644 --- a/bedrock2/src/bedrock2/MetricProgramLogic.v +++ b/bedrock2/src/bedrock2/MetricProgramLogic.v @@ -275,6 +275,7 @@ Ltac straightline := unify e v; exact (eq_refl (Some v))) | |- @coqutil.Map.Interface.map.get String.string Interface.word.rep _ _ _ = Some ?v => let v' := rdelta v in is_evar v'; (change v with v'); exact eq_refl + | |- (_, _) = (_, _) => f_equal (* NOTE: metrics-only case *) | |- ?x = ?y => let y := rdelta y in is_evar y; change (x=y); exact eq_refl | |- ?x = ?y => @@ -329,7 +330,6 @@ Ltac straightline := | |- True => exact I | |- False \/ _ => right | |- _ \/ False => left - (*| |- (_, _) = (_, _) => exact eq_refl (* NOTE: metrics-only case *)*) | |- BinInt.Z.modulo ?z (Memory.bytes_per_word _) = BinInt.Z0 /\ _ => lazymatch Coq.setoid_ring.InitialRing.isZcst z with | true => split; [exact eq_refl|] @@ -353,7 +353,7 @@ Ltac straightline_call := Ltac current_trace_mem_locals := lazymatch goal with - | |- MetricWeakestPrecondition.cmd _ _ ?t ?m ?l _ => constr:((t, m, l)) + | |- MetricWeakestPrecondition.cmd _ _ ?t ?m ?l _ _ => constr:((t, m, l)) end. Ltac seprewrite Hrw := diff --git a/bedrock2/src/bedrock2/ProgramLogic.v b/bedrock2/src/bedrock2/ProgramLogic.v index 2b6970016..5bc8e40b5 100644 --- a/bedrock2/src/bedrock2/ProgramLogic.v +++ b/bedrock2/src/bedrock2/ProgramLogic.v @@ -250,7 +250,7 @@ Ltac straightline := | cmd.interact _ _ _ => fail | _ => unfold1_cmd_goal; cbv beta match delta [cmd_body] end - | |- @list_map _ _ (get _ _) _ _ => unfold1_list_map_goal; cbv beta match delta [list_map_body] + | |- @list_map _ _ (get _) _ _ => unfold1_list_map_goal; cbv beta match delta [list_map_body] | |- @list_map _ _ (expr _ _) _ _ => unfold1_list_map_goal; cbv beta match delta [list_map_body] | |- @list_map _ _ _ nil _ => cbv beta match fix delta [list_map list_map_body] | |- expr _ _ _ _ => unfold1_expr_goal; cbv beta match delta [expr_body] From 387f09c48f8e864a29e965f354074774c0a755e9 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Thu, 4 Apr 2024 05:16:40 -0400 Subject: [PATCH 21/62] fix bsearch proof --- bedrock2/src/bedrock2Examples/bsearch.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/bedrock2/src/bedrock2Examples/bsearch.v b/bedrock2/src/bedrock2Examples/bsearch.v index 2e452db74..32b5ffa52 100644 --- a/bedrock2/src/bedrock2Examples/bsearch.v +++ b/bedrock2/src/bedrock2Examples/bsearch.v @@ -112,7 +112,7 @@ Proof. repeat apply conj; auto; []. (* postcondition *) letexists. split. { exact eq_refl. } - { auto. } + { repeat straightline. } Unshelve. all: exact (word.of_Z 0). From 0fff40067d74bb4568b203496b71461e84da5158 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Thu, 4 Apr 2024 05:17:48 -0400 Subject: [PATCH 22/62] simplify ipow metrics proof --- bedrock2/src/bedrock2Examples/metric_ipow.v | 33 ++++++++++----------- 1 file changed, 15 insertions(+), 18 deletions(-) diff --git a/bedrock2/src/bedrock2Examples/metric_ipow.v b/bedrock2/src/bedrock2Examples/metric_ipow.v index 4b188ced5..7312585f1 100644 --- a/bedrock2/src/bedrock2Examples/metric_ipow.v +++ b/bedrock2/src/bedrock2Examples/metric_ipow.v @@ -151,16 +151,21 @@ Proof. rewrite Z.mul_mod_idemp_r by discriminate. f_equal; ring. } { (* metrics correct *) - cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. + rewrite <- (MetricLog_eq mc0) in H4. applyAddMetrics H4. rewrite msb_shift in H4 by blia. - rewrite MetricArith.mul_sub_distr_r in H4. - rewrite <- MetricArith.add_sub_swap in H4. - rewrite <- MetricArith.le_add_le_sub_r in H4. - eapply MetricArith.le_trans. - 2: exact H4. - unfold iterCost. + unfold iterCost in *. solve_MetricLog. + (*destruct mc0.*) + (*applyAddMetrics H4.*) + (*rewrite msb_shift in H4 by blia.*) + (*rewrite MetricArith.mul_sub_distr_r in H4.*) + (*rewrite <- MetricArith.add_sub_swap in H4.*) + (*rewrite <- MetricArith.le_add_le_sub_r in H4.*) + (*eapply MetricArith.le_trans.*) + (*2: exact H4.*) + (*unfold iterCost.*) + (*solve_MetricLog.*) } } { @@ -181,25 +186,17 @@ Proof. rewrite Z.mul_mod_idemp_r by discriminate. f_equal; ring. } { (* metrics correct *) - cbn [addMetricLoads withLoads instructions stores loads jumps] in H4. + rewrite <- (MetricLog_eq mc0) in H4. applyAddMetrics H4. rewrite msb_shift in H4 by blia. - rewrite MetricArith.mul_sub_distr_r in H4. - rewrite <- MetricArith.add_sub_swap in H4. - rewrite <- MetricArith.le_add_le_sub_r in H4. - eapply MetricArith.le_trans. - 2: exact H4. - unfold iterCost. + unfold iterCost in *. solve_MetricLog. } } } { (* postcondition *) rewrite H, Z.pow_0_r, Z.mul_1_r, word.wrap_unsigned. - repeat match goal with - | |- _ /\ _ => split - | |- _ = _ => reflexivity - end. + split; [reflexivity|]. unfold msb, iterCost, endCost. subst brmc. solve_MetricLog. From b16c3a6cae51c70eefeb31cb34ad50802c3d5d16 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Thu, 25 Apr 2024 00:11:59 -0400 Subject: [PATCH 23/62] fix cost_SOp side effects: - breaks useimmediate - fixes spilling --- compiler/src/compiler/FlatImp.v | 6 ++---- compiler/src/compiler/FlattenExpr.v | 2 +- compiler/src/compiler/RegAlloc.v | 28 ++++++++++++++-------------- compiler/src/compiler/Spilling.v | 24 ++++++------------------ compiler/src/compiler/UseImmediate.v | 13 +++++++------ 5 files changed, 30 insertions(+), 43 deletions(-) diff --git a/compiler/src/compiler/FlatImp.v b/compiler/src/compiler/FlatImp.v index 20478e475..ad4924f01 100644 --- a/compiler/src/compiler/FlatImp.v +++ b/compiler/src/compiler/FlatImp.v @@ -335,7 +335,7 @@ Module exec. end. Definition cost_SOp x y z mc := - match (isReg x, isReg y, isReg z) with + match (isReg x, isReg y, (match z with | Var vo => isReg vo | Const _ => false end)) with | (false, false, false) => (addMetricInstructions 5 (addMetricLoads 7 (addMetricStores 1 mc))) | (false, false, true) | (false, true, false) => (addMetricInstructions 4 (addMetricLoads 5 (addMetricStores 1 mc))) | (false, true, true) => (addMetricInstructions 3 (addMetricLoads 3 (addMetricStores 1 mc))) @@ -466,9 +466,7 @@ Module exec. | op: forall t m l mc x op y y' z z' post, map.get l y = Some y' -> lookup_op_locals l z = Some z' -> - post t m (map.put l x (interp_binop op y' z')) - (addMetricLoads 2 - (addMetricInstructions 2 mc)) -> + post t m (map.put l x (interp_binop op y' z')) (cost_SOp x y z mc) -> exec (SOp x op y z) t m l mc post | set: forall t m l mc x y y' post, map.get l y = Some y' -> diff --git a/compiler/src/compiler/FlattenExpr.v b/compiler/src/compiler/FlattenExpr.v index 8844f8295..ad809d6f2 100644 --- a/compiler/src/compiler/FlattenExpr.v +++ b/compiler/src/compiler/FlattenExpr.v @@ -421,7 +421,7 @@ Section FlattenExpr1. clear IHe1 IHe2. pose_flatten_var_ineqs. set_solver. * intros. simpl in *. simp. clear IHe1 IHe2. - eapply @FlatImp.exec.op; t_safe; t_safe. 2 : solve_MetricLog. + eapply @FlatImp.exec.op; t_safe; t_safe. 2: try solve_MetricLog; admit. eapply flattenExpr_valid_resVar in E1; simpl; maps. - (* expr.ite *) diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index bf1d275f5..a7d21cb05 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -1196,19 +1196,19 @@ Section CheckerCorrect. unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; try blia. Qed. - Lemma check_regs_cost_SOp: forall x x' y y' z z' mc mc', - check_regs x x' = true -> - check_regs y y' = true -> - check_regs z z' = true -> - (MetricLogging.metricsLeq - (MetricLogging.metricsSub (exec.cost_SOp isRegZ x' y' z' mc') mc') - (MetricLogging.metricsSub (exec.cost_SOp isRegStr x y z mc) mc)). - Proof. - intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLit in *; - destr (isRegStr x); destr (isRegZ x'); destr (isRegStr y); destr (isRegZ y'); destr (isRegStr z); destr (isRegZ z'); - try discriminate; - unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; try blia. - Qed. + (*Lemma check_regs_cost_SOp: forall x x' y y' z z' mc mc',*) + (* check_regs x x' = true ->*) + (* check_regs y y' = true ->*) + (* check_regs z z' = true ->*) + (* (MetricLogging.metricsLeq*) + (* (MetricLogging.metricsSub (exec.cost_SOp isRegZ x' y' z' mc') mc')*) + (* (MetricLogging.metricsSub (exec.cost_SOp isRegStr x y z mc) mc)).*) + (*Proof.*) + (* intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLit in *;*) + (* destr (isRegStr x); destr (isRegZ x'); destr (isRegStr y); destr (isRegZ y'); destr (isRegStr z); destr (isRegZ z');*) + (* try discriminate;*) + (* unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; try blia.*) + (*Qed.*) Lemma check_regs_cost_SSet: forall x x' y y' mc mc', check_regs x x' = true -> @@ -1284,7 +1284,7 @@ Section CheckerCorrect. (* Hint Resolve states_compat_then_op : checker_hints.*) Hint Resolve check_regs_cost_SLoad check_regs_cost_SStore check_regs_cost_SInlinetable - check_regs_cost_SStackalloc check_regs_cost_SLit check_regs_cost_SOp + check_regs_cost_SStackalloc check_regs_cost_SLit (*check_regs_cost_SOp*) check_regs_cost_SSet check_regs_cost_SIf check_regs_cost_SLoop_false : checker_hints. diff --git a/compiler/src/compiler/Spilling.v b/compiler/src/compiler/Spilling.v index 10a828a66..993e942ed 100644 --- a/compiler/src/compiler/Spilling.v +++ b/compiler/src/compiler/Spilling.v @@ -1815,7 +1815,6 @@ Section Spilling. + blia. + unfold exec.cost_SLit; repeat isReg_helper. destr (isRegZ x); solve_MetricLog. - (* exec.op *) - unfold exec.lookup_op_locals in *. eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). clear H3. intros. destruct_one_match; fwd. @@ -1825,29 +1824,18 @@ Section Spilling. eapply exec.op. { eapply get_iarg_reg_1; eauto with zarith. } { unfold exec.lookup_op_locals in *. apply map.get_put_same. } - { eapply save_ires_reg_correct; (try eassumption || blia). admit. } + { eapply save_ires_reg_correct; (try eassumption || blia). + unfold exec.cost_SOp; repeat isReg_helper. + destr (isRegZ v); destr (isRegZ x); destr (isRegZ y); solve_MetricLog. } } { eapply exec.seq_cps. eapply exec.op. { apply map.get_put_same. } { unfold exec.lookup_op_locals in *. reflexivity. } - { eapply save_ires_reg_correct; (try eassumption || blia). admit. } + { eapply save_ires_reg_correct; (try eassumption || blia). + unfold exec.cost_SOp; repeat isReg_helper. + destr (isRegZ x); destr (isRegZ y); solve_MetricLog. } } - - (* TODO pratap's proof script may be useful for reconciling the two admits *) - (*eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros.*) - (*eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros.*) - (*eapply exec.seq_cps.*) - (*eapply exec.op.*) - (*1: eapply get_iarg_reg_1; eauto with zarith.*) - (*1: apply map.get_put_same.*) - (*eapply save_ires_reg_correct.*) - (*+ eassumption.*) - (*+ eassumption.*) - (*+ blia.*) - (*+ unfold exec.cost_SOp; repeat isReg_helper.*) - (* destr (isRegZ z); destr (isRegZ x); destr (isRegZ y); solve_MetricLog.*) - - (* exec.set *) eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.seq_cps. diff --git a/compiler/src/compiler/UseImmediate.v b/compiler/src/compiler/UseImmediate.v index b58016caa..ceff40333 100644 --- a/compiler/src/compiler/UseImmediate.v +++ b/compiler/src/compiler/UseImmediate.v @@ -85,11 +85,12 @@ Section WithArguments. all: eapply @exec.op; simpl in *; [ eassumption | reflexivity | try eassumption ]. - { rewrite word.add_comm. assumption. } - { replace (word.add y' (word.of_Z (- v))) with (word.sub y' (word.of_Z v)) by ring. assumption. } - { rewrite word.and_comm. assumption. } - { rewrite word.or_comm. assumption. } - { rewrite word.xor_comm. assumption. } + all: admit. + (*{ rewrite word.add_comm. assumption. }*) + (*{ replace (word.add y' (word.of_Z (- v))) with (word.sub y' (word.of_Z v)) by ring. assumption. }*) + (*{ rewrite word.and_comm. assumption. }*) + (*{ rewrite word.or_comm. assumption. }*) + (*{ rewrite word.xor_comm. assumption. }*) } - Qed. + Admitted. End WithArguments. From b433c2f439d16771b9eb56479c27b314eb52a8e6 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Fri, 7 Jun 2024 04:11:05 -0400 Subject: [PATCH 24/62] utility functions it seems i am trying to do too many things at once so i am committing the shared things so that i do not have to worry about missing dependencies due to commit order --- bedrock2/src/bedrock2/MetricLogging.v | 3 +++ compiler/src/compiler/MetricsToRiscv.v | 7 +++++++ 2 files changed, 10 insertions(+) diff --git a/bedrock2/src/bedrock2/MetricLogging.v b/bedrock2/src/bedrock2/MetricLogging.v index 29644b06b..4706bf935 100644 --- a/bedrock2/src/bedrock2/MetricLogging.v +++ b/bedrock2/src/bedrock2/MetricLogging.v @@ -134,6 +134,9 @@ Module MetricArith. Lemma le_trans : forall n m p, n <= m -> m <= p -> n <= p. Proof. solve_MetricLog. Qed. + Lemma le_refl : forall m, m <= m. + Proof. solve_MetricLog. Qed. + End MetricArith. Lemma applyAddInstructions n a b c d : addMetricInstructions n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a+n; stores := b; loads := c; jumps := d |}. Proof. auto. Qed. diff --git a/compiler/src/compiler/MetricsToRiscv.v b/compiler/src/compiler/MetricsToRiscv.v index ac5423442..3f57d77ba 100644 --- a/compiler/src/compiler/MetricsToRiscv.v +++ b/compiler/src/compiler/MetricsToRiscv.v @@ -23,6 +23,13 @@ Section MetricsToRiscv. storesL := storesH mh; |}. + Definition raiseMetrics (ml: metricsL): metricsH := {| + instructionsH := instructionsL ml; + jumpsH := jumpsL ml; + loadsH := loadsL ml; + storesH := storesL ml; + |}. + End MetricsToRiscv. Ltac solve_MetricLog := From bcd6714f38978706dd22a57675448488c55a01ae Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Sun, 9 Jun 2024 22:54:41 -0400 Subject: [PATCH 25/62] Parametrize FlatImp.exec by compiler phase The spilling phase adds some boilerplate inside functions and around function calls, so we need to account for it by making things artificially more expensive before the spilling phase. In an effort to avoid monolithic commits, this is a partial commit, so the code definitely doesn't compile in the current state since lots of things changed as a side effect. There are a few comments that say "cost_something constraint:", which document the various constraints on the magic numbers chosen as upper bounds for the costs (100, 95, 50). These should probably change to better numbers at some point. --- bedrock2/src/bedrock2/MetricLogging.v | 19 ++ compiler/src/compiler/FlatImp.v | 49 +++-- compiler/src/compiler/FlatToRiscvCommon.v | 2 +- compiler/src/compiler/FlatToRiscvFunctions.v | 15 +- compiler/src/compiler/FlattenExpr.v | 6 +- compiler/src/compiler/Pipeline.v | 15 +- compiler/src/compiler/RegAlloc.v | 4 +- compiler/src/compiler/Spilling.v | 185 ++++++++++++++++--- 8 files changed, 235 insertions(+), 60 deletions(-) diff --git a/bedrock2/src/bedrock2/MetricLogging.v b/bedrock2/src/bedrock2/MetricLogging.v index 4706bf935..2468f050a 100644 --- a/bedrock2/src/bedrock2/MetricLogging.v +++ b/bedrock2/src/bedrock2/MetricLogging.v @@ -118,6 +118,11 @@ Ltac solve_MetricLog := repeat simpl_MetricLog; blia. +Ltac solve_MetricLog_piecewise := + repeat unfold_MetricLog; + repeat simpl_MetricLog; + f_equal; blia. + Module MetricArith. Open Scope MetricH_scope. @@ -137,8 +142,22 @@ Module MetricArith. Lemma le_refl : forall m, m <= m. Proof. solve_MetricLog. Qed. + Lemma le_sub_mono : forall n m p, n - p <= m - p <-> n <= m. + Proof. solve_MetricLog. Qed. + + Lemma add_0_r : forall mc, (mc + EmptyMetricLog)%metricsH = mc. + Proof. destruct mc. unfold EmptyMetricLog. solve_MetricLog_piecewise. Qed. + + Lemma sub_0_r : forall mc, (mc - EmptyMetricLog)%metricsH = mc. + Proof. destruct mc. unfold EmptyMetricLog. solve_MetricLog_piecewise. Qed. + End MetricArith. +Create HintDb metric_arith. +#[export] Hint Resolve MetricArith.le_trans MetricArith.le_refl MetricArith.add_0_r MetricArith.sub_0_r : metric_arith. +#[export] Hint Resolve <- MetricArith.le_sub_mono : metric_arith. +#[export] Hint Resolve -> MetricArith.le_sub_mono : metric_arith. + Lemma applyAddInstructions n a b c d : addMetricInstructions n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a+n; stores := b; loads := c; jumps := d |}. Proof. auto. Qed. Lemma applyAddStores n a b c d : addMetricStores n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a; stores := b+n; loads := c; jumps := d |}. Proof. auto. Qed. Lemma applyAddLoads n a b c d : addMetricLoads n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a; stores := b; loads := c+n; jumps := d |}. Proof. auto. Qed. diff --git a/compiler/src/compiler/FlatImp.v b/compiler/src/compiler/FlatImp.v index ad4924f01..f803a4910 100644 --- a/compiler/src/compiler/FlatImp.v +++ b/compiler/src/compiler/FlatImp.v @@ -26,6 +26,10 @@ Inductive bbinop: Type := | BLtu | BGeu. +Inductive compphase: Type := +| PreSpill +| PostSpill. + Section Syntax. Context {varname: Type}. @@ -282,6 +286,7 @@ Module exec. {env_ok: map.ok env} {ext_spec_ok: ext_spec.ok ext_spec}. + Variable (phase: compphase). Variable (isReg: varname -> bool). Variable (e: env). @@ -298,6 +303,26 @@ Module exec. end. (* Helper functions for computing costs of instructions *) + Definition cost_SInteract mc := + match phase with + | PreSpill => (addMetricInstructions 100 (addMetricJumps 100 (addMetricStores 100 (addMetricLoads 100 mc)))) + | PostSpill => (addMetricInstructions 50 (addMetricJumps 50 (addMetricStores 50 (addMetricLoads 50 mc)))) + end. + + Definition cost_SCall_internal mc := + match phase with + | PreSpill => addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc))) + | PostSpill => addMetricInstructions 50 (addMetricJumps 50 (addMetricLoads 50 (addMetricStores 50 mc))) + end. + + Definition cost_SCall_external mc := + match phase with + | PreSpill => addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc))) + | PostSpill => addMetricInstructions 50 (addMetricJumps 50 (addMetricLoads 50 (addMetricStores 50 mc))) + end. + + (* TODO think about a non-fixed bound on the cost of function preamble and postamble *) + Definition cost_SLoad x a mc := match (isReg x, isReg a) with | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricStores 1 mc))) @@ -412,24 +437,20 @@ Module exec. exists l', map.putmany_of_list_zip resvars resvals l = Some l' /\ forall m', map.split m' mKeep mReceive -> post (((mGive, action, argvals), (mReceive, resvals)) :: t) m' l' - (addMetricInstructions 100 - (addMetricJumps 100 - (addMetricStores 100 - (addMetricLoads 100 mc))))) -> + (cost_SInteract mc)) -> exec (SInteract resvars action argvars) t m l mc post | call: forall t m l mc binds fname args params rets fbody argvs st0 post outcome, map.get e fname = Some (params, rets, fbody) -> map.getmany_of_list l args = Some argvs -> map.putmany_of_list_zip params argvs map.empty = Some st0 -> - exec fbody t m st0 (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc)))) outcome -> + exec fbody t m st0 (cost_SCall_internal mc) outcome -> (forall t' m' mc' st1, outcome t' m' st1 mc' -> exists retvs l', map.getmany_of_list st1 rets = Some retvs /\ map.putmany_of_list_zip binds retvs l = Some l' /\ - post t' m' l' (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc'))))) -> + post t' m' l' (cost_SCall_external mc')) -> exec (SCall binds fname args) t m l mc post - (* TODO think about a non-fixed bound on the cost of function preamble and postamble *) | load: forall t m l mc sz x a o v addr post, map.get l a = Some addr -> load sz m (word.add addr (word.of_Z o)) = Some v -> @@ -531,12 +552,12 @@ Module exec. map.get e fname = Some (params, rets, fbody) -> map.getmany_of_list l args = Some argvs -> map.putmany_of_list_zip params argvs map.empty = Some st -> - exec fbody t m st (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc)))) + exec fbody t m st (cost_SCall_internal mc) (fun t' m' st' mc' => exists retvs l', map.getmany_of_list st' rets = Some retvs /\ map.putmany_of_list_zip binds retvs l = Some l' /\ - post t' m' l' (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc'))))) -> + post t' m' l' (cost_SCall_external mc')) -> exec (SCall binds fname args) t m l mc post. Proof. intros. eapply call; try eassumption. @@ -702,16 +723,17 @@ Section FlatImp2. {env_ok: map.ok env} {ext_spec_ok: ext_spec.ok ext_spec}. + Variable (phase: compphase). Variable (isReg: varname -> bool). Definition SimState: Type := trace * mem * locals * MetricLog. Definition SimExec(e: env)(c: stmt varname): SimState -> (SimState -> Prop) -> Prop := fun '(t, m, l, mc) post => - exec isReg e c t m l mc (fun t' m' l' mc' => post (t', m', l', mc')). + exec phase isReg e c t m l mc (fun t' m' l' mc' => post (t', m', l', mc')). Lemma modVarsSound: forall e s initialT (initialSt: locals) initialM (initialMc: MetricLog) post, - exec isReg e s initialT initialM initialSt initialMc post -> - exec isReg e s initialT initialM initialSt initialMc + exec phase isReg e s initialT initialM initialSt initialMc post -> + exec phase isReg e s initialT initialM initialSt initialMc (fun finalT finalM finalSt _ => map.only_differ initialSt (modVars s) finalSt). Proof. induction 1; @@ -775,6 +797,3 @@ Definition isRegZ (var : Z) : bool := Definition isRegStr (var : String.string) : bool := String.prefix "reg_" var. - - - diff --git a/compiler/src/compiler/FlatToRiscvCommon.v b/compiler/src/compiler/FlatToRiscvCommon.v index c6d6b94f1..017944581 100644 --- a/compiler/src/compiler/FlatToRiscvCommon.v +++ b/compiler/src/compiler/FlatToRiscvCommon.v @@ -288,7 +288,7 @@ Section WithParameters. exists pos, map.get finfo f = Some pos /\ pos mod 4 = 0. Local Notation stmt := (stmt Z). - Local Notation exec := (exec isRegZ). + Local Notation exec := (exec PostSpill isRegZ). (* note: [e_impl_reduced] and [funnames] will shrink one function at a time each time we enter a new function body, to make sure functions cannot call themselves, while diff --git a/compiler/src/compiler/FlatToRiscvFunctions.v b/compiler/src/compiler/FlatToRiscvFunctions.v index 6bbe1d6be..850f769c3 100644 --- a/compiler/src/compiler/FlatToRiscvFunctions.v +++ b/compiler/src/compiler/FlatToRiscvFunctions.v @@ -431,8 +431,13 @@ Section Proofs. Qed. - Local Notation exec := (exec isRegZ). + Local Notation exec := (exec PostSpill isRegZ). + Definition cost_SCall_L mc := + Platform.MetricLogging.addMetricInstructions 95 + (Platform.MetricLogging.addMetricJumps 95 + (Platform.MetricLogging.addMetricLoads 95 + (Platform.MetricLogging.addMetricStores 95 mc))). Lemma compile_function_body_correct: forall (e_impl_full : env) m l mc (argvs : list word) (st0 : locals) (post outcome : Semantics.trace -> mem -> locals -> MetricLog -> Prop) @@ -497,10 +502,7 @@ Section Proofs. (of_list (list_union Z.eqb (List.firstn binds_count (reg_class.all reg_class.arg)) [])) (singleton_set RegisterNames.ra)) (getRegs finalL) /\ - (getMetrics finalL - Platform.MetricLogging.addMetricInstructions 100 - (Platform.MetricLogging.addMetricJumps 100 - (Platform.MetricLogging.addMetricLoads 100 - (Platform.MetricLogging.addMetricStores 100 (getMetrics mach)))) <= + (getMetrics finalL - cost_SCall_L (getMetrics mach) <= lowerMetrics (finalMetricsH - mc))%metricsL /\ goodMachine finalTrace finalMH finalRegsH g finalL). Proof. @@ -1111,6 +1113,7 @@ Section Proofs. end end. cbn in H2p6. + (* cost_SCall constraint: cost_SCall_L >= (...93...) i think? *) blia. + rename l into lH, finalRegsH into lFH', finalRegsH' into lH', st0 into lFH, @@ -1631,6 +1634,8 @@ Section Proofs. split; eauto 8 with map_hints. split; eauto 8 with map_hints. split; eauto 8 with map_hints. + (* cost_SCall constraint: cost_SCall_L + (1,1,1,0) <= cost_SCall_internal + cost_SCall_external *) + unfold exec.cost_SCall_internal, exec.cost_SCall_external, cost_SCall_L in *. MetricsToRiscv.solve_MetricLog. - idtac "Case compile_stmt_correct/SLoad". diff --git a/compiler/src/compiler/FlattenExpr.v b/compiler/src/compiler/FlattenExpr.v index ad809d6f2..224d5447d 100644 --- a/compiler/src/compiler/FlattenExpr.v +++ b/compiler/src/compiler/FlattenExpr.v @@ -343,7 +343,7 @@ Section FlattenExpr1. simpl (disjoint _ _) in *; map_solver locals_ok. - Local Notation exec := (FlatImp.exec FlatImp.isRegStr). + Local Notation exec := (FlatImp.exec FlatImp.PreSpill FlatImp.isRegStr). Lemma seq_with_modVars: forall env t m (l: locals) mc s1 s2 mid post, exec env s1 t m l mc mid -> @@ -906,7 +906,9 @@ Section FlattenExpr1. -- eassumption. -- do 2 eexists. ssplit; try eassumption. ++ simple eapply map.only_differ_putmany; eassumption. - ++ solve_MetricLog. + ++ (* TODO cost_SCall into semantics *) + unfold FlatImp.exec.cost_SCall_internal, FlatImp.exec.cost_SCall_external in *. + solve_MetricLog. - (* interact *) unfold flattenInteract in *. simp. diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index 49340156a..1121c17ca 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -168,7 +168,7 @@ Section WithWordAndMem. Definition FlatWithStrVars: Lang := {| Program := string_keyed_map (list string * list string * FlatImp.stmt string); Valid := map.forall_values ParamsNoDup; - Call := locals_based_call_spec (FlatImp.exec isRegStr); + Call := locals_based_call_spec (FlatImp.exec PreSpill isRegStr); |}. (* | *) @@ -182,7 +182,7 @@ Section WithWordAndMem. Definition FlatWithZVars: Lang := {| Program := string_keyed_map (list Z * list Z * FlatImp.stmt Z); Valid := map.forall_values ParamsNoDup; - Call := locals_based_call_spec (FlatImp.exec isRegZ); + Call := locals_based_call_spec (FlatImp.exec PreSpill isRegZ); |}. (* | *) (* | Spilling *) @@ -190,7 +190,7 @@ Section WithWordAndMem. Definition FlatWithRegs: Lang := {| Program := string_keyed_map (list Z * list Z * FlatImp.stmt Z); Valid := map.forall_values FlatToRiscvDef.valid_FlatImp_fun; - Call := locals_based_call_spec (FlatImp.exec isRegZ); + Call := locals_based_call_spec_spilled (FlatImp.exec PostSpill isRegZ); |}. (* | *) (* | FlatToRiscv *) @@ -408,7 +408,7 @@ Section WithWordAndMem. Proof. unfold FlatWithZVars, FlatWithRegs. split; cbn. 1: exact spilling_preserves_valid. - unfold locals_based_call_spec. intros. fwd. + unfold locals_based_call_spec, locals_based_call_spec_spilled. intros. fwd. pose proof H0 as GL. unfold spill_functions in GL. eapply map.try_map_values_fw in GL. 2: eassumption. @@ -424,8 +424,11 @@ Section WithWordAndMem. fwd. exists argnames2, retnames2, fbody2, l'. split. 1: exact G2. split. 1: eassumption. - intros. eapply spill_fun_correct; try eassumption. - unfold call_spec. intros * E. rewrite E in *. fwd. eauto. + intros. + eapply FlatImp.exec.weaken. + - eapply spill_fun_correct; try eassumption. + unfold call_spec. intros * E. rewrite E in *. fwd. eauto. + - simpl. intros. fwd. repeat (eexists; split; eauto with metric_arith). Qed. Lemma riscv_phase_correct: phase_correct FlatWithRegs RiscvLang (riscvPhase compile_ext_call). diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index a7d21cb05..4f12e9954 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -1290,11 +1290,11 @@ Section CheckerCorrect. Lemma checker_correct: forall (e: srcEnv) (e': impEnv) s t m lH mcH post, check_funcs e e' = Success tt -> - exec isRegStr e s t m lH mcH post -> + exec PreSpill isRegStr e s t m lH mcH post -> forall lL corresp corresp' s' mcL, check corresp s s' = Success corresp' -> states_compat lH (precond corresp s s') lL -> - exec isRegZ e' s' t m lL mcL (fun t' m' lL' mcL' => + exec PreSpill isRegZ e' s' t m lL mcL (fun t' m' lL' mcL' => exists lH' mcH', states_compat lH' corresp' lL' /\ MetricLogging.metricsLeq (MetricLogging.metricsSub mcL' mcL) (MetricLogging.metricsSub mcH' mcH) /\ post t' m' lH' mcH'). diff --git a/compiler/src/compiler/Spilling.v b/compiler/src/compiler/Spilling.v index 993e942ed..61aaadfdc 100644 --- a/compiler/src/compiler/Spilling.v +++ b/compiler/src/compiler/Spilling.v @@ -20,7 +20,8 @@ Open Scope Z_scope. Section Spilling. Notation stmt := (stmt Z). - Notation exec := (exec isRegZ). + Notation execpre := (exec PreSpill isRegZ). + Notation execpost := (exec PostSpill isRegZ). Definition zero := 0. Definition ra := 1. @@ -460,7 +461,7 @@ Section Spilling. (related maxvar frame fpval t1 m1 l1 t2 m2 (map.put l2 (iarg_reg i r) v) -> post t2 m2 (map.put l2 (iarg_reg i r) v) (if isRegZ r then mc2 else (addMetricInstructions 1 (addMetricLoads 2 mc2)))) -> - exec e2 (load_iarg_reg i r) t2 m2 l2 mc2 post. + execpost e2 (load_iarg_reg i r) t2 m2 l2 mc2 post. Proof. intros. unfold load_iarg_reg, stack_loc, iarg_reg, related in *. fwd. @@ -559,7 +560,7 @@ Section Spilling. related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> fp < r <= maxvar /\ (r < a0 \/ a7 < r) -> map.get l1 r = Some v -> - exec e2 (load_iarg_reg i r) t2 m2 l2 mc2 (fun t2' m2' l2' mc2' => + execpost e2 (load_iarg_reg i r) t2 m2 l2 mc2 (fun t2' m2' l2' mc2' => t2' = t2 /\ m2' = m2 /\ l2' = map.put l2 (iarg_reg i r) v /\ related maxvar frame fpval t1 m1 l1 t2' m2' l2' /\ (mc2' <= (if isRegZ r then mc2 else (addMetricInstructions 1 (addMetricLoads 2 mc2))))%metricsH). @@ -611,7 +612,7 @@ Section Spilling. related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> fp < x <= maxvar /\ (x < a0 \/ a7 < x) -> (mc2' - mc2 <= mc1' - (if isRegZ x then mc1 else (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc1)))))%metricsH -> - exec e (save_ires_reg x) t2 m2 (map.put l2 (ires_reg x) v) mc2' + execpost e (save_ires_reg x) t2 m2 (map.put l2 (ires_reg x) v) mc2' (fun t2' m2' l2' mc2'' => exists t1' m1' l1' mc1'', related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ post t1' m1' l1' mc1'' /\ @@ -716,7 +717,7 @@ Section Spilling. (forall t2' m2' l2', related maxvar frame fpval t1 m1 (map.put l1 x v) t2' m2' l2' -> post t2' m2' l2' (if isRegZ x then mc2 else (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc2))))) -> - exec e (save_ires_reg x) t2 m2 (map.put l2 (ires_reg x) v) mc2 post. + execpost e (save_ires_reg x) t2 m2 (map.put l2 (ires_reg x) v) mc2 post. Proof. intros. unfold save_ires_reg, stack_loc, ires_reg, related in *. fwd. @@ -903,7 +904,7 @@ Section Spilling. related maxvar frame fpval t1 m1 l1' t2 m2' l2' -> mc2' = (cost_set_vars_to_reg_range args start mc2) -> post t2 m2' l2' mc2') -> - exec e (set_vars_to_reg_range args start) t2 m2 l2 mc2 post. + execpost e (set_vars_to_reg_range args start) t2 m2 l2 mc2 post. Proof. induction args; intros. - simpl. eapply exec.skip. fwd. eauto. @@ -992,7 +993,7 @@ Section Spilling. map.getmany_of_list l2' (List.unfoldn (Z.add 1) (List.length args) start) = Some argvs -> mc2' = (cost_set_reg_range_to_vars start args mc2) -> post t2 m2 l2' mc2') -> - exec e (set_reg_range_to_vars start args) t2 m2 l2 mc2 post. + execpost e (set_reg_range_to_vars start args) t2 m2 l2 mc2 post. Proof. induction args; intros. - simpl. eapply exec.skip. eapply H5. 1: eassumption. @@ -1064,6 +1065,98 @@ Section Spilling. rewrite H6; trivial. Qed. + (* XXX this section seems silly *) + + Lemma factor_out_set_reg_range_to_vars : forall args start mc, + cost_set_reg_range_to_vars start args mc = + (mc + cost_set_reg_range_to_vars start args EmptyMetricLog)%metricsH. + Proof. + induction args; eauto with metric_arith. + simpl; intros. + rewrite IHargs. + destruct (isRegZ a); solve_MetricLog_piecewise. + Qed. + + Lemma factor_out_set_vars_to_reg_range : forall args start mc, + cost_set_vars_to_reg_range args start mc = + (mc + cost_set_vars_to_reg_range args start EmptyMetricLog)%metricsH. + Proof. + induction args; eauto with metric_arith. + simpl; intros. + rewrite IHargs. + destruct (isRegZ a); solve_MetricLog_piecewise. + Qed. + + Lemma factor_out_cost_external : forall phase mc, + exec.cost_SCall_external phase mc = + (mc + exec.cost_SCall_external phase EmptyMetricLog)%metricsH. + Proof. destruct phase; eauto with metric_arith. Qed. + + Lemma factor_out_cost_internal : forall phase mc, + exec.cost_SCall_internal phase mc = + (mc + exec.cost_SCall_internal phase EmptyMetricLog)%metricsH. + Proof. destruct phase; eauto with metric_arith. Qed. + + Lemma factor_out_cost_stackalloc : forall x mc, + exec.cost_SStackalloc isRegZ x mc = + (mc + exec.cost_SStackalloc isRegZ x EmptyMetricLog)%metricsH. + Proof. + intros. + unfold exec.cost_SStackalloc, EmptyMetricLog. + destruct (isRegZ x); solve_MetricLog_piecewise. + Qed. + + Lemma cost_set_reg_range_to_vars_bound : forall args start mc len, + Z.of_nat (Datatypes.length args) <= len -> + (cost_set_reg_range_to_vars start args mc <= addMetricInstructions len (addMetricLoads (2 * len) mc))%metricsH. + Proof. + induction args. + - unfold cost_set_reg_range_to_vars, Z.of_nat. intros. simpl in H. solve_MetricLog. + - intros. + cbn [cost_set_reg_range_to_vars]. + specialize (IHargs (start+1) mc (Z.of_nat (Datatypes.length args)) (Z.le_refl _)). + subst. + simpl in H. + destruct (isRegZ a); solve_MetricLog. + Qed. + + Lemma cost_set_vars_to_reg_range_bound : forall args start mc len, + Z.of_nat (Datatypes.length args) <= len -> + (cost_set_vars_to_reg_range args start mc <= addMetricInstructions len (addMetricLoads len (addMetricStores len mc)))%metricsH. + Proof. + induction args. + - unfold cost_set_vars_to_reg_range, Z.of_nat. intros. simpl in H. solve_MetricLog. + - intros. + cbn [cost_set_vars_to_reg_range]. + specialize (IHargs (start+1) mc (Z.of_nat (Datatypes.length args)) (Z.le_refl _)). + subst. + simpl in H. + destruct (isRegZ a); solve_MetricLog. + Qed. + + (* pulled and modified from Coq.Program.Tactics *) + Ltac add_hypothesis p := + match type of p with + ?X => match goal with + | [ H : X |- _ ] => fail 1 + | _ => pose proof p + end + end. + + Ltac add_bounds := + repeat match goal with + | _: context[cost_set_reg_range_to_vars ?x ?y ?z] |- _ => + add_hypothesis (cost_set_reg_range_to_vars_bound y x z 8 ltac:(blia)) + | _: context[cost_set_vars_to_reg_range ?x ?y ?z] |- _ => + add_hypothesis (cost_set_vars_to_reg_range_bound x y z 8 ltac:(blia)) + | |- context[cost_set_reg_range_to_vars ?x ?y ?z] => + add_hypothesis (cost_set_reg_range_to_vars_bound y x z 8 ltac:(blia)) + | |- context[cost_set_vars_to_reg_range ?x ?y ?z] => + add_hypothesis (cost_set_vars_to_reg_range_bound x y z 8 ltac:(blia)) + end. + + (* end silly seeming section *) + Lemma grow_related_mem: forall maxvar frame t1 mSmall1 l1 t2 mSmall2 l2 mStack mCombined2 fpval, related maxvar frame fpval t1 mSmall1 l1 t2 mSmall2 l2 -> map.split mCombined2 mSmall2 mStack -> @@ -1230,12 +1323,12 @@ Section Spilling. Definition spilling_correct_for(e1 e2 : env)(s1 : stmt): Prop := forall (t1 : Semantics.trace) (m1 : mem) (l1 : locals) (mc1 : MetricLog) (post : Semantics.trace -> mem -> locals -> MetricLog -> Prop), - exec e1 s1 t1 m1 l1 mc1 post -> + execpre e1 s1 t1 m1 l1 mc1 post -> forall (frame : mem -> Prop) (maxvar : Z), valid_vars_src maxvar s1 -> forall (t2 : Semantics.trace) (m2 : mem) (l2 : locals) (mc2 : MetricLog) (fpval : word), related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> - exec e2 (spill_stmt s1) t2 m2 l2 mc2 + execpost e2 (spill_stmt s1) t2 m2 l2 mc2 (fun (t2' : Semantics.trace) (m2' : mem) (l2' : locals) (mc2' : MetricLog) => exists t1' m1' l1' mc1', related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ @@ -1243,12 +1336,20 @@ Section Spilling. (mc2' - mc2 <= mc1' - mc1)%metricsH). Definition call_spec(e: env) '(argnames, retnames, fbody) - (t: Semantics.trace)(m: mem)(argvals: list word) - (post: Semantics.trace -> mem -> list word -> Prop): Prop := - forall l mc, map.of_list_zip argnames argvals = Some l -> - exec e fbody t m l mc (fun t' m' l' mc' => + (t: Semantics.trace)(m: mem)(argvals: list word)(mc: MetricLog) + (post: Semantics.trace -> mem -> list word -> MetricLog -> Prop): Prop := + forall l, map.of_list_zip argnames argvals = Some l -> + execpre e fbody t m l (exec.cost_SCall_internal PreSpill mc) (fun t' m' l' mc' => + exists retvals, map.getmany_of_list l' retnames = Some retvals /\ + post t' m' retvals mc'). + + Definition call_spec_spilled(e: env) '(argnames, retnames, fbody) + (t: Semantics.trace)(m: mem)(argvals: list word)(mc: MetricLog) + (post: Semantics.trace -> mem -> list word -> MetricLog -> Prop): Prop := + forall l, map.of_list_zip argnames argvals = Some l -> + execpost e fbody t m l mc (fun t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ - post t' m' retvals). + post t' m' retvals mc'). (* In exec.call, there are many maps of locals involved: @@ -1276,14 +1377,16 @@ Section Spilling. what happens in the callee. TODO: actually use that lemma in case exec.call. Moreover, this lemma will also be used in the pipeline, where phases are composed based on the semantics of function calls. *) + Lemma spill_fun_correct_aux: forall e1 e2 argnames1 retnames1 body1 argnames2 retnames2 body2, spill_fun (argnames1, retnames1, body1) = Success (argnames2, retnames2, body2) -> spilling_correct_for e1 e2 body1 -> - forall argvals t m (post: Semantics.trace -> mem -> list word -> Prop), - call_spec e1 (argnames1, retnames1, body1) t m argvals post -> - call_spec e2 (argnames2, retnames2, body2) t m argvals post. + forall argvals t m mcH mcL (post: Semantics.trace -> mem -> list word -> MetricLog -> Prop), + call_spec e1 (argnames1, retnames1, body1) t m argvals mcH post -> + call_spec_spilled e2 (argnames2, retnames2, body2) t m argvals mcL + (fun t' m' l' mcL' => exists mcH', metricsLeq (mcL' - mcL) (mcH' - mcH) /\ post t' m' l' mcH'). Proof. - unfold call_spec, spilling_correct_for. intros * Sp IHexec * Ex lFL3 mc OL2. + unfold call_spec, spilling_correct_for. intros * Sp IHexec * Ex lFL3 OL2. unfold spill_fun in Sp. fwd. apply_in_hyps @map.getmany_of_list_length. apply_in_hyps @map.putmany_of_list_zip_sameLength. @@ -1402,7 +1505,7 @@ Section Spilling. | |- exists _, _ => eexists | |- _ /\ _ => split end. - 4: eassumption. + 5: eassumption. 2: { unfold map.split. eauto. } @@ -1418,8 +1521,11 @@ Section Spilling. } blia. } { eassumption. } - Unshelve. - all: try assumption. + { + add_bounds. + unfold exec.cost_SStackalloc, exec.cost_SCall_internal in *. + destruct (isRegZ fp); solve_MetricLog. + } Qed. @@ -1453,12 +1559,12 @@ Section Spilling. (l1 : locals) (mc1 : MetricLog) (post : Semantics.trace -> mem -> locals -> MetricLog -> Prop): - exec e1 s1 t1 m1 l1 mc1 post -> + execpre e1 s1 t1 m1 l1 mc1 post -> forall (frame : mem -> Prop) (maxvar : Z), valid_vars_src maxvar s1 -> forall (t2 : Semantics.trace) (m2 : mem) (l2 : locals) (mc2 : MetricLog) (fpval : word), related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> - exec e2 (spill_stmt s1) t2 m2 l2 mc2 + execpost e2 (spill_stmt s1) t2 m2 l2 mc2 (fun (t2' : Semantics.trace) (m2' : mem) (l2' : locals) (mc2' : MetricLog) => exists t1' m1' l1' mc1', related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ @@ -1514,7 +1620,10 @@ Section Spilling. unfold sep at 1 in C. destruct C as (mKeepL' & mRest & SC & ? & _). subst mKeepL'. move H2 at bottom. unfold map.split in H2. fwd. eapply map.shrink_disjoint_l; eassumption. } - cbn in *. rewrite H5; rewrite CSet. admit. (* currently wrong *) + cbn in *. subst. + add_bounds. + solve_MetricLog. + (* cost_SInteract constraint: prespill - postspill >= (...32...) i think? *) } (* related for set_vars_to_reg_range_correct: *) unfold related. @@ -1730,7 +1839,24 @@ Section Spilling. { unfold a0, a7. blia. } { eassumption. } { intros m22 l22 mc22 R22. do 4 eexists. split. 1: eassumption. - split; try eassumption. Fail solve_MetricLog. admit. } + split; try eassumption. + subst. + move Hmetrics at bottom. + rewrite factor_out_set_vars_to_reg_range. + rewrite factor_out_cost_external. + rewrite factor_out_set_reg_range_to_vars. + rewrite (factor_out_cost_external PreSpill). + rewrite factor_out_set_vars_to_reg_range in Hmetrics. + rewrite factor_out_cost_stackalloc in Hmetrics. + rewrite factor_out_cost_internal in Hmetrics. + rewrite factor_out_set_reg_range_to_vars in Hmetrics. + rewrite (factor_out_cost_internal PreSpill) in Hmetrics. + add_bounds. + unfold exec.cost_SStackalloc, exec.cost_SCall_internal, exec.cost_SCall_external, EmptyMetricLog in *. + destruct (isRegZ fp); solve_MetricLog. + (* cost_SCall constraint: prespill - postspill >= (...66...) i think? *) + (* TODO XXX extremely slow *) + } - (* exec.load *) eapply exec.seq_cps. @@ -1957,14 +2083,15 @@ Section Spilling. solve_MetricLog. - (* exec.skip *) eapply exec.skip. exists t, m, l, mc. repeat split; eauto; solve_MetricLog. - Admitted. + Qed. Lemma spill_fun_correct: forall e1 e2 argnames1 retnames1 body1 argnames2 retnames2 body2, spill_functions e1 = Success e2 -> spill_fun (argnames1, retnames1, body1) = Success (argnames2, retnames2, body2) -> - forall argvals t m (post: Semantics.trace -> mem -> list word -> Prop), - call_spec e1 (argnames1, retnames1, body1) t m argvals post -> - call_spec e2 (argnames2, retnames2, body2) t m argvals post. + forall argvals t m mcH mcL (post: Semantics.trace -> mem -> list word -> MetricLog -> Prop), + call_spec e1 (argnames1, retnames1, body1) t m argvals mcH post -> + call_spec_spilled e2 (argnames2, retnames2, body2) t m argvals mcL + (fun t' m' l' mcL' => exists mcH', metricsLeq (mcL' - mcL) (mcH' - mcH) /\ post t' m' l' mcH'). Proof. intros. eapply spill_fun_correct_aux; try eassumption. unfold spilling_correct_for. From 68059d6bfd3955b1d792a25b61a0789dd45b8273 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Sun, 9 Jun 2024 22:57:38 -0400 Subject: [PATCH 26/62] Incorporate metrics into compiler pipeline proofs The pipeline now knows that metrics are preserved throughout. Many further changes are needed in RegAlloc.v and UseImmediate.v, so changes to theorems in those files are not included in this commit (and it will therefore not compile by itself as a result). --- compiler/src/compiler/LowerPipeline.v | 18 ++++--- compiler/src/compiler/Pipeline.v | 78 +++++++++++++++++++-------- 2 files changed, 66 insertions(+), 30 deletions(-) diff --git a/compiler/src/compiler/LowerPipeline.v b/compiler/src/compiler/LowerPipeline.v index 48d1900a4..1a8ba12e7 100644 --- a/compiler/src/compiler/LowerPipeline.v +++ b/compiler/src/compiler/LowerPipeline.v @@ -376,8 +376,8 @@ Section LowerPipeline. (FlatImp.SInteract resvars extcall argvars). Definition riscv_call(p: list Instruction * pos_map * Z) - (f_name: string)(t: Semantics.trace)(mH: mem)(argvals: list word) - (post: Semantics.trace -> mem -> list word -> Prop): Prop := + (f_name: string)(t: Semantics.trace)(mH: mem)(argvals: list word)(mc: MetricLog) + (post: Semantics.trace -> mem -> list word -> MetricLog -> Prop): Prop := let '(instrs, finfo, req_stack_size) := p in exists f_rel_pos, map.get finfo f_name = Some f_rel_pos /\ @@ -392,7 +392,7 @@ Section LowerPipeline. machine_ok p_funcs stack_start stack_pastend instrs mH Rdata Rexec initial -> runsTo initial (fun final => exists mH' retvals, arg_regs_contain final.(getRegs) retvals /\ - post final.(getLog) mH' retvals /\ + post final.(getLog) mH' retvals (raiseMetrics final.(getMetrics)) /\ map.only_differ initial.(getRegs) reg_class.caller_saved final.(getRegs) /\ final.(getPc) = ret_addr /\ machine_ok p_funcs stack_start stack_pastend instrs mH' Rdata Rexec final). @@ -470,15 +470,18 @@ Section LowerPipeline. Lemma flat_to_riscv_correct: forall p1 p2, map.forall_values FlatToRiscvDef.valid_FlatImp_fun p1 -> riscvPhase p1 = Success p2 -> - forall fname t m argvals post, + forall fname t m argvals mc post, (exists argnames retnames fbody l, map.get p1 fname = Some (argnames, retnames, fbody) /\ map.of_list_zip argnames argvals = Some l /\ - forall mc, FlatImp.exec isRegZ p1 fbody t m l mc (fun t' m' l' mc' => + FlatImp.exec PostSpill isRegZ p1 fbody t m l mc (fun t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ - post t' m' retvals)) -> - riscv_call p2 fname t m argvals post. + post t' m' retvals mc')) -> + riscv_call p2 fname t m argvals mc post. Proof. + admit. (* XXX *) + Admitted. + (* unfold riscv_call. intros. destruct p2 as ((finstrs & finfo) & req_stack_size). match goal with @@ -773,5 +776,6 @@ Section LowerPipeline. Unshelve. all: try exact EmptyMetricLog. Qed. + *) End LowerPipeline. diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index 1121c17ca..a92ce3be9 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -47,6 +47,7 @@ Require Export coqutil.Word.SimplWordExpr. Require Export compiler.Registers. Require Import compiler.ForeverSafe. Require Import FunctionalExtensionality. +Require Import PropExtensionality. Require Import coqutil.Tactics.autoforward. Require Import compiler.FitsStack. Require Import compiler.LowerPipeline. @@ -62,8 +63,8 @@ Section WithWordAndMem. Program: Type; Valid: Program -> Prop; Call(p: Program)(funcname: string) - (t: trace)(m: mem)(argvals: list word) - (post: trace -> mem -> list word -> Prop): Prop; + (t: trace)(m: mem)(argvals: list word)(mc: MetricLog) + (post: trace -> mem -> list word -> MetricLog -> Prop): Prop; }. Record phase_correct{L1 L2: Lang} @@ -76,10 +77,13 @@ Section WithWordAndMem. phase_preserves_post: forall p1 p2, L1.(Valid) p1 -> compile p1 = Success p2 -> - forall fname t m argvals post, - L1.(Call) p1 fname t m argvals post -> - L2.(Call) p2 fname t m argvals post; + forall fname t m argvals mc post, + L1.(Call) p1 fname t m argvals mc post -> + L2.(Call) p2 fname t m argvals mc (fun t m a mc' => + exists mc'', metricsLeq mc' mc'' /\ post t m a mc'' + ); }. + (* XXX think about the merits of this kind of metrics statement *) Arguments phase_correct : clear implicits. @@ -90,6 +94,10 @@ Section WithWordAndMem. | Failure e => Failure e end. + Ltac post_ext := + repeat (eapply functional_extensionality; intro); + apply propositional_extensionality. + Lemma compose_phases_correct{L1 L2 L3: Lang} {compile12: L1.(Program) -> result L2.(Program)} {compile23: L2.(Program) -> result L3.(Program)}: @@ -99,7 +107,10 @@ Section WithWordAndMem. Proof. unfold compose_phases. intros [V12 C12] [V23 C23]. - split; intros; fwd; eauto. + split; intros; fwd; [eauto|]. + erewrite f_equal; [eauto|]. + post_ext. + split; [|destruct 1 as (?&?&?&?&?)]; eauto with metric_arith. Qed. Section WithMoreParams. @@ -145,14 +156,29 @@ Section WithWordAndMem. Cmd -> trace -> mem -> locals -> MetricLog -> (trace -> mem -> locals -> MetricLog -> Prop) -> Prop) (e: string_keyed_map (list Var * list Var * Cmd)%type)(f: string) - (t: trace)(m: mem)(argvals: list word) - (post: trace -> mem -> list word -> Prop): Prop := + (t: trace)(m: mem)(argvals: list word)(mc: MetricLog) + (post: trace -> mem -> list word -> MetricLog -> Prop): Prop := exists argnames retnames fbody l, map.get e f = Some (argnames, retnames, fbody) /\ map.of_list_zip argnames argvals = Some l /\ - forall mc, Exec e fbody t m l mc (fun t' m' l' mc' => + Exec e fbody t m l (exec.cost_SCall_internal PreSpill mc) (fun t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ - post t' m' retvals). + post t' m' retvals mc'). + + Definition locals_based_call_spec_spilled{Var Cmd: Type}{locals: map.map Var word} + {string_keyed_map: forall T: Type, map.map string T} + (Exec: string_keyed_map (list Var * list Var * Cmd)%type -> + Cmd -> trace -> mem -> locals -> MetricLog -> + (trace -> mem -> locals -> MetricLog -> Prop) -> Prop) + (e: string_keyed_map (list Var * list Var * Cmd)%type)(f: string) + (t: trace)(m: mem)(argvals: list word)(mc: MetricLog) + (post: trace -> mem -> list word -> MetricLog -> Prop): Prop := + exists argnames retnames fbody l, + map.get e f = Some (argnames, retnames, fbody) /\ + map.of_list_zip argnames argvals = Some l /\ + Exec e fbody t m l mc (fun t' m' l' mc' => + exists retvals, map.getmany_of_list l' retnames = Some retvals /\ + post t' m' retvals mc'). Definition ParamsNoDup{Var: Type}: (list Var * list Var * FlatImp.stmt Var) -> Prop := fun '(argnames, retnames, body) => NoDup argnames /\ NoDup retnames. @@ -236,7 +262,7 @@ Section WithWordAndMem. eexists _, _, _, _. split. 1: eassumption. split. 1: eassumption. intros. eapply FlatImp.exec.weaken. - - eapply flattenStmt_correct_aux with (mcH := mc). + - eapply flattenStmt_correct_aux. + eassumption. + eauto. + reflexivity. @@ -257,7 +283,9 @@ Section WithWordAndMem. eapply start_state_spec. 2: exact A. eapply ListSet.In_list_union_l. eapply ListSet.In_list_union_l. assumption. + eapply @freshNameGenState_disjoint_fbody. - - simpl. intros. fwd. eauto using map.getmany_of_list_extends. + - simpl. intros. fwd. eexists. split. + + eauto using map.getmany_of_list_extends. + + eauto with metric_arith. Qed. Lemma useimmediate_functions_NoDup: forall funs funs', @@ -295,9 +323,10 @@ Section WithWordAndMem. eexists _, _, _, _. split. 1: eassumption. split. 1: eassumption. intros. eapply exec.weaken. - - eapply useImmediate_correct_aux. - all: eauto. - - eauto. + - eapply useImmediate_correct_aux; eauto. + - simpl. destruct 1 as (?&?&?&?&?). + repeat (eexists; split; try eassumption). + eauto with metric_arith. Qed. Lemma regalloc_functions_NoDup: forall funs funs', @@ -351,8 +380,9 @@ Section WithWordAndMem. edestruct putmany_of_list_zip_states_compat as (lL' & P' & Cp); try eassumption. 1: eapply states_compat_empty. rewrite H1 in P'. inversion P'. exact Cp. - - simpl. intros. fwd. eexists. split. 2: eassumption. - eauto using states_compat_getmany. + - simpl. intros. fwd. eexists. split. + + eauto using states_compat_getmany. + + eauto with metric_arith. Unshelve. all: repeat constructor. Qed. @@ -436,8 +466,9 @@ Section WithWordAndMem. unfold FlatWithRegs, RiscvLang. split; cbn. - intros p1 ((? & finfo) & ?). intros. exact I. - - eapply flat_to_riscv_correct; eassumption. - Qed. + - admit. + (*- eapply flat_to_riscv_correct; eassumption.*) + Admitted. Definition composed_compile: Semantics.env -> @@ -548,10 +579,11 @@ Section WithWordAndMem. rewrite map.of_list_tuples. reflexivity. } specialize C with (1 := H0'). - specialize C with (1 := H1). - cbv iota in C. - fwd. eauto 10. - Qed. + admit. + (*specialize C with (1 := H1).*) + (*cbv iota in C.*) + (*fwd. eauto 10.*) + Admitted. (* XXX *) Definition instrencode(p: list Instruction): list byte := List.flat_map (fun inst => HList.tuple.to_list (LittleEndian.split 4 (encode inst))) p. From 86a0f8a35d362b15691a6abb27f6cfbec2b5798c Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Sun, 9 Jun 2024 23:48:10 -0400 Subject: [PATCH 27/62] Fix admits in RegAlloc step This is now 2/5 compiler stages fully verified with metrics. --- compiler/src/compiler/Pipeline.v | 2 -- compiler/src/compiler/RegAlloc.v | 56 +++++++++++++++++++------------- 2 files changed, 33 insertions(+), 25 deletions(-) diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index a92ce3be9..8b87732ad 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -383,8 +383,6 @@ Section WithWordAndMem. - simpl. intros. fwd. eexists. split. + eauto using states_compat_getmany. + eauto with metric_arith. - Unshelve. - all: repeat constructor. Qed. Ltac debool := diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index 4f12e9954..3b655c452 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -581,6 +581,13 @@ Definition mapping_eqb: srcvar * impvar -> srcvar * impvar -> bool := Definition check_regs (x: srcvar) (x': impvar) : bool := negb (andb (isRegStr x) (negb (isRegZ x'))). +Definition check_regs_op (x: @operand srcvar) (x': @operand impvar) : bool := + match x, x' with + | Var vx, Var vx' => check_regs vx vx' + | Const cx, Const cx' => Z.eqb cx cx' + | _, _ => false + end. + Definition assert_in(y: srcvar)(y': impvar)(m: list (srcvar * impvar)): result unit := match List.find (mapping_eqb (y, y')) m with | Some _ => if check_regs y y' then Success tt else @@ -1129,7 +1136,13 @@ Section CheckerCorrect. Opaque isRegStr. Opaque isRegZ. - + + Lemma assert_in_then_check_regs_op : forall x x' corresp, + assert_in_op x x' corresp = Success tt -> check_regs_op x x' = true. + Proof. + unfold check_regs_op, assert_in_op, assert_in; intros; destruct x'; fwd; reflexivity. + Qed. + Lemma check_regs_cost_SLoad: forall x x' a a' mc mc', check_regs a a' = true -> check_regs x x' = true -> @@ -1196,19 +1209,19 @@ Section CheckerCorrect. unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; try blia. Qed. - (*Lemma check_regs_cost_SOp: forall x x' y y' z z' mc mc',*) - (* check_regs x x' = true ->*) - (* check_regs y y' = true ->*) - (* check_regs z z' = true ->*) - (* (MetricLogging.metricsLeq*) - (* (MetricLogging.metricsSub (exec.cost_SOp isRegZ x' y' z' mc') mc')*) - (* (MetricLogging.metricsSub (exec.cost_SOp isRegStr x y z mc) mc)).*) - (*Proof.*) - (* intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLit in *;*) - (* destr (isRegStr x); destr (isRegZ x'); destr (isRegStr y); destr (isRegZ y'); destr (isRegStr z); destr (isRegZ z');*) - (* try discriminate;*) - (* unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; try blia.*) - (*Qed.*) + Lemma check_regs_cost_SOp: forall x x' y y' z z' mc mc', + check_regs x x' = true -> + check_regs y y' = true -> + check_regs_op z z' = true -> + (MetricLogging.metricsLeq + (MetricLogging.metricsSub (exec.cost_SOp isRegZ x' y' z' mc') mc') + (MetricLogging.metricsSub (exec.cost_SOp isRegStr x y z mc) mc)). + Proof. + intros; unfold check_regs_op, check_regs in *; cbn in *; + destr (isRegStr x); destr (isRegZ x'); destr (isRegStr y); destr (isRegZ y'); destruct z as [z|z]; destruct z' as [z'|z']; try destr (isRegStr z); try destr (isRegZ z'); + try discriminate; + unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; blia. + Qed. Lemma check_regs_cost_SSet: forall x x' y y' mc mc', check_regs x x' = true -> @@ -1280,8 +1293,8 @@ Section CheckerCorrect. Hint Constructors exec.exec : checker_hints. Hint Resolve states_compat_get : checker_hints. Hint Resolve states_compat_put : checker_hints. -(* Hint Resolve states_compat_get_op : checker_hints.*) -(* Hint Resolve states_compat_then_op : checker_hints.*) + Hint Resolve states_compat_get_op : checker_hints. + Hint Resolve states_compat_then_op : checker_hints. Hint Resolve check_regs_cost_SLoad check_regs_cost_SStore check_regs_cost_SInlinetable check_regs_cost_SStackalloc check_regs_cost_SLit (*check_regs_cost_SOp*) @@ -1379,8 +1392,7 @@ Section CheckerCorrect. - (* Case exec.op *) repeat econstructor; eauto 10 with checker_hints. all: unfold assert_in, assignment in *; fwd. - all: admit. - (*all: eapply check_regs_cost_SOp; eauto.*) + all: eapply check_regs_cost_SOp; eauto; eapply assert_in_then_check_regs_op; eauto. - (* Case exec.set *) repeat econstructor; eauto 10 with checker_hints. all: unfold assert_in, assignment in *; fwd. @@ -1412,7 +1424,7 @@ Section CheckerCorrect. eapply exec.loop with (mid2 := (fun (t'0 : Semantics.trace) (m'0 : mem) (lL' : impLocals) (mcL' : MetricLogging.MetricLog) => exists (lH' : srcLocals) (mcH' : MetricLogging.MetricLog), - states_compat lH' a lL' /\ + states_compat lH' a1 lL' /\ (exists mcHmid mcLmid, MetricLogging.metricsLeq (MetricLogging.metricsSub mcLmid mcL) (MetricLogging.metricsSub mcHmid mc) /\ MetricLogging.metricsLeq (MetricLogging.metricsSub mcL' mcLmid) (MetricLogging.metricsSub mcH' mcHmid)) /\ @@ -1425,7 +1437,6 @@ Section CheckerCorrect. + cbv beta. intros. fwd. eapply exec.weaken. 1: eapply IH2; eauto using states_compat_eval_bcond_bw. 1: eapply states_compat_precond; eassumption. cbv beta. intros. fwd. eexists. eexists. split. 2: split. 1,3: eauto. - 1: admit. exists mcH'. exists mc'. split; eauto. + cbv beta. intros. fwd. eapply exec.weaken. 1: eapply IH12. 1: eassumption. 1: eassumption. @@ -1438,8 +1449,7 @@ Section CheckerCorrect. rewrite E in P. rewrite E. specialize (P eq_refl). rewrite P. - (*eapply extends_intersect_r.*) - eapply extends_refl. + eapply extends_intersect_r. * cbv beta. intros. fwd. eexists. eexists. split. 2: split. 1: eauto. 2: eauto. intros. repeat (unfold check_bcond, assert_in, assignment in *; fwd). @@ -1464,6 +1474,6 @@ Section CheckerCorrect. - (* case exec.skip *) eapply exec.skip. eexists. exists mc. split. 2: split. 1,3: eauto. destr mc; destr mcL; unfold_metrics; cbn in *; unfold_metrics; cbn in *; blia. - Admitted. + Qed. End CheckerCorrect. From 568b87d6755f73c1c26637375a3d3106423f76c6 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Mon, 10 Jun 2024 01:44:19 -0400 Subject: [PATCH 28/62] Move cost_SStackalloc into postcondition It's nicer to work with this way, since all the exec judgments have the same starting metrics then. --- compiler/src/compiler/FlatImp.v | 6 +++--- compiler/src/compiler/RegAlloc.v | 15 +++++++-------- compiler/src/compiler/Spilling.v | 4 ++-- 3 files changed, 12 insertions(+), 13 deletions(-) diff --git a/compiler/src/compiler/FlatImp.v b/compiler/src/compiler/FlatImp.v index f803a4910..4c0e32506 100644 --- a/compiler/src/compiler/FlatImp.v +++ b/compiler/src/compiler/FlatImp.v @@ -346,7 +346,7 @@ Module exec. | ( true, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) | ( true, true) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) end. - + Definition cost_SStackalloc x mc := match isReg x with | false => (addMetricInstructions 2 (addMetricLoads 2 (addMetricStores 1 mc))) @@ -474,12 +474,12 @@ Module exec. (forall a mStack mCombined, anybytes a n mStack -> map.split mCombined mSmall mStack -> - exec body t mCombined (map.put l x a) (cost_SStackalloc x mc) + exec body t mCombined (map.put l x a) mc (fun t' mCombined' l' mc' => exists mSmall' mStack', anybytes a n mStack' /\ map.split mCombined' mSmall' mStack' /\ - post t' mSmall' l' mc')) -> + post t' mSmall' l' (cost_SStackalloc x mc'))) -> exec (SStackalloc x n body) t mSmall l mc post | lit: forall t m l mc x v post, post t m (map.put l x (word.of_Z v)) (cost_SLit x mc) -> diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index 3b655c452..751728174 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -1184,17 +1184,16 @@ Section CheckerCorrect. Lemma check_regs_cost_SStackalloc: forall x x' mcL mcL' mcH mcH', check_regs x x' = true -> + (MetricLogging.metricsLeq (MetricLogging.metricsSub mcL' mcL) + (MetricLogging.metricsSub mcH' mcH)) -> (MetricLogging.metricsLeq - (MetricLogging.metricsSub mcL' (exec.cost_SStackalloc isRegZ x' mcL)) - (MetricLogging.metricsSub mcH' (exec.cost_SStackalloc isRegStr x mcH))) -> - (MetricLogging.metricsLeq - (MetricLogging.metricsSub mcL' mcL) - (MetricLogging.metricsSub mcH' mcH)). + (MetricLogging.metricsSub (exec.cost_SStackalloc isRegZ x' mcL') mcL) + (MetricLogging.metricsSub (exec.cost_SStackalloc isRegStr x mcH') mcH)). Proof. intros; unfold check_regs in *; cbn in *; unfold exec.cost_SStackalloc in *; destr (isRegStr x); destr (isRegZ x'); try discriminate; - unfold_metrics; cbn; repeat split; destruct mcL; destruct mcH; destruct mcL'; destruct mcH'; cbn in *; unfold_metrics; cbn in *; try blia. + unfold_metrics; cbn; repeat split; destruct mcL; destruct mcH; destruct mcL'; destruct mcH'; cbn in *; unfold_metrics; cbn in *; try blia. Qed. Lemma check_regs_cost_SLit: forall x x' mc mc', @@ -1375,7 +1374,7 @@ Section CheckerCorrect. repeat econstructor; eauto 10 with checker_hints. all: unfold assert_in, assignment in *; fwd. all: eapply check_regs_cost_SInlinetable; eauto. - - (* Case exec.stackalloc *) + - (* Case exec.stackalloc *) eapply exec.stackalloc. 1: assumption. intros. eapply exec.weaken. + eapply H2; try eassumption. @@ -1384,7 +1383,7 @@ Section CheckerCorrect. eexists. eexists. repeat (split; try eassumption). eexists. eexists. repeat (split; try eassumption). all: unfold assert_in, assignment in *; fwd. - all: eapply check_regs_cost_SStackalloc; eauto. + all: eapply check_regs_cost_SStackalloc; eauto. - (* Case exec.lit *) repeat econstructor; eauto 10 with checker_hints. all: unfold assert_in, assignment in *; fwd. diff --git a/compiler/src/compiler/Spilling.v b/compiler/src/compiler/Spilling.v index 61aaadfdc..82239a7a4 100644 --- a/compiler/src/compiler/Spilling.v +++ b/compiler/src/compiler/Spilling.v @@ -1844,10 +1844,10 @@ Section Spilling. move Hmetrics at bottom. rewrite factor_out_set_vars_to_reg_range. rewrite factor_out_cost_external. + rewrite factor_out_cost_stackalloc. rewrite factor_out_set_reg_range_to_vars. rewrite (factor_out_cost_external PreSpill). rewrite factor_out_set_vars_to_reg_range in Hmetrics. - rewrite factor_out_cost_stackalloc in Hmetrics. rewrite factor_out_cost_internal in Hmetrics. rewrite factor_out_set_reg_range_to_vars in Hmetrics. rewrite (factor_out_cost_internal PreSpill) in Hmetrics. @@ -1855,7 +1855,7 @@ Section Spilling. unfold exec.cost_SStackalloc, exec.cost_SCall_internal, exec.cost_SCall_external, EmptyMetricLog in *. destruct (isRegZ fp); solve_MetricLog. (* cost_SCall constraint: prespill - postspill >= (...66...) i think? *) - (* TODO XXX extremely slow *) + (* TODO XXX fairly slow *) } - (* exec.load *) From 56a2834e4d69b96d20d9ad0e4040ce4808df826d Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Mon, 10 Jun 2024 15:12:37 -0400 Subject: [PATCH 29/62] Fixed UseImmediate proof Largely rewritten from scratch except for the SSeq part. There are now 3/5 compiler phases fully verified with metrics! --- compiler/src/compiler/FlatImp.v | 2 +- compiler/src/compiler/UseImmediate.v | 174 +++++++++++++++++++++------ 2 files changed, 135 insertions(+), 41 deletions(-) diff --git a/compiler/src/compiler/FlatImp.v b/compiler/src/compiler/FlatImp.v index 4c0e32506..cb92e77ee 100644 --- a/compiler/src/compiler/FlatImp.v +++ b/compiler/src/compiler/FlatImp.v @@ -360,7 +360,7 @@ Module exec. end. Definition cost_SOp x y z mc := - match (isReg x, isReg y, (match z with | Var vo => isReg vo | Const _ => false end)) with + match (isReg x, isReg y, (match z with | Var vo => isReg vo | Const _ => true end)) with | (false, false, false) => (addMetricInstructions 5 (addMetricLoads 7 (addMetricStores 1 mc))) | (false, false, true) | (false, true, false) => (addMetricInstructions 4 (addMetricLoads 5 (addMetricStores 1 mc))) | (false, true, true) => (addMetricInstructions 3 (addMetricLoads 3 (addMetricStores 1 mc))) diff --git a/compiler/src/compiler/UseImmediate.v b/compiler/src/compiler/UseImmediate.v index ceff40333..8a197bf2b 100644 --- a/compiler/src/compiler/UseImmediate.v +++ b/compiler/src/compiler/UseImmediate.v @@ -25,40 +25,131 @@ Section WithArguments. morphism (word.ring_morph (word := word)), constants [word_cst]). - Local Hint Constructors exec: core. - Local Notation exec := (exec isRegStr). + Local Notation exec := (exec PreSpill isRegStr). + + Open Scope MetricH_scope. + + Ltac tandem H := ( + repeat match goal with + | |- exists _, _ => let x := fresh in destruct H as (x&?); exists x + | |- _ /\ _ => destruct H as (?&H); split; try eassumption + | _ => let x := fresh in intro x; specialize (H x) + end + ). + + Ltac finish := ( + simpl; + intros *; + repeat match goal with + | |- (exists _, _ /\ _) -> _ => intros (?&?&?) + | |- _ /\ _ => split; eauto + | |- exists _, _ => eexists + end; + repeat match goal with + | H : _ <= _ |- _ => revert H + end; + clear; + unfold exec.cost_SInteract, exec.cost_SCall_internal, + exec.cost_SCall_external, exec.cost_SLoad, exec.cost_SStore, + exec.cost_SInlinetable, exec.cost_SStackalloc, exec.cost_SLit, + exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, exec.cost_SLoop_true, + exec.cost_SLoop_false; + repeat match goal with + | x : operand |- _ => destr x + | x : bcond _ |- _ => destr x + | |- context[isRegStr ?x] => destr (isRegStr x) + end; + try solve_MetricLog + ). + + (* TODO these two lemmas are somewhat slow *) + Lemma op_cost_y : forall x0 y v0 mcH' mc v mcL lit, + exec.cost_SOp isRegStr x0 y (Var v0) EmptyMetricLog - EmptyMetricLog <= + mcH' - exec.cost_SLit isRegStr lit mc -> + exec.cost_SOp isRegStr x0 y (Const v) (exec.cost_SLit isRegStr lit mcL) - mcL <= + mcH' - mc. + Proof. finish. Qed. + + Lemma op_cost_v0 : forall x0 y v0 mcH' mc v mcL lit, + exec.cost_SOp isRegStr x0 y (Var v0) EmptyMetricLog - EmptyMetricLog <= + mcH' - exec.cost_SLit isRegStr lit mc -> + exec.cost_SOp isRegStr x0 v0 (Const v) (exec.cost_SLit isRegStr lit mcL) - mcL <= + mcH' - mc. + Proof. finish. Qed. Lemma useImmediate_correct_aux: forall eH eL, (useimmediate_functions is5BitImmediate is12BitImmediate) eH = Success eL -> forall sH t m mcH lH post, exec eH sH t m lH mcH post -> - exec eL (useImmediate is5BitImmediate is12BitImmediate sH) t m lH mcH post. + forall mcL, + exec eL (useImmediate is5BitImmediate is12BitImmediate sH) t m lH mcL + (fun t' m' l' mcL' => exists mcH', metricsLeq (mcL' - mcL) (mcH' - mcH) /\ post t' m' l' mcH'). Proof. - induction 2. - (* most cases stay the same *) - all: try solve [simpl; eauto]. - - (* SCall *) - { simpl. - eapply @exec.call; try eassumption. - assert (exists v2, (useimmediate_function is5BitImmediate is12BitImmediate) (params, rets, fbody) = Success v2 /\ map.get eL fname = Some v2). - { eapply map.try_map_values_fw. - - simpl in H; eapply H. - - eassumption. - } - destruct H5. destruct H5. simpl in H5. inversion H5. fwd. - eassumption. - } - - (* SSeq *) - { simpl. - repeat (match goal with - | |- context[match ?x with _ => _ end] => destr x - | |- context[if ?x then _ else _ ] => destr x - end; - try solve [eapply @exec.seq; eassumption]); - simpl in *. + induction 2; try solve [ + simpl; econstructor; eauto; + tandem H3; + finish + ]. + + - (* SCall *) + simpl; econstructor; eauto. + { unfold useimmediate_functions in H. + destruct (map.try_map_values_fw _ _ _ H _ _ H0) as (?&[=Huseimm]&Hmap). + rewrite <- Huseimm in Hmap. + exact Hmap. } + cbv beta. + intros * (?&?&Houtcome). + destruct (H4 _ _ _ _ Houtcome) as (retvs&l'&Hpost). + exists retvs, l'. + tandem Hpost. + finish. + + - (* SStackalloc *) + simpl; econstructor; eauto. + tandem H2. + eapply exec.weaken; [eauto|]. + simpl; intros * (?&?&?&?&?&?&?). + finish. + + - (* SIf true *) + simpl; econstructor; eauto. + eapply exec.weaken; [eauto|]. + finish. + + - (* SIf false *) + simpl; intro; eapply exec.if_false; eauto. + eapply exec.weaken; [eauto|]. + finish. + + - (* SLoop *) + simpl; econstructor; eauto; simpl. + { intros * (?&?&?); eauto. } + { intros * (?&?&?) **. finish. } + { intros * (?&?&?) **. + eapply exec.weaken; [eauto|]. + simpl; intros * (?&?&?). + instantiate (1 := fun t m l MC1 => exists MC2, MC1 - mcL <= MC2 - mc /\ mid2 t m l MC2). + finish. } + { intros * (?&?&?). + eapply exec.weaken; [eauto|]. + finish. } + + - (* SSeq *) + simpl. intro. + + repeat ( + match goal with + | |- context[match ?x with _ => _ end] => destr x + | |- context[if ?x then _ else _ ] => destr x + end; + try solve [ + eapply @exec.seq; eauto; simpl; + intros * (?&?&?); + eapply @exec.weaken; [eauto|]; + finish + ] + ). all: eapply @exec.seq_cps; eapply @exec.lit. @@ -71,10 +162,10 @@ Section WithArguments. all: match goal with | H: ?mid _ _ _ _, - H0: forall t m l mc, - ?mid t m l mc -> exec ?eL _ _ _ _ _ ?post - |- exec ?eL _ _ _ _ _ ?post - => apply H0 in H; inversion H + H0: forall t m l mc, + ?mid t m l mc -> forall mcL, exec ?eL _ _ _ _ mcL _ + |- exec ?eL _ _ _ _ _ _ + => specialize (H0 _ _ _ _ H EmptyMetricLog); inversion H0 end. all: simpl in *; @@ -83,14 +174,17 @@ Section WithArguments. => rewrite map.get_put_same in H; fwd end. - all: eapply @exec.op; simpl in *; [ eassumption | reflexivity | try eassumption ]. + all: eapply @exec.op; simpl in *; [ eassumption | reflexivity | ]. + + all: exists mcH'; split; [solve [eapply op_cost_y; eauto | eapply op_cost_v0; eauto]|eauto]. + + + rewrite word.add_comm. assumption. + + replace (word.add y' (word.of_Z (-v))) with (word.sub y' (word.of_Z v)) by ring. assumption. + + rewrite word.and_comm. assumption. + + rewrite word.or_comm. assumption. + + rewrite word.xor_comm. assumption. + + (* this is slightly slow also *) + Qed. - all: admit. - (*{ rewrite word.add_comm. assumption. }*) - (*{ replace (word.add y' (word.of_Z (- v))) with (word.sub y' (word.of_Z v)) by ring. assumption. }*) - (*{ rewrite word.and_comm. assumption. }*) - (*{ rewrite word.or_comm. assumption. }*) - (*{ rewrite word.xor_comm. assumption. }*) - } - Admitted. End WithArguments. From 39ff02b3b8e01d797070c8c95ca8bfe52ef853cb Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Mon, 10 Jun 2024 23:09:16 -0400 Subject: [PATCH 30/62] Incorporate metrics into DCE Also introduce MetricTactics, which I expect will be very useful elsewhere. --- compiler/src/compiler/DeadCodeElim.v | 109 +++++++++++++----------- compiler/src/compiler/DeadCodeElimDef.v | 3 +- compiler/src/compiler/MetricTactics.v | 66 ++++++++++++++ compiler/src/compiler/Pipeline.v | 21 ++--- 4 files changed, 136 insertions(+), 63 deletions(-) create mode 100644 compiler/src/compiler/MetricTactics.v diff --git a/compiler/src/compiler/DeadCodeElim.v b/compiler/src/compiler/DeadCodeElim.v index 0a29b06bb..305ad6af4 100644 --- a/compiler/src/compiler/DeadCodeElim.v +++ b/compiler/src/compiler/DeadCodeElim.v @@ -10,7 +10,7 @@ Require Import coqutil.Datatypes.ListSet. Local Notation var := String.string (only parsing). Require Import compiler.util.Common. Require Import bedrock2.MetricLogging. -Require Import coqutil.Tactics.fwd. +Require Import compiler.MetricTactics. (* below only for of_list_list_diff *) Require Import compiler.DeadCodeElimDef. @@ -109,6 +109,7 @@ Section WithArguments1. rewrite ListSet.of_list_removeb end. + Ltac mcsolve := eexists; split; [|split; cycle 1; [eauto|exec_cost_hammer]]; try assumption. Lemma dce_correct_aux : forall eH eL, @@ -117,7 +118,7 @@ Section WithArguments1. exec eH sH t m lH mcH postH -> forall used_after lL mcL, map.agree_on (of_list (live sH used_after)) lH lL -> - exec eL (dce sH used_after) t m lL mcL (compile_post used_after postH). + exec eL (dce sH used_after) t m lL mcL (compile_post mcH mcL used_after postH). Proof. induction 2; match goal with @@ -141,10 +142,9 @@ Section WithArguments1. * eapply H5. * intros. unfold compile_post. - exists l'. eexists. split. - -- agree_on_solve. repeat listset_to_set. - subset_union_solve. - -- eauto. + exists l'. mcsolve. + agree_on_solve. repeat listset_to_set. + subset_union_solve. - intros. eapply @exec.call; try solve [ eassumption ]. + unfold dce_functions, dce_function in *. @@ -158,16 +158,16 @@ Section WithArguments1. eapply agree_on_refl. + intros. unfold compile_post in *. - fwd. eapply H4 in H6p1. fwd. + fwd. eapply H4 in H6p2. fwd. let Heq := fresh in - pose proof H6p1p1 as Heq; - eapply map.putmany_of_list_zip_sameLength, map.sameLength_putmany_of_list in H6p1p1. fwd. + pose proof H6p2p1 as Heq; + eapply map.putmany_of_list_zip_sameLength, map.sameLength_putmany_of_list in H6p2p1. fwd. exists retvs. eexists. repeat split. * erewrite agree_on_getmany. - -- eapply H6p1p0. + -- eapply H6p2p0. -- listset_to_set. agree_on_solve. - * eapply H6p1p1. - * do 2 eexists. split; [ | eassumption ]. + * eapply H6p2p1. + * eexists. mcsolve. agree_on_solve. repeat listset_to_set. subset_union_solve. @@ -178,12 +178,12 @@ Section WithArguments1. * rewrite <- H3p1. eassumption. * eauto. * unfold compile_post. - exists (map.put l x v); eexists; split; [ | eassumption ]. + exists (map.put l x v); mcsolve. repeat listset_to_set. agree_on_solve. + eapply @exec.skip. * unfold compile_post. - exists (map.put l x v); eexists; split; [ | eassumption ]. + exists (map.put l x v); mcsolve. repeat listset_to_set. agree_on_solve. - intros. repeat listset_to_set. @@ -195,17 +195,17 @@ Section WithArguments1. + erewrite <- H4p0; eauto. unfold elem_of; destr (a =? v)%string; [ eapply in_eq | eapply in_cons, in_eq ]. + eassumption. - + unfold compile_post. exists l; eexists; split; eassumption. + + unfold compile_post. exists l; mcsolve. - intros. eapply agree_on_find in H4; fwd. destr (existsb (eqb x) used_after); fwd. + eapply @exec.inlinetable; eauto. * rewrite <- H4p1. eassumption. - * unfold compile_post; do 2 eexists; split ; [ | eassumption ]. + * unfold compile_post; eexists; mcsolve. repeat listset_to_set; agree_on_solve. + eapply @exec.skip; eauto. unfold compile_post. - do 2 eexists; split; [ | eassumption ]. + eexists; mcsolve. repeat listset_to_set; agree_on_solve. - intros. repeat listset_to_set. @@ -220,17 +220,17 @@ Section WithArguments1. ++ eassumption. ++ split. ** eassumption. - ** do 2 eexists; split; [ eassumption | eapply H6p1p2 ]. + ** eexists; mcsolve; eauto. - intros. destr (existsb (eqb x) used_after). + eapply @exec.lit. unfold compile_post. repeat listset_to_set. - do 2 eexists; split; [ | eassumption ]. + eexists; mcsolve. agree_on_solve. + eapply @exec.skip. unfold compile_post. repeat listset_to_set. - do 2 eexists; split; [ | eassumption ]. + eexists; mcsolve. agree_on_solve. - destr z. + intros. repeat listset_to_set. @@ -248,11 +248,11 @@ Section WithArguments1. ++ eapply in_eq. ++ eapply in_cons, in_eq. -- unfold compile_post. - do 2 eexists; split; [ | eassumption ]. + eexists; mcsolve. agree_on_solve. * eapply @exec.skip. unfold compile_post. - do 2 eexists; split; [ | eassumption ]. + eexists; mcsolve. agree_on_solve. + intros. eapply agree_on_find in H3; fwd. @@ -260,11 +260,11 @@ Section WithArguments1. * eapply @exec.op. -- rewrite <- H3p1. eassumption. -- simpl. constructor. - -- unfold compile_post. simpl in *. inversion H1. fwd. do 2 eexists; split; [ | eassumption ]. + -- unfold compile_post. simpl in *. inversion H1. fwd. eexists; mcsolve. repeat listset_to_set. agree_on_solve. * eapply @exec.skip. unfold compile_post. - do 2 eexists; split ; [ | eassumption ]. + eexists; mcsolve. repeat listset_to_set. agree_on_solve. - intros. @@ -273,11 +273,11 @@ Section WithArguments1. destr (existsb (eqb x) used_after). { eapply @exec.set. - rewrite <- H2p1; eassumption. - - unfold compile_post. do 2 eexists; split; [ | eassumption ]. + - unfold compile_post. eexists; mcsolve. agree_on_solve. } { eapply @exec.skip. - - unfold compile_post. do 2 eexists; split; [ | eassumption ]. + - unfold compile_post. eexists; mcsolve. agree_on_solve. } - intros. @@ -287,7 +287,10 @@ Section WithArguments1. eapply @exec.if_true. + erewrite agree_on_eval_bcond; [ eassumption | ]. pose agree_on_comm; eauto. - + eauto. + + eapply @exec.weaken; [eauto|]. + unfold compile_post. + intros * (?&?&?&?&?). + eexists. mcsolve. - intros. repeat listset_to_set. eapply agree_on_union in H2; fwd. @@ -295,14 +298,17 @@ Section WithArguments1. eapply @exec.if_false. + erewrite agree_on_eval_bcond; [ eassumption | ]. pose agree_on_comm; eauto. - + eauto. + + eapply @exec.weaken; [eauto|]. + unfold compile_post. + intros * (?&?&?&?&?). + eexists. mcsolve. - intros. cbn - [live]. rename IHexec into IH1. rename H6 into IH12. rename H4 into IH2. cbn - [live] in IH12. - eapply @exec.loop with (mid2 := compile_post (live (SLoop body1 cond body2) used_after) mid2). + eapply @exec.loop with (mid2 := compile_post mc mcL (live (SLoop body1 cond body2) used_after) mid2). { eapply IH1. eapply agree_on_subset. - let Heq := fresh in @@ -312,8 +318,8 @@ Section WithArguments1. } { intros. unfold compile_post in *. - repeat destr H4. - eapply H1 in H6. + repeat destr H4. destr H6. + eapply H1 in H8. erewrite agree_on_eval_bcond; [ eassumption | ]. eapply agree_on_comm. repeat listset_to_set. @@ -324,15 +330,12 @@ Section WithArguments1. } { intros. unfold compile_post in *. - repeat destr H4. - eapply H2 in H8. - - exists x. - eexists. - split. - + repeat listset_to_set. - eapply agree_on_subset; [ | eapply H4 ]. - subset_union_solve. - + eapply H8. + repeat destr H4. destr H8. + eapply H2 in H9. + - exists x. mcsolve. + repeat listset_to_set. + eapply agree_on_subset; [ | eapply H4 ]. + subset_union_solve. - erewrite agree_on_eval_bcond; [ eassumption | ]. repeat listset_to_set. eapply agree_on_subset; [ | eapply H4 ]. @@ -341,24 +344,30 @@ Section WithArguments1. { intros. unfold compile_post in *. - repeat destr H4. - eapply IH2. - - eapply H8. - - erewrite agree_on_eval_bcond; [ eassumption | ]. + repeat destr H4. destr H8. + assert (eval_bcond x cond = Some true) as Hbcond. + { erewrite agree_on_eval_bcond; [ eassumption | ]. repeat listset_to_set. eapply agree_on_subset; [ | eapply H4 ]. - subset_union_solve. + subset_union_solve. } + eapply @exec.weaken; [eapply IH2|]. + - eapply H9. + - exact Hbcond. - repeat listset_to_set. eapply agree_on_subset; [ | eapply H4 ]. subset_union_solve. + - cbv beta. intros * (?&?&?&?&?). + eexists. mcsolve. } { intros. unfold compile_post in *. - repeat destr H4. - eapply IH12. - - eapply H6. + repeat destr H4. destr H6. + eapply @exec.weaken; [eapply IH12|]. + - eapply H8. - eapply H4. + - cbv beta. intros * (?&?&?&?&?). + eexists. mcsolve. } - intros. eapply @exec.seq. @@ -369,9 +378,9 @@ Section WithArguments1. -- eassumption. -- eassumption. * unfold compile_post. intros. fwd. - do 2 eexists; split; eassumption. + eexists. mcsolve. - intros. eapply @exec.skip. - unfold compile_post. do 2 eexists; split; eassumption. + unfold compile_post. eexists. mcsolve. Qed. End WithArguments1. diff --git a/compiler/src/compiler/DeadCodeElimDef.v b/compiler/src/compiler/DeadCodeElimDef.v index 7ebf13a81..5e905c4c1 100644 --- a/compiler/src/compiler/DeadCodeElimDef.v +++ b/compiler/src/compiler/DeadCodeElimDef.v @@ -518,7 +518,7 @@ Section WithArguments1. Definition compile_post - used_after + mcH mcL used_after (postH: Semantics.trace -> mem -> locals -> MetricLog -> Prop) : Semantics.trace -> mem -> locals -> MetricLog -> Prop @@ -526,6 +526,7 @@ Section WithArguments1. (fun t' m' lL' mcL' => exists lH' mcH', map.agree_on (PropSet.of_list used_after) lH' lL' + /\ metricsLeq (mcL' - mcL) (mcH' - mcH) /\ postH t' m' lH' mcH'). Lemma agree_on_eval_bcond: diff --git a/compiler/src/compiler/MetricTactics.v b/compiler/src/compiler/MetricTactics.v new file mode 100644 index 000000000..3db8f2d9c --- /dev/null +++ b/compiler/src/compiler/MetricTactics.v @@ -0,0 +1,66 @@ +Require Import compiler.util.Common. +Require Import bedrock2.MetricLogging. +Require Import compiler.FlatImp. + +Open Scope MetricH_scope. + +Definition metrics_additive f := forall mc, f mc - mc = f EmptyMetricLog. +#[global] Transparent metrics_additive. + +Ltac exec_cost_unfold := + unfold exec.cost_SInteract, exec.cost_SCall_internal, + exec.cost_SCall_external, exec.cost_SLoad, exec.cost_SStore, + exec.cost_SInlinetable, exec.cost_SStackalloc, exec.cost_SLit, exec.cost_SOp, + exec.cost_SSet, exec.cost_SIf, exec.cost_SLoop_true, exec.cost_SLoop_false in + *. + +Ltac exec_cost_destr := + repeat match goal with + | x : compphase |- _ => destr x + | x : operand |- _ => destr x + | x : bcond _ |- _ => destr x + | _ : context[if ?x then _ else _] |- _ => destr x + | |- context[if ?x then _ else _] => destr x + end; unfold EmptyMetricLog. + +Ltac exec_cost_solve := exec_cost_unfold; exec_cost_destr; try solve_MetricLog. +Ltac exec_cost_solve_piecewise := exec_cost_unfold; exec_cost_destr; try solve_MetricLog_piecewise. + +Ltac t := unfold metrics_additive; intros; exec_cost_solve_piecewise. + +Lemma exec_cost_additive_SInteract : forall phase, metrics_additive (exec.cost_SInteract phase). Proof. t. Qed. +Lemma exec_cost_additive_SCall_internal : forall phase, metrics_additive (exec.cost_SCall_internal phase). Proof. t. Qed. +Lemma exec_cost_additive_SCall_external : forall phase, metrics_additive (exec.cost_SCall_external phase). Proof. t. Qed. +Lemma exec_cost_additive_SLoad : forall varname isReg x a, metrics_additive (@exec.cost_SLoad varname isReg x a). Proof. t. Qed. +Lemma exec_cost_additive_SStore : forall varname isReg x a, metrics_additive (@exec.cost_SStore varname isReg x a). Proof. t. Qed. +Lemma exec_cost_additive_SInlinetable : forall varname isReg x a, metrics_additive (@exec.cost_SInlinetable varname isReg x a). Proof. t. Qed. +Lemma exec_cost_additive_SStackalloc : forall varname isReg x, metrics_additive (@exec.cost_SStackalloc varname isReg x). Proof. t. Qed. +Lemma exec_cost_additive_SLit : forall varname isReg x, metrics_additive (@exec.cost_SLit varname isReg x). Proof. t. Qed. +Lemma exec_cost_additive_SOp : forall varname isReg x y z, metrics_additive (@exec.cost_SOp varname isReg x y z). Proof. t. Qed. +Lemma exec_cost_additive_SSet : forall varname isReg x y, metrics_additive (@exec.cost_SSet varname isReg x y). Proof. t. Qed. +Lemma exec_cost_additive_SIf : forall varname isReg x, metrics_additive (@exec.cost_SIf varname isReg x). Proof. t. Qed. +Lemma exec_cost_additive_SLoop_true : forall varname isReg x, metrics_additive (@exec.cost_SLoop_true varname isReg x). Proof. t. Qed. +Lemma exec_cost_additive_SLoop_false : forall varname isReg x, metrics_additive (@exec.cost_SLoop_false varname isReg x). Proof. t. Qed. + +(* COQBUG: + these cannot have "in *", because this causes a crash if the environment contains a Semantics.ExtSpec. + Anomaly "Unable to handle arbitrary u+k <= v constraints." + Please report at http://coq.inria.fr/bugs/. + to use these in hypotheses, revert them first *) +Ltac exec_cost_additive := repeat ( + rewrite exec_cost_additive_SInteract || + rewrite exec_cost_additive_SCall_internal || + rewrite exec_cost_additive_SCall_external || + rewrite exec_cost_additive_SLoad || + rewrite exec_cost_additive_SStore || + rewrite exec_cost_additive_SInlinetable || + rewrite exec_cost_additive_SStackalloc || + rewrite exec_cost_additive_SLit || + rewrite exec_cost_additive_SOp || + rewrite exec_cost_additive_SSet || + rewrite exec_cost_additive_SIf || + rewrite exec_cost_additive_SLoop_true || + rewrite exec_cost_additive_SLoop_false + ). + +Ltac exec_cost_hammer := exec_cost_additive; try solve [eauto 3 with metric_arith | exec_cost_solve]. diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index c24e1118a..ec4621378 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -368,18 +368,15 @@ Section WithWordAndMem. unfold dce_function in GI. fwd. eexists _, _, _, _. split. 1: eassumption. split. 1: eassumption. intros. - admit. - Admitted. - (*eapply @exec.weaken.*) - (*- eapply dce_correct_aux; eauto.*) - (* eapply MapEauto.agree_on_refl.*) - (*- unfold compile_post. intros. fwd.*) - (* exists retvals.*) - (* split.*) - (* + erewrite MapEauto.agree_on_getmany; [ eauto | eapply MapEauto.agree_on_comm; [ eassumption ] ].*) - (* + eassumption.*) - (* Unshelve. eauto.*) - (*Qed.*) + eapply @exec.weaken. + - eapply dce_correct_aux; eauto. + eapply MapEauto.agree_on_refl. + - unfold compile_post. intros. fwd. + exists retvals. + split. + + erewrite MapEauto.agree_on_getmany; [ eauto | eapply MapEauto.agree_on_comm; [ eassumption ] ]. + + eexists; split; eauto. solve_MetricLog. + Qed. Lemma regalloc_functions_NoDup: forall funs funs', regalloc_functions funs = Success funs' -> From 5abef1ca86ebb9175e5a5322704aa64302d3bab7 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Tue, 11 Jun 2024 02:06:38 -0400 Subject: [PATCH 31/62] Use alternative metrics statement style in pipeline I now realize why the rest of the codebase uses this style; it's just easier to work with as a user. --- compiler/src/compiler/Pipeline.v | 18 +++++++++--------- 1 file changed, 9 insertions(+), 9 deletions(-) diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index ec4621378..a8680872f 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -79,13 +79,13 @@ Section WithWordAndMem. phase_preserves_post: forall p1 p2, L1.(Valid) p1 -> compile p1 = Success p2 -> - forall fname t m argvals mc post, - L1.(Call) p1 fname t m argvals mc post -> - L2.(Call) p2 fname t m argvals mc (fun t m a mc' => - exists mc'', metricsLeq mc' mc'' /\ post t m a mc'' + forall fname t m argvals mcH post, + L1.(Call) p1 fname t m argvals mcH post -> + forall mcL, + L2.(Call) p2 fname t m argvals mcL (fun t m a mcL' => + exists mcH', metricsLeq (mcL' - mcL) (mcH' - mcH) /\ post t m a mcH' ); }. - (* XXX think about the merits of this kind of metrics statement *) Arguments phase_correct : clear implicits. @@ -292,7 +292,7 @@ Section WithWordAndMem. + eapply @freshNameGenState_disjoint_fbody. - simpl. intros. fwd. eexists. split. + eauto using map.getmany_of_list_extends. - + eauto with metric_arith. + + eexists. split; [|eassumption]. solve_MetricLog. Qed. Lemma useimmediate_functions_NoDup: forall funs funs', @@ -333,7 +333,7 @@ Section WithWordAndMem. - eapply useImmediate_correct_aux; eauto. - simpl. destruct 1 as (?&?&?&?&?). repeat (eexists; split; try eassumption). - eauto with metric_arith. + solve_MetricLog. Qed. @@ -375,7 +375,7 @@ Section WithWordAndMem. exists retvals. split. + erewrite MapEauto.agree_on_getmany; [ eauto | eapply MapEauto.agree_on_comm; [ eassumption ] ]. - + eexists; split; eauto. solve_MetricLog. + + eexists; split; eauto. unfold exec.cost_SCall_internal in *. solve_MetricLog. Qed. Lemma regalloc_functions_NoDup: forall funs funs', @@ -431,7 +431,7 @@ Section WithWordAndMem. rewrite H1 in P'. inversion P'. exact Cp. - simpl. intros. fwd. eexists. split. + eauto using states_compat_getmany. - + eauto with metric_arith. + + eexists. split; [|eassumption]. solve_MetricLog. Qed. Ltac debool := From 6484871becef2eef10152a99fb10fd53b88e6629 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Tue, 11 Jun 2024 02:07:39 -0400 Subject: [PATCH 32/62] Finish metrics proof of lower pipeline This seemed too easy, so I'm suspicious something's wrong, but everything seems to work. The definition of riscv_call now incurs a penalty equal to cost_SCall_L. I don't fully understand why, but presumably there's some boilerplate added somewhere. --- compiler/src/compiler/LowerPipeline.v | 40 +++++++++++++++++++-------- compiler/src/compiler/Pipeline.v | 5 ++-- 2 files changed, 31 insertions(+), 14 deletions(-) diff --git a/compiler/src/compiler/LowerPipeline.v b/compiler/src/compiler/LowerPipeline.v index 0cb319aad..1541fc50f 100644 --- a/compiler/src/compiler/LowerPipeline.v +++ b/compiler/src/compiler/LowerPipeline.v @@ -32,6 +32,17 @@ Local Arguments Z.modulo : simpl never. Local Arguments Z.pow: simpl never. Local Arguments Z.sub: simpl never. +Lemma raise_metrics_ineq : forall m1 m2, + (m1 <= m2)%metricsL -> (raiseMetrics m1 <= raiseMetrics m2)%metricsH. +Proof. + intros. + destr m1; destr m2. + destruct H as (?&?&?&?). + unfold_MetricLog. + simpl in *. + repeat split; assumption. +Qed. + Section WithWordAndMem. Context {width: Z} {word: word.word width} {mem: map.map word byte}. @@ -384,6 +395,7 @@ Section LowerPipeline. forall p_funcs stack_start stack_pastend ret_addr Rdata Rexec (initial: MetricRiscvMachine), map.get initial.(getRegs) RegisterNames.ra = Some ret_addr -> initial.(getLog) = t -> + raiseMetrics (cost_SCall_L initial.(getMetrics)) = mc -> word.unsigned ret_addr mod 4 = 0 -> arg_regs_contain initial.(getRegs) argvals -> req_stack_size <= word.unsigned (word.sub stack_pastend stack_start) / bytes_per_word -> @@ -470,18 +482,17 @@ Section LowerPipeline. Lemma flat_to_riscv_correct: forall p1 p2, map.forall_values FlatToRiscvDef.valid_FlatImp_fun p1 -> riscvPhase p1 = Success p2 -> - forall fname t m argvals mc post, + forall fname t m argvals mcH post, (exists argnames retnames fbody l, map.get p1 fname = Some (argnames, retnames, fbody) /\ map.of_list_zip argnames argvals = Some l /\ - FlatImp.exec PostSpill isRegZ p1 fbody t m l mc (fun t' m' l' mc' => + FlatImp.exec PostSpill isRegZ p1 fbody t m l mcH (fun t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals mc')) -> - riscv_call p2 fname t m argvals mc post. + forall mcL, + riscv_call p2 fname t m argvals mcL (fun t m a mcL' => + exists mcH', metricsLeq (mcL' - mcL) (mcH' - mcH) /\ post t m a mcH'). Proof. - admit. (* XXX *) - Admitted. - (* unfold riscv_call. intros. destruct p2 as ((finstrs & finfo) & req_stack_size). match goal with @@ -520,7 +531,7 @@ Section LowerPipeline. (l := l) (post := fun t' m' l' mc' => (exists retvals, - map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals)). + map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals mc')). eapply Q with (g := {| rem_stackwords := word.unsigned (word.sub stack_pastend stack_start) / bytes_per_word; @@ -679,7 +690,17 @@ Section LowerPipeline. symmetry. eapply map.getmany_of_list_length. exact GM. - + eassumption. + + eexists. split; [|eassumption]. + apply raise_metrics_ineq in H10p3. + destruct (getMetrics final). + destruct (getMetrics initial). + unfold raiseMetrics, lowerMetrics in *. + unfold_MetricLog. + simpl_MetricLog. + simpl in *. + inversion H2. + solve_MetricLog. + (* XXX ???? *) + eapply only_differ_subset. 1: eassumption. rewrite ListSet.of_list_list_union. rewrite ?singleton_set_eq_of_list. @@ -773,9 +794,6 @@ Section LowerPipeline. + assumption. + assumption. + assumption. - Unshelve. - all: try exact EmptyMetricLog. Qed. - *) End LowerPipeline. diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index a8680872f..de2fa8bb1 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -513,9 +513,8 @@ Section WithWordAndMem. unfold FlatWithRegs, RiscvLang. split; cbn. - intros p1 ((? & finfo) & ?). intros. exact I. - - admit. - (*- eapply flat_to_riscv_correct; eassumption.*) - Admitted. + - eapply flat_to_riscv_correct; eassumption. + Qed. Definition composed_compile: Semantics.env -> From 7cb778356e80ce739efc258f6a42b971e70951b3 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Tue, 11 Jun 2024 02:48:41 -0400 Subject: [PATCH 33/62] Finish metrics patches in pipeline Only one file left with admits! (FlattenExpr) --- compiler/src/compiler/Pipeline.v | 45 +++++++++++++++++++------------- 1 file changed, 27 insertions(+), 18 deletions(-) diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index de2fa8bb1..e2350f161 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -15,7 +15,6 @@ Require Import riscv.Utility.runsToNonDet. Require Export riscv.Platform.MetricRiscvMachine. Require Import coqutil.Z.Lia. Require Import coqutil.Tactics.fwd. -Require Import bedrock2.MetricLogging. Require Import compiler.ExprImp. Require Import compiler.FlattenExprDef. Require Import compiler.FlattenExpr. @@ -39,6 +38,7 @@ Require Import compiler.SeparationLogic. Require Import compiler.Spilling. Require Import compiler.RegAlloc. Require Import compiler.RiscvEventLoop. +Require Import compiler.MetricsToRiscv. Require Import bedrock2.MetricLogging. Require Import compiler.FlatToRiscvCommon. Require Import compiler.FlatToRiscvFunctions. @@ -51,7 +51,7 @@ Require Import PropExtensionality. Require Import coqutil.Tactics.autoforward. Require Import compiler.FitsStack. Require Import compiler.LowerPipeline. -Require Import bedrock2.WeakestPreconditionProperties. +Require Import bedrock2.MetricWeakestPreconditionProperties. Require Import compiler.UseImmediateDef. Require Import compiler.UseImmediate. Require Import compiler.DeadCodeElimDef. @@ -582,18 +582,19 @@ Section WithWordAndMem. (* function we choose to call: *) (fname: string) (* high-level initial state & post on final state: *) - (t: trace) (mH: mem) (argvals: list word) (post: trace -> mem -> list word -> Prop), + (t: trace) (mH: mem) (argvals: list word) (mc: MetricLog) (post: trace -> mem -> list word -> MetricLog -> Prop), valid_src_funs functions = true -> compile functions = Success (instrs, finfo, req_stack_size) -> (exists (argnames retnames: list string) (fbody: cmd) l, map.get (map.of_list (map := Semantics.env) functions) fname = Some (argnames, retnames, fbody) /\ map.of_list_zip argnames argvals = Some l /\ - forall mc, - MetricSemantics.exec (map.of_list functions) fbody t mH l mc + MetricSemantics.exec (map.of_list functions) fbody t mH l + (exec.cost_SCall_internal PreSpill mc) (fun t' m' l' mc' => exists retvals: list word, map.getmany_of_list l' retnames = Some retvals /\ - post t' m' retvals)) -> + post t' m' retvals mc')) -> + forall mcL, exists (f_rel_pos: Z), map.get (map.of_list finfo) fname = Some f_rel_pos /\ forall (* low-level machine on which we're going to run the compiled program: *) @@ -607,11 +608,14 @@ Section WithWordAndMem. word.unsigned ret_addr mod 4 = 0 -> arg_regs_contain initial.(getRegs) argvals -> initial.(getLog) = t -> + raiseMetrics (cost_SCall_L initial.(getMetrics)) = mcL -> machine_ok p_funcs stack_lo stack_hi instrs mH Rdata Rexec initial -> runsTo initial (fun final : MetricRiscvMachine => exists mH' retvals, arg_regs_contain (getRegs final) retvals /\ - post final.(getLog) mH' retvals /\ + (exists mcH' : MetricLog, + ((raiseMetrics final.(getMetrics)) - mcL <= mcH' - mc)%metricsH /\ + post (getLog final) mH' retvals mcH') /\ map.only_differ initial.(getRegs) reg_class.caller_saved final.(getRegs) /\ final.(getPc) = ret_addr /\ machine_ok p_funcs stack_lo stack_hi instrs mH' Rdata Rexec final). @@ -627,11 +631,11 @@ Section WithWordAndMem. rewrite map.of_list_tuples. reflexivity. } specialize C with (1 := H0'). - admit. - (*specialize C with (1 := H1).*) - (*cbv iota in C.*) - (*fwd. eauto 10.*) - Admitted. (* XXX *) + specialize C with (1 := H1). + specialize (C mcL). + cbv iota in C. + fwd. eauto 10. + Qed. Definition instrencode(p: list Instruction): list byte := List.flat_map (fun inst => HList.tuple.to_list (LittleEndian.split 4 (encode inst))) p. @@ -652,7 +656,7 @@ Section WithWordAndMem. (* function we choose to call: *) (fname: string) (f_rel_pos: Z) (* high-level initial state & post on final state: *) - (t: trace) (mH: mem) (argvals: list word) (post: trace -> mem -> list word -> Prop) + (t: trace) (mH: mem) (argvals: list word) (mc: MetricLog) (post: trace -> mem -> list word -> MetricLog -> Prop) (* ghost vars that help describe the low-level machine: *) (stack_lo stack_hi ret_addr p_funcs: word) (Rdata Rexec: mem -> Prop) (* low-level machine on which we're going to run the compiled program: *) @@ -660,7 +664,9 @@ Section WithWordAndMem. valid_src_funs fs = true -> NoDup (map fst fs) -> compile fs = Success (instrs, finfo, req_stack_size) -> - WeakestPrecondition.call (map.of_list fs) fname t mH argvals post -> + MetricWeakestPrecondition.call (map.of_list fs) fname t mH argvals + (exec.cost_SCall_internal PreSpill mc) post -> + forall mcL, map.get (map.of_list finfo) fname = Some f_rel_pos -> req_stack_size <= word.unsigned (word.sub stack_hi stack_lo) / bytes_per_word -> word.unsigned (word.sub stack_hi stack_lo) mod bytes_per_word = 0 -> @@ -669,24 +675,27 @@ Section WithWordAndMem. word.unsigned ret_addr mod 4 = 0 -> arg_regs_contain initial.(getRegs) argvals -> initial.(getLog) = t -> + raiseMetrics (cost_SCall_L initial.(getMetrics)) = mcL -> machine_ok p_funcs stack_lo stack_hi instrs mH Rdata Rexec initial -> runsTo initial (fun final : MetricRiscvMachine => exists mH' retvals, arg_regs_contain (getRegs final) retvals /\ - post final.(getLog) mH' retvals /\ + (exists mcH' : MetricLog, + ((raiseMetrics final.(getMetrics)) - mcL <= mcH' - mc)%metricsH /\ + post (getLog final) mH' retvals mcH') /\ map.only_differ initial.(getRegs) reg_class.caller_saved final.(getRegs) /\ final.(getPc) = ret_addr /\ machine_ok p_funcs stack_lo stack_hi instrs mH' Rdata Rexec final). Proof. intros. - let H := hyp WeakestPrecondition.call in rename H into WP. + let H := hyp MetricWeakestPrecondition.call in rename H into WP. edestruct compiler_correct with (fname := fname) (argvals := argvals) (post := post) as (f_rel_pos' & G & C); try eassumption. 2: { eapply C; clear C; try assumption; try congruence; try eassumption. } intros. - unfold Semantics.call in WP. fwd. + unfold MetricSemantics.call in WP. fwd. do 5 eexists. 1: eassumption. split. 1: eassumption. - intros. eapply MetricSemantics.of_metrics_free. assumption. + intros. assumption. Qed. End WithMoreParams. From 98f5ca2bf84fd5501de1dfddb097b3891994b2ec Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Tue, 11 Jun 2024 02:59:58 -0400 Subject: [PATCH 34/62] Patch ToplevelLoop for metrics changes --- compiler/src/compiler/ToplevelLoop.v | 16 +++++++++------- 1 file changed, 9 insertions(+), 7 deletions(-) diff --git a/compiler/src/compiler/ToplevelLoop.v b/compiler/src/compiler/ToplevelLoop.v index 7b82edff5..99a63f567 100644 --- a/compiler/src/compiler/ToplevelLoop.v +++ b/compiler/src/compiler/ToplevelLoop.v @@ -297,16 +297,15 @@ Section Pipeline1. compile_ext_call_length_ignores_positions as P. unfold runsTo in P. specialize P with (argvals := []) - (post := fun t' m' retvals => isReady spec t' m' /\ goodTrace spec t') + (post := fun t' m' retvals mc' => isReady spec t' m' /\ goodTrace spec t') (fname := "init"%string). edestruct P as (init_rel_pos & G & P'); clear P; cycle -1. 1: eapply P' with (p_funcs := word.add loop_pos (word.of_Z 8)) (Rdata := R). all: simpl_MetricRiscvMachine_get_set. - 11: { + 12: { unfold hl_inv in init_code_correct. move init_code_correct at bottom. do 4 eexists. split. 1: eassumption. split. 1: reflexivity. - intros mc. eapply ExprImp.weaken_exec. - refine (init_code_correct _ _ _). replace (datamem_start spec) with (heap_start ml) by congruence. @@ -314,7 +313,7 @@ Section Pipeline1. exact HMem. - cbv beta. intros * _ HP. exists []. split. 1: reflexivity. exact HP. } - 10: { unfold compile. rewrite_match. reflexivity. } + 11: { unfold compile. rewrite_match. reflexivity. } all: try eassumption. { apply stack_length_divisible. } { cbn. clear CP. @@ -324,6 +323,7 @@ Section Pipeline1. { destruct mlOk. solve_divisibleBy4. } { reflexivity. } { reflexivity. } + { reflexivity. } unfold machine_ok. clear P'. rewrite GetPos in G. fwd. @@ -409,6 +409,7 @@ Section Pipeline1. * eapply iff1ToEq. unfold init_sp_insts, init_insts, loop_insts, backjump_insts. wwcancel. + Unshelve. exact MetricLogging.EmptyMetricLog. Qed. Lemma ll_inv_is_invariant: forall (st: MetricRiscvMachine), @@ -505,14 +506,13 @@ Section Pipeline1. unfold runsTo in P. specialize P with (argvals := []) (fname := "loop"%string) - (post := fun t' m' retvals => isReady spec t' m' /\ goodTrace spec t'). + (post := fun t' m' retvals mc' => isReady spec t' m' /\ goodTrace spec t'). edestruct P as (loop_rel_pos & G & P'); clear P; cycle -1. 1: eapply P' with (p_funcs := word.add loop_pos (word.of_Z 8)) (Rdata := R) (ret_addr := word.add loop_pos (word.of_Z 4)). - 11: { + 12: { move loop_body_correct at bottom. do 4 eexists. split. 1: eassumption. split. 1: reflexivity. - intros mc. eapply ExprImp.weaken_exec. - eapply loop_body_correct; eauto. - cbv beta. intros * _ HP. exists []. split. 1: reflexivity. exact HP. @@ -526,6 +526,7 @@ Section Pipeline1. { subst loop_pos init_pos. destruct mlOk. solve_divisibleBy4. } { reflexivity. } { reflexivity. } + { reflexivity. } unfold loop_pos, init_pos. unfold machine_ok. unfold_RiscvMachine_get_set. @@ -588,6 +589,7 @@ Section Pipeline1. * wcancel_assumption. * eapply rearrange_footpr_subset. 1: eassumption. wwcancel. + Unshelve. exact MetricLogging.EmptyMetricLog. Qed. Lemma ll_inv_implies_prefix_of_good: forall st, From 2e3ada6fef7d587144623be94257024c49ccdc65 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Sat, 15 Jun 2024 02:53:11 -0400 Subject: [PATCH 35/62] Factor out expr/stmt costs into MetricCosts.v Also update bedrock2 semantics for new metrics, which fixes FlattenExpr.v. All the other pipeline stages are now broken and need to be patched. --- bedrock2/src/bedrock2/MetricCosts.v | 175 ++++++++++++++++++ bedrock2/src/bedrock2/MetricLoops.v | 37 ++-- bedrock2/src/bedrock2/MetricSemantics.v | 58 +++--- .../src/bedrock2/MetricWeakestPrecondition.v | 43 +++-- .../MetricWeakestPreconditionProperties.v | 11 +- compiler/src/compiler/FlatImp.v | 172 +++++------------ compiler/src/compiler/FlatToRiscvCommon.v | 1 + compiler/src/compiler/FlatToRiscvFunctions.v | 89 ++++----- compiler/src/compiler/FlattenExpr.v | 55 +++--- compiler/src/compiler/LowerPipeline.v | 1 + compiler/src/compiler/MetricTactics.v | 66 ------- 11 files changed, 351 insertions(+), 357 deletions(-) create mode 100644 bedrock2/src/bedrock2/MetricCosts.v delete mode 100644 compiler/src/compiler/MetricTactics.v diff --git a/bedrock2/src/bedrock2/MetricCosts.v b/bedrock2/src/bedrock2/MetricCosts.v new file mode 100644 index 000000000..a681aef8b --- /dev/null +++ b/bedrock2/src/bedrock2/MetricCosts.v @@ -0,0 +1,175 @@ +Require Import BinIntDef. +Require Import Coq.Strings.String. +Require Import bedrock2.MetricLogging. +From coqutil.Tactics Require Import destr. + +Local Open Scope MetricH_scope. + +Inductive compphase: Type := +| PreSpill +| PostSpill. + +Section FlatImpExec. + + Context {varname: Type}. + Variable (phase: compphase). + Variable (isReg: varname -> bool). + + Definition cost_interact mc := + match phase with + | PreSpill => (addMetricInstructions 100 (addMetricJumps 100 (addMetricStores 100 (addMetricLoads 100 mc)))) + | PostSpill => (addMetricInstructions 50 (addMetricJumps 50 (addMetricStores 50 (addMetricLoads 50 mc)))) + end. + + Definition cost_call_internal mc := + match phase with + | PreSpill => addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc))) + | PostSpill => addMetricInstructions 50 (addMetricJumps 50 (addMetricLoads 50 (addMetricStores 50 mc))) + end. + + Definition cost_call_external mc := + match phase with + | PreSpill => addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc))) + | PostSpill => addMetricInstructions 50 (addMetricJumps 50 (addMetricLoads 50 (addMetricStores 50 mc))) + end. + + (* TODO think about a non-fixed bound on the cost of function preamble and postamble *) + + Definition cost_load x a mc := + match (isReg x, isReg a) with + | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricStores 1 mc))) + | (false, true) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) + | ( true, false) => (addMetricInstructions 2 (addMetricLoads 4 mc)) + | ( true, true) => (addMetricInstructions 1 (addMetricLoads 2 mc)) + end. + + Definition cost_store a v mc := + match (isReg a, isReg v) with + | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricStores 1 mc))) + | (false, true) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) + | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) + | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc))) + end. + + Definition cost_inlinetable x i mc := + match (isReg x, isReg i) with + | (false, false) => (addMetricInstructions 5 (addMetricLoads 7 (addMetricStores 1 (addMetricJumps 1 mc)))) + | (false, true) => (addMetricInstructions 4 (addMetricLoads 5 (addMetricStores 1 (addMetricJumps 1 mc)))) + | ( true, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) + | ( true, true) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) + end. + + Definition cost_stackalloc x mc := + match isReg x with + | false => (addMetricInstructions 2 (addMetricLoads 2 (addMetricStores 1 mc))) + | true => (addMetricInstructions 1 (addMetricLoads 1 mc)) + end. + + Definition cost_lit x mc := + match isReg x with + | false => (addMetricInstructions 9 (addMetricLoads 9 (addMetricStores 1 mc))) + | true => (addMetricInstructions 8 (addMetricLoads 8 mc)) + end. + + Definition cost_op x y z mc := + match (isReg x, isReg y, isReg z) with + | (false, false, false) => (addMetricInstructions 5 (addMetricLoads 7 (addMetricStores 1 mc))) + | (false, false, true) | (false, true, false) => (addMetricInstructions 4 (addMetricLoads 5 (addMetricStores 1 mc))) + | (false, true, true) => (addMetricInstructions 3 (addMetricLoads 3 (addMetricStores 1 mc))) + | ( true, false, false) => (addMetricInstructions 4 (addMetricLoads 6 mc)) + | ( true, false, true) | ( true, true, false) => (addMetricInstructions 3 (addMetricLoads 4 mc)) + | ( true, true, true) => (addMetricInstructions 2 (addMetricLoads 2 mc)) + end. + + Definition cost_set x y mc := + match (isReg x, isReg y) with + | (false, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricStores 1 mc))) + | (false, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricStores 1 mc))) + | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 mc)) + | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 mc)) + end. + + Definition cost_if x y mc := + match (isReg x, match y with | Some y' => isReg y' | None => true end) with + | (false, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) + | (false, true) | ( true, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) + | ( true, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) + end. + + Definition cost_loop_true x y mc := + match (isReg x, match y with | Some y' => isReg y' | None => true end) with + | (false, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) + | (false, true) | ( true, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) + | ( true, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) + end. + + Definition cost_loop_false x y mc := + match (isReg x, match y with | Some y' => isReg y' | None => true end) with + | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricJumps 1 mc))) + | (false, true) | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricJumps 1 mc))) + | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mc))) + end. + +End FlatImpExec. + +Definition isRegZ (var : Z) : bool := + Z.leb var 31. + +Definition isRegStr (var : String.string) : bool := + String.prefix "reg_" var. + +Ltac cost_unfold := + unfold cost_interact, cost_call_internal, cost_call_external, cost_load, + cost_store, cost_inlinetable, cost_stackalloc, cost_lit, cost_op, cost_set, + cost_if, cost_loop_true, cost_loop_false in *. + +Ltac cost_destr := + repeat match goal with + | x : compphase |- _ => destr x + | _ : context[if ?x then _ else _] |- _ => destr x; try discriminate + | |- context[if ?x then _ else _] => destr x; try discriminate + end; unfold EmptyMetricLog. + +Ltac cost_solve := cost_unfold; cost_destr; try solve_MetricLog. +Ltac cost_solve_piecewise := cost_unfold; cost_destr; try solve_MetricLog_piecewise. + +Definition metrics_additive f := forall mc, f mc - mc = f EmptyMetricLog. +#[global] Transparent metrics_additive. + +Ltac t := unfold metrics_additive; intros; cost_solve_piecewise. +Lemma cost_additive_interact : forall phase, metrics_additive (cost_interact phase). Proof. t. Qed. +Lemma cost_additive_call_internal : forall phase, metrics_additive (cost_call_internal phase). Proof. t. Qed. +Lemma cost_additive_call_external : forall phase, metrics_additive (cost_call_external phase). Proof. t. Qed. +Lemma cost_additive_load : forall varname isReg x a, metrics_additive (@cost_load varname isReg x a). Proof. t. Qed. +Lemma cost_additive_store : forall varname isReg x a, metrics_additive (@cost_store varname isReg x a). Proof. t. Qed. +Lemma cost_additive_inlinetable : forall varname isReg x a, metrics_additive (@cost_inlinetable varname isReg x a). Proof. t. Qed. +Lemma cost_additive_stackalloc : forall varname isReg x, metrics_additive (@cost_stackalloc varname isReg x). Proof. t. Qed. +Lemma cost_additive_lit : forall varname isReg x, metrics_additive (@cost_lit varname isReg x). Proof. t. Qed. +Lemma cost_additive_op : forall varname isReg x y z, metrics_additive (@cost_op varname isReg x y z). Proof. t. Qed. +Lemma cost_additive_set : forall varname isReg x y, metrics_additive (@cost_set varname isReg x y). Proof. t. Qed. +Lemma cost_additive_if : forall varname isReg x y, metrics_additive (@cost_if varname isReg x y). Proof. t. Qed. +Lemma cost_additive_loop_true : forall varname isReg x y, metrics_additive (@cost_loop_true varname isReg x y). Proof. t. Qed. +Lemma cost_additive_loop_false : forall varname isReg x y, metrics_additive (@cost_loop_false varname isReg x y). Proof. t. Qed. + +(* COQBUG: + these cannot have "in *", because this causes a crash if the environment contains a Semantics.ExtSpec. + Anomaly "Unable to handle arbitrary u+k <= v constraints." + Please report at http://coq.inria.fr/bugs/. + to use these in hypotheses, revert them first *) +Ltac cost_additive := repeat ( + rewrite cost_additive_interact || + rewrite cost_additive_call_internal || + rewrite cost_additive_call_external || + rewrite cost_additive_load || + rewrite cost_additive_store || + rewrite cost_additive_inlinetable || + rewrite cost_additive_stackalloc || + rewrite cost_additive_lit || + rewrite cost_additive_op || + rewrite cost_additive_set || + rewrite cost_additive_if || + rewrite cost_additive_loop_true || + rewrite cost_additive_loop_false + ). + +Ltac cost_hammer := cost_additive; try solve [eauto 3 with metric_arith | cost_solve]. diff --git a/bedrock2/src/bedrock2/MetricLoops.v b/bedrock2/src/bedrock2/MetricLoops.v index 723953561..28417652e 100644 --- a/bedrock2/src/bedrock2/MetricLoops.v +++ b/bedrock2/src/bedrock2/MetricLoops.v @@ -10,6 +10,7 @@ From bedrock2 Require Semantics. From bedrock2 Require Import MetricWeakestPrecondition MetricWeakestPreconditionProperties. Require Import bedrock2.MetricLogging. +Require Import bedrock2.MetricCosts. Section Loops. Context {width: Z} {BW: Bitwidth width} {word: word.word width} {mem: map.map word Byte.byte}. @@ -22,17 +23,19 @@ Section Loops. Context {fs : Semantics.env}. Let call := fs. + Local Notation UNK := String.EmptyString. + Lemma wp_while: forall e c t m l mc (post: _ -> _ -> _ -> _ -> Prop), (exists measure (lt:measure->measure->Prop) (inv:measure->Semantics.trace->mem->locals->MetricLog->Prop), Coq.Init.Wf.well_founded lt /\ (exists v, inv v t m l mc) /\ (forall v t m l mc, inv v t m l mc -> exists bv bmc, dexpr m l e mc (bv, bmc) /\ - (word.unsigned bv <> 0%Z -> cmd call c t m l bmc (fun t' m l mc => - exists v', inv v' t' m l (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) + (word.unsigned bv <> 0%Z -> cmd call c t m l bmc (fun t' m' l' mc' => + exists v', inv v' t' m' l' (cost_loop_true isRegStr UNK (Some UNK) mc') /\ lt v' v)) /\ (word.unsigned bv = 0%Z -> post t m l - (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 bmc)))))) -> + (cost_loop_false isRegStr UNK (Some UNK) bmc)))) -> cmd call (cmd.while e c) t m l mc post. Proof. intros. destruct H as (measure & lt & inv & Hwf & HInit & Hbody). @@ -61,11 +64,11 @@ Section Loops. exists brv brmc, expr m l e mc (eq (brv, brmc)) /\ (word.unsigned brv <> 0%Z -> cmd call c t m l brmc (fun t' m' l' mc' => exists v' g', - P v' g' t' m' l' (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc'))) /\ + P v' g' t' m' l' (cost_loop_true isRegStr UNK (Some UNK) mc') /\ lt v' v /\ (forall t'' m'' l'' mc'', Q v' g' t'' m'' l'' mc'' -> Q v g t'' m'' l'' mc''))) /\ (word.unsigned brv = 0%Z -> Q v g t m l - (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 brmc))))) + (cost_loop_false isRegStr UNK (Some UNK) brmc))) (Hpost: forall t m l mc, Q v0 g0 t m l mc -> post t m l mc) : cmd call (cmd.while e c) t m l mc post. Proof. @@ -97,11 +100,11 @@ Section Loops. exists brv brmc, expr m l e mc (eq (brv, brmc)) /\ (word.unsigned brv <> 0%Z -> cmd call c t m l brmc (fun t' m' l' mc' => exists v' g', - P v' g' t' m' l' (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc'))) /\ + P v' g' t' m' l' (cost_loop_true isRegStr UNK (Some UNK) mc') /\ lt v' v /\ (forall t'' m'' l'' mc'', Q v' g' t'' m'' l'' mc'' -> Q v g t'' m'' l'' mc''))) /\ (word.unsigned brv = 0%Z -> cmd call rest t m l - (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 brmc))) (Q v g))) + (cost_loop_false isRegStr UNK (Some UNK) brmc) (Q v g))) : cmd call (cmd.seq (cmd.while e c) rest) t m l mc (Q v0 g0). Proof. cbn. eapply tailrec_localsmap_1ghost with @@ -194,10 +197,10 @@ Section Loops. exists brv brmc, expr m l e mc (eq (Datatypes.pair brv brmc)) /\ (word.unsigned brv <> 0 -> cmd fs c t m l brmc (fun t m l mc => exists v', - invariant v' t m l (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) + invariant v' t m l (cost_loop_true isRegStr UNK (Some UNK) mc) /\ lt v' v)) /\ (word.unsigned brv = 0 -> post t m l - (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 brmc))))) + (cost_loop_false isRegStr UNK (Some UNK) brmc))) : cmd fs (cmd.while e c) t m l mc post. Proof. eapply wp_while. @@ -225,8 +228,8 @@ Section Loops. Markers.unique (Markers.left (tuple.existss (fun localstuple => enforce variables localstuple l /\ Markers.right (Markers.unique (exists v', - tuple.apply (invariant v' t m (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc)))) localstuple /\ lt v' v))))))) /\ - (word.unsigned brv = 0 -> post t m l (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 brmc)))))) + tuple.apply (invariant v' t m (cost_loop_true isRegStr UNK (Some UNK) mc)) localstuple /\ lt v' v))))))) /\ + (word.unsigned brv = 0 -> post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc)))) : cmd call (cmd.while e c) t m l mc post. Proof. eapply (while_localsmap (fun v t m l mc => @@ -272,11 +275,11 @@ Section Loops. (fun t' m' localsmap' mc' => Markers.unique (Markers.left (hlist.existss (fun l' => enforce variables l' localsmap' /\ Markers.right ( Markers.unique (Markers.left (hlist.existss (fun g' => exists v', - match tuple.apply (hlist.apply (spec v') g' t' m') l' (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc'))) with S' => + match tuple.apply (hlist.apply (spec v') g' t' m') l' (cost_loop_true isRegStr UNK (Some UNK) mc') with S' => S'.(1) /\ Markers.right ( lt v' v /\ forall T M, hlist.foralls (fun L => forall MC, tuple.apply (S'.(2) T M) L MC -> tuple.apply (S_.(2) T M) L MC)) end))))))))) /\ - (word.unsigned brv = 0%Z -> tuple.apply (S_.(2) t m) l (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 brmc)))))))end)))) + (word.unsigned brv = 0%Z -> tuple.apply (S_.(2) t m) l (cost_loop_false isRegStr UNK (Some UNK) brmc)))))end)))) (Hpost : match (tuple.apply (hlist.apply (spec v0) g0 t m) l0 mc).(2) with Q0 => forall t m mc, hlist.foralls (fun l => tuple.apply (Q0 t m) l mc -> post t m (reconstruct variables l) mc)end) , cmd call (cmd.while e c) t m localsmap mc post ). Proof. @@ -317,15 +320,11 @@ Section Loops. exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ (word.unsigned br <> 0%Z -> cmd call c t m l mc' (fun t' m' l' mc''=> exists v', - let S' := spec v' t' m' l' (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc''))) in let '(P', Q') := S' in + let S' := spec v' t' m' l' (cost_loop_true isRegStr UNK (Some UNK) mc'') in let '(P', Q') := S' in P' /\ lt v' v /\ forall T M L MC, Q' T M L MC -> Q T M L MC)) /\ - (word.unsigned br = 0%Z -> Q t m l (addMetricInstructions 1 - (addMetricLoads 1 - (addMetricJumps 1 mc'))))) + (word.unsigned br = 0%Z -> Q t m l (cost_loop_false isRegStr UNK (Some UNK) mc'))) (Hpost : forall t m l mc, Q0 t m l mc -> post t m l mc) : cmd call (cmd.while e c) t m l mc post. Proof. diff --git a/bedrock2/src/bedrock2/MetricSemantics.v b/bedrock2/src/bedrock2/MetricSemantics.v index 41936c903..bde80ac64 100644 --- a/bedrock2/src/bedrock2/MetricSemantics.v +++ b/bedrock2/src/bedrock2/MetricSemantics.v @@ -7,9 +7,12 @@ Require Import BinIntDef coqutil.Word.Interface coqutil.Word.Bitwidth. Require Export bedrock2.Memory. Require Import Coq.Lists.List. Require Import bedrock2.MetricLogging. +Require Import bedrock2.MetricCosts. Require Import bedrock2.Semantics. Require Import Coq.Lists.List. +Local Notation UNK := String.EmptyString. + Section semantics. Context {width: Z} {BW: Bitwidth width} {word: word.word width} {mem: map.map word byte}. Context {locals: map.map String.string word}. @@ -23,37 +26,31 @@ Section semantics. Local Notation "' x <- a | y ; f" := (match a with x => f | _ => y end) (right associativity, at level 70, x pattern). + (* TODO XXX possibly be a bit smarter about whether things are registers, + for tighter metrics bounds at bedrock2 level *) Fixpoint eval_expr (e : expr) (mc : metrics) : option (word * metrics) := match e with - | expr.literal v => Some (word.of_Z v, addMetricInstructions 8 - (addMetricLoads 8 mc)) + | expr.literal v => Some (word.of_Z v, cost_lit isRegStr UNK mc) | expr.var x => match map.get l x with - | Some v => Some (v, addMetricInstructions 1 - (addMetricLoads 2 mc)) + | Some v => Some (v, cost_set isRegStr UNK x mc) | None => None end | expr.inlinetable aSize t index => 'Some (index', mc') <- eval_expr index mc | None; 'Some v <- load aSize (map.of_list_word t) index' | None; - Some (v, (addMetricInstructions 3 - (addMetricLoads 4 - (addMetricJumps 1 mc')))) + Some (v, cost_inlinetable isRegStr UNK UNK mc') | expr.load aSize a => 'Some (a', mc') <- eval_expr a mc | None; 'Some v <- load aSize m a' | None; - Some (v, addMetricInstructions 1 - (addMetricLoads 2 mc')) + Some (v, cost_load isRegStr UNK UNK mc') | expr.op op e1 e2 => 'Some (v1, mc') <- eval_expr e1 mc | None; 'Some (v2, mc'') <- eval_expr e2 mc' | None; - Some (interp_binop op v1 v2, addMetricInstructions 2 - (addMetricLoads 2 mc'')) + Some (interp_binop op v1 v2, cost_op isRegStr UNK UNK UNK mc'') | expr.ite c e1 e2 => 'Some (vc, mc') <- eval_expr c mc | None; eval_expr (if word.eqb vc (word.of_Z 0) then e2 else e1) - (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc'))) + (cost_if isRegStr UNK (Some UNK) mc') end. Fixpoint eval_call_args (arges : list expr) (mc : metrics) := @@ -88,8 +85,7 @@ Module exec. Section WithParams. | set x e t m l mc post v mc' (_ : eval_expr m l e mc = Some (v, mc')) - (_ : post t m (map.put l x v) (addMetricInstructions 1 - (addMetricLoads 1 mc'))) + (_ : post t m (map.put l x v) (cost_set isRegStr x UNK mc')) : exec (cmd.set x e) t m l mc post | unset x t m l mc post @@ -100,9 +96,7 @@ Module exec. Section WithParams. a mc' (_ : eval_expr m l ea mc = Some (a, mc')) v mc'' (_ : eval_expr m l ev mc' = Some (v, mc'')) m' (_ : store sz m a v = Some m') - (_ : post t m' l (addMetricInstructions 1 - (addMetricLoads 1 - (addMetricStores 1 mc'')))) + (_ : post t m' l (cost_store isRegStr UNK UNK mc'')) : exec (cmd.store sz ea ev) t m l mc post | stackalloc x n body t mSmall l mc post @@ -110,7 +104,7 @@ Module exec. Section WithParams. (_ : forall a mStack mCombined, anybytes a n mStack -> map.split mCombined mSmall mStack -> - exec body t mCombined (map.put l x a) (addMetricInstructions 1 (addMetricLoads 1 mc)) + exec body t mCombined (map.put l x a) (cost_stackalloc isRegStr x mc) (fun t' mCombined' l' mc' => exists mSmall' mStack', anybytes a n mStack' /\ @@ -120,17 +114,13 @@ Module exec. Section WithParams. | if_true t m l mc e c1 c2 post v mc' (_ : eval_expr m l e mc = Some (v, mc')) (_ : word.unsigned v <> 0) - (_ : exec c1 t m l (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc'))) post) + (_ : exec c1 t m l (cost_if isRegStr UNK (Some UNK) mc') post) : exec (cmd.cond e c1 c2) t m l mc post | if_false e c1 c2 t m l mc post v mc' (_ : eval_expr m l e mc = Some (v, mc')) (_ : word.unsigned v = 0) - (_ : exec c2 t m l (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc'))) post) + (_ : exec c2 t m l (cost_if isRegStr UNK (Some UNK) mc') post) : exec (cmd.cond e c1 c2) t m l mc post | seq c1 c2 t m l mc post @@ -141,9 +131,7 @@ Module exec. Section WithParams. t m l mc post v mc' (_ : eval_expr m l e mc = Some (v, mc')) (_ : word.unsigned v = 0) - (_ : post t m l (addMetricInstructions 1 - (addMetricLoads 1 - (addMetricJumps 1 mc')))) + (_ : post t m l (cost_loop_false isRegStr UNK (Some UNK) mc')) : exec (cmd.while e c) t m l mc post | while_true e c t m l mc post @@ -151,20 +139,18 @@ Module exec. Section WithParams. (_ : word.unsigned v <> 0) mid (_ : exec c t m l mc' mid) (_ : forall t' m' l' mc'', mid t' m' l' mc'' -> - exec (cmd.while e c) t' m' l' (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc''))) post) + exec (cmd.while e c) t' m' l' (cost_loop_true isRegStr UNK (Some UNK) mc'') post) : exec (cmd.while e c) t m l mc post | call binds fname arges t m l mc post params rets fbody (_ : map.get e fname = Some (params, rets, fbody)) args mc' (_ : eval_call_args m l arges mc = Some (args, mc')) lf (_ : map.of_list_zip params args = Some lf) - mid (_ : exec fbody t m lf (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc')))) mid) + mid (_ : exec fbody t m lf (cost_call_internal PreSpill mc') mid) (_ : forall t' m' st1 mc'', mid t' m' st1 mc'' -> exists retvs, map.getmany_of_list st1 rets = Some retvs /\ exists l', map.putmany_of_list_zip binds retvs l = Some l' /\ - post t' m' l' (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc''))))) + post t' m' l' (cost_call_external PreSpill mc'')) : exec (cmd.call binds fname arges) t m l mc post | interact binds action arges t m l mc post @@ -175,9 +161,7 @@ Module exec. Section WithParams. exists l', map.putmany_of_list_zip binds resvals l = Some l' /\ forall m', map.split m' mKeep mReceive -> post (cons ((mGive, action, args), (mReceive, resvals)) t) m' l' - (addMetricInstructions 1 - (addMetricStores 1 - (addMetricLoads 2 mc')))) + (cost_interact PreSpill mc')) : exec (cmd.interact binds action arges) t m l mc post . diff --git a/bedrock2/src/bedrock2/MetricWeakestPrecondition.v b/bedrock2/src/bedrock2/MetricWeakestPrecondition.v index 63570bd97..2e2e906e6 100644 --- a/bedrock2/src/bedrock2/MetricWeakestPrecondition.v +++ b/bedrock2/src/bedrock2/MetricWeakestPrecondition.v @@ -2,6 +2,7 @@ Require Import coqutil.Macros.subst coqutil.Macros.unique coqutil.Map.Interface Require Import Coq.ZArith.BinIntDef coqutil.Word.Interface coqutil.Word.Bitwidth. Require Import coqutil.dlet bedrock2.Syntax bedrock2.Semantics. Require Import bedrock2.MetricLogging. +Require Import bedrock2.MetricCosts. Require Import bedrock2.MetricSemantics. Section WeakestPrecondition. @@ -11,13 +12,15 @@ Section WeakestPrecondition. Implicit Types (t : trace) (m : mem) (l : locals). Local Notation metrics := MetricLog. - + Local Notation UNK := String.EmptyString. + + (* TODO XXX address inconsistency in where metrics are added *) Definition literal v mc (post : (word * metrics) -> Prop) : Prop := - dlet! v := word.of_Z v in post (v, addMetricInstructions 8 (addMetricLoads 8 mc)). + dlet! v := word.of_Z v in post (v, cost_lit isRegStr UNK mc). Definition get (l : locals) (x : String.string) mc (post : (word * metrics) -> Prop) : Prop := - exists v, map.get l x = Some v /\ post (v, addMetricInstructions 1 (addMetricLoads 2 mc)). + exists v, map.get l x = Some v /\ post (v, cost_set isRegStr UNK x mc). Definition load s m a mc (post: (word * metrics) -> Prop) : Prop := - exists v, load s m a = Some v /\ post(v, addMetricInstructions 1 (addMetricLoads 2 mc)). + exists v, load s m a = Some v /\ post (v, mc). Definition store sz m a v post := exists m', store sz m a v = Some m' /\ post m'. @@ -32,17 +35,15 @@ Section WeakestPrecondition. | expr.op op e1 e2 => rec e1 mc (fun '(v1, mc') => rec e2 mc' (fun '(v2, mc'') => - post (interp_binop op v1 v2, addMetricInstructions 2 (addMetricLoads 2 mc'')))) + post (interp_binop op v1 v2, cost_op isRegStr UNK UNK UNK mc''))) | expr.load s e => rec e mc (fun '(a, mc') => - load s m a mc' post) + load s m a (cost_load isRegStr UNK UNK mc') post) | expr.inlinetable s t e => rec e mc (fun '(a, mc') => - load s (map.of_list_word t) a (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc'))) post) + load s (map.of_list_word t) a (cost_inlinetable isRegStr UNK UNK mc') post) | expr.ite c e1 e2 => - rec c mc (fun '(b, mc') => rec (if word.eqb b (word.of_Z 0) then e2 else e1) (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc'))) post) + rec c mc (fun '(b, mc') => rec (if word.eqb b (word.of_Z 0) then e2 else e1) (cost_if isRegStr UNK (Some UNK) mc') post) end. Fixpoint expr e := expr_body expr e. End WithMemAndLocals. @@ -77,7 +78,7 @@ Section WeakestPrecondition. | cmd.set x ev => exists v mc', dexpr m l ev mc (v, mc') /\ dlet! l := map.put l x v in - post t m l (addMetricInstructions 1 (addMetricLoads 1 mc')) + post t m l (cost_set isRegStr x UNK mc') | cmd.unset x => dlet! l := map.remove l x in post t m l mc @@ -85,43 +86,41 @@ Section WeakestPrecondition. exists a mc', dexpr m l ea mc (a, mc') /\ exists v mc'', dexpr m l ev mc (v, mc'') /\ store sz m a v (fun m => - post t m l (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc'')))) + post t m l (cost_store isRegStr UNK UNK mc'')) | cmd.stackalloc x n c => Z.modulo n (bytes_per_word width) = 0 /\ forall a mStack mCombined, anybytes a n mStack -> map.split mCombined m mStack -> dlet! l := map.put l x a in - rec c t mCombined l (addMetricInstructions 1 (addMetricLoads 1 mc)) + rec c t mCombined l (cost_stackalloc isRegStr x mc) (fun t' mCombined' l' mc' => exists m' mStack', anybytes a n mStack' /\ map.split mCombined' m' mStack' /\ post t' m' l' mc') - | cmd.cond br ct cf => + | cmd.cond br ct cf => exists v mc', dexpr m l br mc (v, mc') /\ - dlet! mc'' := addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc')) in + dlet! mc'' := cost_if isRegStr UNK (Some UNK) mc' in (word.unsigned v <> 0%Z -> rec ct t m l mc'' post) /\ - (word.unsigned v = 0%Z -> rec cf t m l mc'' post) + (word.unsigned v = 0%Z -> rec cf t m l mc'' post) | cmd.seq c1 c2 => rec c1 t m l mc (fun t m l mc => rec c2 t m l mc post) | cmd.while _ _ => MetricSemantics.exec e c t m l mc post | cmd.call binds fname arges => exists args mc', dexprs m l arges mc (args, mc') /\ MetricSemantics.call e fname t m args mc (fun t m rets mc'' => - exists l', map.putmany_of_list_zip binds rets l = Some l' /\ post t m l' (addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc''))))) + exists l', map.putmany_of_list_zip binds rets l = Some l' /\ post t m l' (cost_call_internal PreSpill mc'')) | cmd.interact binds action arges => - exists args mc', dexprs m l arges mc (args, mc') /\ + exists args mc', dexprs m l arges mc (args, mc') /\ exists mKeep mGive, map.split m mKeep mGive /\ ext_spec t mGive action args (fun mReceive rets => exists l', map.putmany_of_list_zip binds rets l = Some l' /\ forall m', map.split m' mKeep mReceive -> - post (cons ((mGive, action, args), (mReceive, rets)) t) m' l' (addMetricInstructions 1 - (addMetricStores 1 - (addMetricLoads 2 mc')))) + post (cons ((mGive, action, args), (mReceive, rets)) t) m' l' (cost_interact PreSpill mc')) end. Fixpoint cmd c := cmd_body cmd c. End WithFunctions. - + Definition func call '(innames, outnames, c) (t : trace) (m : mem) (args : list word) (mc : metrics) (post : trace -> mem -> list word -> metrics -> Prop) := exists l, map.of_list_zip innames args = Some l /\ cmd call c t m l mc (fun t m l mc => diff --git a/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v b/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v index 4a7fdcefb..9248a2973 100644 --- a/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v +++ b/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v @@ -175,12 +175,7 @@ locals_ok mem mem_ok width word word_ok. ind_on Syntax.expr; t. { destruct H. destruct H. eexists. eexists. rewrite H. eauto. } { eapply IHe in H; t. cbv [MetricWeakestPrecondition.load] in H0; t. rewrite H. rewrite H0. eauto. } { eapply IHe in H; t. cbv [MetricWeakestPrecondition.load] in H0; t. rewrite H. rewrite H0. - (* now we prove that 1+2=3 and 2+2=4 *) - eexists. eexists. split; eauto. - MetricLogging.unfold_MetricLog. MetricLogging.simpl_MetricLog. - repeat rewrite <- BinInt.Z.add_assoc in H1. - simpl in H1. - eauto. } + eexists. eexists. split; eauto. } { eapply IHe1 in H; t. eapply IHe2 in H0; t. rewrite H, H0; eauto. } { eapply IHe1 in H; t. rewrite H. Tactics.destruct_one_match. { apply IHe3; t. } @@ -207,9 +202,7 @@ locals_ok mem mem_ok width word word_ok. 2: { eapply IHe. rewrite E. reflexivity. } intros (addr, oldmc) ?. apply pair_equal_spec in H0; destruct H0. subst r m0. unfold MetricWeakestPrecondition.load. eexists; split; eauto. - apply Option.eq_of_eq_Some in H; auto. rewrite <- H. - apply pair_equal_spec; split; auto. - destruct oldmc. applyAddMetricsGoal. reflexivity. + apply Option.eq_of_eq_Some in H; auto. - repeat (destruct_one_match_hyp; try discriminate; []). eapply Proper_expr. 2: { eapply IHe1. rewrite E. reflexivity. } diff --git a/compiler/src/compiler/FlatImp.v b/compiler/src/compiler/FlatImp.v index 7c5c03f40..d5c558b76 100644 --- a/compiler/src/compiler/FlatImp.v +++ b/compiler/src/compiler/FlatImp.v @@ -2,6 +2,7 @@ Require Import Coq.Bool.Bool. Require Import Coq.ZArith.ZArith. Require Import Coq.Lists.List. Import ListNotations. Require Import bedrock2.MetricLogging. +Require Import bedrock2.MetricCosts. Require Import coqutil.Macros.unique. Require Import bedrock2.Memory. Require Import compiler.util.Common. @@ -26,10 +27,6 @@ Inductive bbinop: Set := | BLtu | BGeu. -Inductive compphase: Type := -| PreSpill -| PostSpill. - Section Syntax. Context {varname: Type}. @@ -301,123 +298,27 @@ Module exec. end. (* Helper functions for computing costs of instructions *) - Definition cost_SInteract mc := - match phase with - | PreSpill => (addMetricInstructions 100 (addMetricJumps 100 (addMetricStores 100 (addMetricLoads 100 mc)))) - | PostSpill => (addMetricInstructions 50 (addMetricJumps 50 (addMetricStores 50 (addMetricLoads 50 mc)))) - end. - - Definition cost_SCall_internal mc := - match phase with - | PreSpill => addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc))) - | PostSpill => addMetricInstructions 50 (addMetricJumps 50 (addMetricLoads 50 (addMetricStores 50 mc))) - end. - - Definition cost_SCall_external mc := - match phase with - | PreSpill => addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc))) - | PostSpill => addMetricInstructions 50 (addMetricJumps 50 (addMetricLoads 50 (addMetricStores 50 mc))) - end. - - (* TODO think about a non-fixed bound on the cost of function preamble and postamble *) - - Definition cost_SLoad x a mc := - match (isReg x, isReg a) with - | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricStores 1 mc))) - | (false, true) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) - | ( true, false) => (addMetricInstructions 2 (addMetricLoads 4 mc)) - | ( true, true) => (addMetricInstructions 1 (addMetricLoads 2 mc)) - end. - - Definition cost_SStore a v mc := - match (isReg a, isReg v) with - | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricStores 1 mc))) - | (false, true) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) - | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) - | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc))) - end. - - Definition cost_SInlinetable x i mc := - match (isReg x, isReg i) with - | (false, false) => (addMetricInstructions 5 (addMetricLoads 7 (addMetricStores 1 (addMetricJumps 1 mc)))) - | (false, true) => (addMetricInstructions 4 (addMetricLoads 5 (addMetricStores 1 (addMetricJumps 1 mc)))) - | ( true, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) - | ( true, true) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) - end. - - Definition cost_SStackalloc x mc := - match isReg x with - | false => (addMetricInstructions 2 (addMetricLoads 2 (addMetricStores 1 mc))) - | true => (addMetricInstructions 1 (addMetricLoads 1 mc)) - end. - - Definition cost_SLit x mc := - match isReg x with - | false => (addMetricInstructions 9 (addMetricLoads 9 (addMetricStores 1 mc))) - | true => (addMetricInstructions 8 (addMetricLoads 8 mc)) - end. Definition cost_SOp x y z mc := - match (isReg x, isReg y, (match z with | Var vo => isReg vo | Const _ => true end)) with - | (false, false, false) => (addMetricInstructions 5 (addMetricLoads 7 (addMetricStores 1 mc))) - | (false, false, true) | (false, true, false) => (addMetricInstructions 4 (addMetricLoads 5 (addMetricStores 1 mc))) - | (false, true, true) => (addMetricInstructions 3 (addMetricLoads 3 (addMetricStores 1 mc))) - | ( true, false, false) => (addMetricInstructions 4 (addMetricLoads 6 mc)) - | ( true, false, true) | ( true, true, false) => (addMetricInstructions 3 (addMetricLoads 4 mc)) - | ( true, true, true) => (addMetricInstructions 2 (addMetricLoads 2 mc)) - end. - - Definition cost_SSet x y mc := - match (isReg x, isReg y) with - | (false, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricStores 1 mc))) - | (false, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricStores 1 mc))) - | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 mc)) - | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 mc)) - end. + cost_op (fun v => match v with | Var vo => isReg vo | Const _ => true end) + (Var x) (Var y) z mc. Definition cost_SIf bcond mc := match bcond with - | CondBinary _ x y => - match (isReg x, isReg y) with - | (false, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) - | (false, true) | ( true, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) - | ( true, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) - end - | CondNez x => - match isReg x with - | false => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) - | true => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) - end + | CondBinary _ x y => cost_if isReg x (Some y) mc + | CondNez x => cost_if isReg x None mc end. Definition cost_SLoop_true bcond mc := match bcond with - | CondBinary _ x y => - match (isReg x, isReg y) with - | (false, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) - | (false, true) | ( true, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) - | ( true, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) - end - | CondNez x => - match isReg x with - | false => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) - | true => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) - end + | CondBinary _ x y => cost_loop_true isReg x (Some y) mc + | CondNez x => cost_loop_true isReg x None mc end. Definition cost_SLoop_false bcond mc := match bcond with - | CondBinary _ x y => - match (isReg x, isReg y) with - | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricJumps 1 mc))) - | (false, true) | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricJumps 1 mc))) - | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mc))) - end - | CondNez x => - match isReg x with - | false => (addMetricInstructions 2 (addMetricLoads 3 (addMetricJumps 1 mc))) - | true => (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mc))) - end + | CondBinary _ x y => cost_loop_false isReg x (Some y) mc + | CondNez x => cost_loop_false isReg x None mc end. (* alternative semantics which allow non-determinism *) @@ -435,37 +336,37 @@ Module exec. exists l', map.putmany_of_list_zip resvars resvals l = Some l' /\ forall m', map.split m' mKeep mReceive -> post (((mGive, action, argvals), (mReceive, resvals)) :: t) m' l' - (cost_SInteract mc)) -> + (cost_interact phase mc)) -> exec (SInteract resvars action argvars) t m l mc post | call: forall t m l mc binds fname args params rets fbody argvs st0 post outcome, map.get e fname = Some (params, rets, fbody) -> map.getmany_of_list l args = Some argvs -> map.putmany_of_list_zip params argvs map.empty = Some st0 -> - exec fbody t m st0 (cost_SCall_internal mc) outcome -> + exec fbody t m st0 (cost_call_internal phase mc) outcome -> (forall t' m' mc' st1, outcome t' m' st1 mc' -> exists retvs l', map.getmany_of_list st1 rets = Some retvs /\ map.putmany_of_list_zip binds retvs l = Some l' /\ - post t' m' l' (cost_SCall_external mc')) -> + post t' m' l' (cost_call_external phase mc')) -> exec (SCall binds fname args) t m l mc post | load: forall t m l mc sz x a o v addr post, map.get l a = Some addr -> load sz m (word.add addr (word.of_Z o)) = Some v -> - post t m (map.put l x v) (cost_SLoad x a mc)-> + post t m (map.put l x v) (cost_load isReg x a mc)-> exec (SLoad sz x a o) t m l mc post | store: forall t m m' mc l sz a o addr v val post, map.get l a = Some addr -> map.get l v = Some val -> store sz m (word.add addr (word.of_Z o)) val = Some m' -> - post t m' l (cost_SStore a v mc) -> + post t m' l (cost_store isReg a v mc) -> exec (SStore sz a v o) t m l mc post | inlinetable: forall sz x table i v index t m l mc post, (* compiled riscv code uses x as a tmp register and this shouldn't overwrite i *) x <> i -> map.get l i = Some index -> load sz (map.of_list_word table) index = Some v -> - post t m (map.put l x v) (cost_SInlinetable x i mc) -> + post t m (map.put l x v) (cost_inlinetable isReg x i mc) -> exec (SInlinetable sz x table i) t m l mc post | stackalloc: forall t mSmall l mc x n body post, n mod (bytes_per_word width) = 0 -> @@ -477,10 +378,10 @@ Module exec. exists mSmall' mStack', anybytes a n mStack' /\ map.split mCombined' mSmall' mStack' /\ - post t' mSmall' l' (cost_SStackalloc x mc'))) -> + post t' mSmall' l' (cost_stackalloc isReg x mc'))) -> exec (SStackalloc x n body) t mSmall l mc post | lit: forall t m l mc x v post, - post t m (map.put l x (word.of_Z v)) (cost_SLit x mc) -> + post t m (map.put l x (word.of_Z v)) (cost_lit isReg x mc) -> exec (SLit x v) t m l mc post | op: forall t m l mc x op y y' z z' post, map.get l y = Some y' -> @@ -489,7 +390,7 @@ Module exec. exec (SOp x op y z) t m l mc post | set: forall t m l mc x y y' post, map.get l y = Some y' -> - post t m (map.put l x y') (cost_SSet x y mc) -> + post t m (map.put l x y') (cost_set isReg x y mc) -> exec (SSet x y) t m l mc post | if_true: forall t m l mc cond bThen bElse post, eval_bcond l cond = Some true -> @@ -550,12 +451,12 @@ Module exec. map.get e fname = Some (params, rets, fbody) -> map.getmany_of_list l args = Some argvs -> map.putmany_of_list_zip params argvs map.empty = Some st -> - exec fbody t m st (cost_SCall_internal mc) + exec fbody t m st (cost_call_internal phase mc) (fun t' m' st' mc' => exists retvs l', map.getmany_of_list st' rets = Some retvs /\ map.putmany_of_list_zip binds retvs l = Some l' /\ - post t' m' l' (cost_SCall_external mc')) -> + post t' m' l' (cost_call_external phase mc')) -> exec (SCall binds fname args) t m l mc post. Proof. intros. eapply call; try eassumption. @@ -792,8 +693,33 @@ Section FlatImp2. End FlatImp2. -Definition isRegZ (var : Z) : bool := - Z.leb var 31. +(* various helper tactics extending the ones from MetricCosts *) + +Ltac scost_unfold := + unfold exec.cost_SOp, exec.cost_SIf, exec.cost_SLoop_true, exec.cost_SLoop_false in *; cost_unfold. + +Ltac scost_destr := + repeat match goal with + | x : operand |- _ => destr x + | x : bbinop _ |- _ => destr x + | x : bcond _ |- _ => destr x + | _ => cost_destr + end. + +Ltac scost_solve := scost_unfold; scost_destr; try solve_MetricLog. +Ltac scost_solve_piecewise := scost_unfold; scost_destr; try solve_MetricLog_piecewise. + +Ltac t := unfold metrics_additive; intros; scost_solve_piecewise. +Lemma scost_additive_op : forall varname isReg x y z, metrics_additive (@exec.cost_SOp varname isReg x y z). Proof. t. Qed. +Lemma scost_additive_if : forall varname isReg x, metrics_additive (@exec.cost_SIf varname isReg x). Proof. t. Qed. +Lemma scost_additive_loop_true : forall varname isReg x, metrics_additive (@exec.cost_SLoop_true varname isReg x). Proof. t. Qed. +Lemma scost_additive_loop_false : forall varname isReg x, metrics_additive (@exec.cost_SLoop_false varname isReg x). Proof. t. Qed. + +Ltac scost_additive := repeat ( + rewrite scost_additive_op || + rewrite scost_additive_if || + rewrite scost_additive_loop_true || + rewrite scost_additive_loop_false + ); cost_additive. -Definition isRegStr (var : String.string) : bool := - String.prefix "reg_" var. +Ltac scost_hammer := scost_additive; try solve [eauto 3 with metric_arith | scost_solve]. diff --git a/compiler/src/compiler/FlatToRiscvCommon.v b/compiler/src/compiler/FlatToRiscvCommon.v index 017944581..7cdb7c49b 100644 --- a/compiler/src/compiler/FlatToRiscvCommon.v +++ b/compiler/src/compiler/FlatToRiscvCommon.v @@ -45,6 +45,7 @@ Require Import compiler.RunInstruction. Require Import compiler.DivisibleBy4. Require Import compiler.MetricsToRiscv. Require Export compiler.regs_initialized. +Require Import bedrock2.MetricCosts. Require Import coqutil.Word.Interface. Local Hint Mode Word.Interface.word - : typeclass_instances. diff --git a/compiler/src/compiler/FlatToRiscvFunctions.v b/compiler/src/compiler/FlatToRiscvFunctions.v index 8a276e2d6..7900c1edc 100644 --- a/compiler/src/compiler/FlatToRiscvFunctions.v +++ b/compiler/src/compiler/FlatToRiscvFunctions.v @@ -32,6 +32,7 @@ Require Import coqutil.Word.DebugWordEq. Require Import compiler.MemoryLayout. Require Import coqutil.Map.MapEauto. Require Import compiler.Registers. +Require Import bedrock2.MetricCosts. Import MetricLogging. @@ -204,12 +205,12 @@ Section Proofs. Qed. Lemma valid_FlatImp_var_isRegZ : forall x, - valid_FlatImp_var x -> isRegZ x = true. + valid_FlatImp_var x -> isRegZ x = true. Proof. - unfold valid_FlatImp_var, isRegZ; blia. - Qed. + unfold valid_FlatImp_var, isRegZ; blia. + Qed. Hint Resolve valid_FlatImp_var_isRegZ. - + Ltac run1done := apply runsToDone; simpl_MetricRiscvMachine_get_set; @@ -219,13 +220,11 @@ Section Proofs. end; ssplit; simpl_word_exprs word_ok; match goal with | |- _ => solve_word_eq word_ok - | |- (_ <= _)%metricsL => - unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, - exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, - exec.cost_SLoop_true, exec.cost_SLoop_false in *; + | |- (_ <= _)%metricsL => + scost_unfold; repeat match goal with | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end; + end; MetricsToRiscv.solve_MetricLog | |- iff1 ?x ?x => reflexivity (* `exists stack_trash frame_trash, ...` from goodMachine *) @@ -1617,7 +1616,7 @@ Section Proofs. forget (Datatypes.length binds) as binds_count. subst binds. eapply runsTo_weaken. - 1:{ + 1:{ match goal with | H: (binds_count <= 8)%nat |- _ => rename H into BC end. @@ -1635,7 +1634,7 @@ Section Proofs. split; eauto 8 with map_hints. split; eauto 8 with map_hints. (* cost_SCall constraint: cost_SCall_L + (1,1,1,0) <= cost_SCall_internal + cost_SCall_external *) - unfold exec.cost_SCall_internal, exec.cost_SCall_external, cost_SCall_L in *. + unfold cost_SCall_L in *. MetricsToRiscv.solve_MetricLog. - idtac "Case compile_stmt_correct/SLoad". @@ -1647,8 +1646,8 @@ Section Proofs. } inline_iff1. run1det. clear H0. (* <-- TODO this should not be needed *) run1done. - - + + - idtac "Case compile_stmt_correct/SStore". inline_iff1. simpl_MetricRiscvMachine_get_set. @@ -1859,11 +1858,9 @@ Section Proofs. } run1done. cbn. - unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, - exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, - exec.cost_SLoop_true, exec.cost_SLoop_false in *. + cost_unfold. repeat match goal with - | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; try rewrite H in * end. remember (updateMetricsForLiteral v initialL_metrics) as finalMetrics; symmetry in HeqfinalMetrics; @@ -1880,7 +1877,7 @@ Section Proofs. match goal with | op: Syntax.bopname.bopname |- _ => destr op end. - all: match goal with + all: scost_unfold; match goal with | y: operand, H: context[Syntax.bopname.eq] |- _ => destr y; simpl in *; [ run1det; run1det; run1done; @@ -1952,24 +1949,21 @@ Section Proofs. simpl_MetricRiscvMachine_get_set. intros. destruct_RiscvMachine mid. fwd. run1done. - destruct cond eqn:E; - unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, - exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, - exec.cost_SLoop_true, exec.cost_SLoop_false in *; - unfold ForallVars_bcond in *; + destruct cond eqn:E; scost_unfold; + unfold ForallVars_bcond in *; repeat match goal with | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H end; repeat match goal with | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end; - try MetricsToRiscv.solve_MetricLog. + end; + try MetricsToRiscv.solve_MetricLog. destruct H6p0. repeat match goal with | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * end. - MetricsToRiscv.solve_MetricLog. - + MetricsToRiscv.solve_MetricLog. + - idtac "Case compile_stmt_correct/SIf/Else". (* execute branch instruction, which will jump over then-branch *) eapply runsToStep. @@ -1994,24 +1988,21 @@ Section Proofs. * (* at end of else-branch, i.e. also at end of if-then-else, just prove that computed post satisfies required post *) simpl. intros. destruct_RiscvMachine middle. fwd. subst. run1done. - destruct cond eqn:E; - unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, - exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, - exec.cost_SLoop_true, exec.cost_SLoop_false in *; - unfold ForallVars_bcond in *; + destruct cond eqn:E; scost_unfold; + unfold ForallVars_bcond in *; repeat match goal with | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H end; repeat match goal with | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end; - try MetricsToRiscv.solve_MetricLog. + end; + try MetricsToRiscv.solve_MetricLog. destruct H6p0. repeat match goal with | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * end. - MetricsToRiscv.solve_MetricLog. - + MetricsToRiscv.solve_MetricLog. + - idtac "Case compile_stmt_correct/SLoop". match goal with | H: context[FlatImpConstraints.uses_standard_arg_regs body1 -> _] |- _ => rename H into IH1 @@ -2083,24 +2074,21 @@ Section Proofs. } (* at end of loop, just prove that computed post satisfies required post *) simpl. intros. destruct_RiscvMachine middle. fwd. run1done. - destruct cond eqn:blah; - unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, - exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, - exec.cost_SLoop_true, exec.cost_SLoop_false in *; - unfold ForallVars_bcond in *; + destruct cond eqn:blah; scost_unfold; + unfold ForallVars_bcond in *; repeat match goal with | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H end; repeat match goal with | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end; - try MetricsToRiscv.solve_MetricLog. + end; + try MetricsToRiscv.solve_MetricLog. destruct H11p0. repeat match goal with | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * end. MetricsToRiscv.solve_MetricLog. - + * (* false: done, jump over body2 *) eapply runsToStep. { eapply compile_bcond_by_inverting_correct with (l := lH') (b := false); @@ -2109,18 +2097,15 @@ Section Proofs. } simpl_MetricRiscvMachine_get_set. intros. destruct_RiscvMachine mid. fwd. run1done. - destruct cond eqn:blah; - unfold exec.cost_SLoad, exec.cost_SStore, exec.cost_SInlinetable, exec.cost_SStackalloc, - exec.cost_SLit, exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, - exec.cost_SLoop_true, exec.cost_SLoop_false in *; - unfold ForallVars_bcond in *; + destruct cond eqn:blah; scost_unfold; + unfold ForallVars_bcond in *; repeat match goal with | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H end; repeat match goal with | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end; - try MetricsToRiscv.solve_MetricLog. + end; + try MetricsToRiscv.solve_MetricLog. destruct H11p0. repeat match goal with | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * @@ -2158,4 +2143,4 @@ Section Proofs. End Proofs. - + diff --git a/compiler/src/compiler/FlattenExpr.v b/compiler/src/compiler/FlattenExpr.v index 62a4f8e10..516c483e3 100644 --- a/compiler/src/compiler/FlattenExpr.v +++ b/compiler/src/compiler/FlattenExpr.v @@ -7,6 +7,7 @@ Require Import coqutil.Decidable. Require Import coqutil.Word.Bitwidth. Require Import bedrock2.Syntax. Require Import bedrock2.MetricLogging. +Require Import bedrock2.MetricCosts. Require Import bedrock2.Semantics bedrock2.MetricSemantics. Require Import coqutil.Macros.unique. Require Import Coq.Bool.Bool. @@ -343,8 +344,8 @@ Section FlattenExpr1. simpl (disjoint _ _) in *; map_solver locals_ok. - Local Notation exec := (FlatImp.exec FlatImp.PreSpill FlatImp.isRegStr). - + Local Notation exec := (FlatImp.exec PreSpill isRegStr). + Lemma seq_with_modVars: forall env t m (l: locals) mc s1 s2 mid post, exec env s1 t m l mc mid -> (forall t' m' l' mc', @@ -379,21 +380,19 @@ Section FlattenExpr1. Proof. induction e; intros *; intros F Ex U D Ev; simpl in *; simp. - (* Ltac solve_MetricLog ::= try solve [solve_MetricLog; idtac " Admitted metrics goal"; admit]. *) - - (* expr.literal *) - eapply @FlatImp.exec.lit; t_safe. Fail solve_MetricLog. admit. + eapply @FlatImp.exec.lit; t_safe. cost_hammer. - (* expr.var *) destruct oResVar; simp. - + eapply @FlatImp.exec.set; t_safe; [maps | try solve_MetricLog]. admit. - + eapply @FlatImp.exec.skip; t_safe. solve_MetricLog. + + eapply @FlatImp.exec.set; t_safe; [maps | cost_hammer]. + + eapply @FlatImp.exec.skip; t_safe. cost_hammer. - (* expr.load *) eapply @FlatImp.exec.seq. + eapply IHe; try eassumption. maps. + intros. simpl in *. simp. - eapply @FlatImp.exec.load; t_safe; rewrite ?word.add_0_r; try eassumption. try solve_MetricLog. admit. + eapply @FlatImp.exec.load; t_safe; rewrite ?word.add_0_r; try eassumption. cost_hammer. - (* expr.inlinetable *) repeat match goal with @@ -404,7 +403,7 @@ Section FlattenExpr1. + eapply IHe; try eassumption. 1: maps. set_solver; destr (String.eqb s0 x); subst; tauto. (* TODO improve set_solver? *) + intros. simpl in *. simp. - eapply @FlatImp.exec.inlinetable; t_safe; try eassumption. 2: try solve_MetricLog; admit. + eapply @FlatImp.exec.inlinetable; t_safe; try eassumption. 2: cost_hammer. apply_in_hyps flattenExpr_uses_Some_resVar. subst s0. intro C. subst s2. destruct oResVar. @@ -421,7 +420,7 @@ Section FlattenExpr1. clear IHe1 IHe2. pose_flatten_var_ineqs. set_solver. * intros. simpl in *. simp. clear IHe1 IHe2. - eapply @FlatImp.exec.op; t_safe; t_safe. 2: try solve_MetricLog; admit. + eapply @FlatImp.exec.op; t_safe; t_safe. 2: FlatImp.scost_hammer. eapply flattenExpr_valid_resVar in E1; simpl; maps. - (* expr.ite *) @@ -463,7 +462,7 @@ Section FlattenExpr1. eapply subset_union_rr. eapply subset_refl. } - ++ cbv beta. intros. simp. t_safe. 2: try solve_MetricLog; admit. + ++ cbv beta. intros. simp. t_safe. 2: FlatImp.scost_hammer. clear IHe1 IHe2 IHe3. apply_in_hyps flattenExpr_uses_Some_resVar. subst. assumption. * eapply FlatImp.exec.if_true. @@ -497,10 +496,10 @@ Section FlattenExpr1. eapply subset_union_rl. eapply subset_refl. } - ++ cbv beta. intros. simp. t_safe. 2: try solve_MetricLog; admit. + ++ cbv beta. intros. simp. t_safe. 2: FlatImp.scost_hammer. clear IHe1 IHe2 IHe3. apply_in_hyps flattenExpr_uses_Some_resVar. subst. assumption. - Admitted. + Qed. Goal True. idtac "FlattenExpr: flattenExpr_correct_aux done". Abort. Lemma flattenExpr_correct_with_modVars : forall e fenv oResVar ngs1 ngs2 resVar s t m lH lL initialMcH initialMcL finalMcH res, @@ -622,7 +621,7 @@ Section FlattenExpr1. | intros; simpl in *; simp; default_flattenBooleanExpr ]. - 2, 4, 6 : try solve_MetricLog. + 2, 4, 6 : cost_hammer. all: rewrite bool_to_word_to_bool_id; destruct_one_match; @@ -725,10 +724,10 @@ Section FlattenExpr1. specialize P with (1 := E). subst. maps. * maps. - * solve_MetricLog. - * solve_MetricLog. - * solve_MetricLog. - * solve_MetricLog. + * cost_hammer. + * cost_hammer. + * cost_hammer. + * cost_hammer. - (* exec.unset *) eapply @FlatImp.exec.skip. @@ -743,7 +742,7 @@ Section FlattenExpr1. * intros. simpl in *. simp. eapply @FlatImp.exec.store; rewrite ?word.add_0_r; try eassumption. { eapply flattenExpr_valid_resVar in E; maps. } - { repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all:admit. } + { repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all: cost_hammer. } - (* exec.stackalloc *) eapply @FlatImp.exec.stackalloc. 1: eassumption. @@ -751,7 +750,7 @@ Section FlattenExpr1. eapply @FlatImp.exec.weaken. { eapply IHexec; try reflexivity; try eassumption; maps. } { intros. simpl in *. simp. do 2 eexists. ssplit; try eassumption. - do 2 eexists. ssplit; try eassumption; try solve_MetricLog; try maps. admit. } + do 2 eexists. ssplit; try eassumption; try solve_MetricLog; try maps. cost_hammer. } - (* if_true *) eapply @FlatImp.exec.seq. @@ -768,7 +767,7 @@ Section FlattenExpr1. eapply @FlatImp.exec.weaken. { eapply IHexec; try reflexivity; try eassumption; maps. } { intros. simpl in *. simp. - repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all:admit. } + repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all: FlatImp.scost_hammer. } - (* if_false *) eapply @FlatImp.exec.seq. @@ -785,7 +784,7 @@ Section FlattenExpr1. eapply @FlatImp.exec.weaken. { eapply IHexec; try reflexivity; try eassumption; maps. } { intros. simpl in *. simp. - repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all:admit. } + repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all: FlatImp.scost_hammer. } - (* seq *) eapply seq_with_modVars. @@ -817,7 +816,7 @@ Section FlattenExpr1. | intros; simpl in *; simp .. ]. + maps. + congruence. - + repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all:admit. + + repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all: FlatImp.scost_hammer. + exfalso. match goal with | H: context [word.eqb _ _] |- _ => rewrite word.eqb_eq in H @@ -862,7 +861,7 @@ Section FlattenExpr1. 1,3: solve [maps]. pose proof (ExprImp.modVars_subset_allVars c). maps. * simpl. intros. simp. - repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all:admit. + repeat eexists; repeat (split || eassumption || solve_MetricLog); try maps. all: FlatImp.scost_hammer. - (* call *) unfold flattenCall in *. simp. @@ -906,9 +905,7 @@ Section FlattenExpr1. -- eassumption. -- do 2 eexists. ssplit; try eassumption. ++ simple eapply map.only_differ_putmany; eassumption. - ++ (* TODO cost_SCall into semantics *) - unfold FlatImp.exec.cost_SCall_internal, FlatImp.exec.cost_SCall_external in *. - solve_MetricLog. + ++ FlatImp.scost_hammer. - (* interact *) unfold flattenInteract in *. simp. @@ -929,8 +926,8 @@ Section FlattenExpr1. split; eauto. simple apply conj; [eassumption|]. split; [simple eapply map.only_differ_putmany; eassumption|]. - Fail solve_MetricLog. admit. - Admitted. + cost_hammer. + Qed. Goal True. idtac "FlattenExpr: flattenStmt_correct_aux done". Abort. Lemma flattenStmt_correct: forall eH eL sH sL lL t m mc post, diff --git a/compiler/src/compiler/LowerPipeline.v b/compiler/src/compiler/LowerPipeline.v index 1541fc50f..15be08c5f 100644 --- a/compiler/src/compiler/LowerPipeline.v +++ b/compiler/src/compiler/LowerPipeline.v @@ -21,6 +21,7 @@ Require Import compiler.FlatToRiscvDef. Require Import compiler.FlatToRiscvCommon. Require Import compiler.FlatToRiscvFunctions. Require Import bedrock2.MetricLogging. +Require Import bedrock2.MetricCosts. Require Import compiler.FitsStack. Require Import compiler.Registers. Require Import riscv.Utility.InstructionCoercions. diff --git a/compiler/src/compiler/MetricTactics.v b/compiler/src/compiler/MetricTactics.v deleted file mode 100644 index 3db8f2d9c..000000000 --- a/compiler/src/compiler/MetricTactics.v +++ /dev/null @@ -1,66 +0,0 @@ -Require Import compiler.util.Common. -Require Import bedrock2.MetricLogging. -Require Import compiler.FlatImp. - -Open Scope MetricH_scope. - -Definition metrics_additive f := forall mc, f mc - mc = f EmptyMetricLog. -#[global] Transparent metrics_additive. - -Ltac exec_cost_unfold := - unfold exec.cost_SInteract, exec.cost_SCall_internal, - exec.cost_SCall_external, exec.cost_SLoad, exec.cost_SStore, - exec.cost_SInlinetable, exec.cost_SStackalloc, exec.cost_SLit, exec.cost_SOp, - exec.cost_SSet, exec.cost_SIf, exec.cost_SLoop_true, exec.cost_SLoop_false in - *. - -Ltac exec_cost_destr := - repeat match goal with - | x : compphase |- _ => destr x - | x : operand |- _ => destr x - | x : bcond _ |- _ => destr x - | _ : context[if ?x then _ else _] |- _ => destr x - | |- context[if ?x then _ else _] => destr x - end; unfold EmptyMetricLog. - -Ltac exec_cost_solve := exec_cost_unfold; exec_cost_destr; try solve_MetricLog. -Ltac exec_cost_solve_piecewise := exec_cost_unfold; exec_cost_destr; try solve_MetricLog_piecewise. - -Ltac t := unfold metrics_additive; intros; exec_cost_solve_piecewise. - -Lemma exec_cost_additive_SInteract : forall phase, metrics_additive (exec.cost_SInteract phase). Proof. t. Qed. -Lemma exec_cost_additive_SCall_internal : forall phase, metrics_additive (exec.cost_SCall_internal phase). Proof. t. Qed. -Lemma exec_cost_additive_SCall_external : forall phase, metrics_additive (exec.cost_SCall_external phase). Proof. t. Qed. -Lemma exec_cost_additive_SLoad : forall varname isReg x a, metrics_additive (@exec.cost_SLoad varname isReg x a). Proof. t. Qed. -Lemma exec_cost_additive_SStore : forall varname isReg x a, metrics_additive (@exec.cost_SStore varname isReg x a). Proof. t. Qed. -Lemma exec_cost_additive_SInlinetable : forall varname isReg x a, metrics_additive (@exec.cost_SInlinetable varname isReg x a). Proof. t. Qed. -Lemma exec_cost_additive_SStackalloc : forall varname isReg x, metrics_additive (@exec.cost_SStackalloc varname isReg x). Proof. t. Qed. -Lemma exec_cost_additive_SLit : forall varname isReg x, metrics_additive (@exec.cost_SLit varname isReg x). Proof. t. Qed. -Lemma exec_cost_additive_SOp : forall varname isReg x y z, metrics_additive (@exec.cost_SOp varname isReg x y z). Proof. t. Qed. -Lemma exec_cost_additive_SSet : forall varname isReg x y, metrics_additive (@exec.cost_SSet varname isReg x y). Proof. t. Qed. -Lemma exec_cost_additive_SIf : forall varname isReg x, metrics_additive (@exec.cost_SIf varname isReg x). Proof. t. Qed. -Lemma exec_cost_additive_SLoop_true : forall varname isReg x, metrics_additive (@exec.cost_SLoop_true varname isReg x). Proof. t. Qed. -Lemma exec_cost_additive_SLoop_false : forall varname isReg x, metrics_additive (@exec.cost_SLoop_false varname isReg x). Proof. t. Qed. - -(* COQBUG: - these cannot have "in *", because this causes a crash if the environment contains a Semantics.ExtSpec. - Anomaly "Unable to handle arbitrary u+k <= v constraints." - Please report at http://coq.inria.fr/bugs/. - to use these in hypotheses, revert them first *) -Ltac exec_cost_additive := repeat ( - rewrite exec_cost_additive_SInteract || - rewrite exec_cost_additive_SCall_internal || - rewrite exec_cost_additive_SCall_external || - rewrite exec_cost_additive_SLoad || - rewrite exec_cost_additive_SStore || - rewrite exec_cost_additive_SInlinetable || - rewrite exec_cost_additive_SStackalloc || - rewrite exec_cost_additive_SLit || - rewrite exec_cost_additive_SOp || - rewrite exec_cost_additive_SSet || - rewrite exec_cost_additive_SIf || - rewrite exec_cost_additive_SLoop_true || - rewrite exec_cost_additive_SLoop_false - ). - -Ltac exec_cost_hammer := exec_cost_additive; try solve [eauto 3 with metric_arith | exec_cost_solve]. From 88c314d982aa0d3a3cdd66cf76883d6839539647 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Sat, 15 Jun 2024 03:11:37 -0400 Subject: [PATCH 36/62] MetricCosts migration: Spilling --- bedrock2/src/bedrock2/MetricCosts.v | 4 +- compiler/src/compiler/Spilling.v | 171 +++++++++++++--------------- 2 files changed, 81 insertions(+), 94 deletions(-) diff --git a/bedrock2/src/bedrock2/MetricCosts.v b/bedrock2/src/bedrock2/MetricCosts.v index a681aef8b..b6f56d6b3 100644 --- a/bedrock2/src/bedrock2/MetricCosts.v +++ b/bedrock2/src/bedrock2/MetricCosts.v @@ -121,14 +121,14 @@ Definition isRegStr (var : String.string) : bool := Ltac cost_unfold := unfold cost_interact, cost_call_internal, cost_call_external, cost_load, cost_store, cost_inlinetable, cost_stackalloc, cost_lit, cost_op, cost_set, - cost_if, cost_loop_true, cost_loop_false in *. + cost_if, cost_loop_true, cost_loop_false, EmptyMetricLog in *. Ltac cost_destr := repeat match goal with | x : compphase |- _ => destr x | _ : context[if ?x then _ else _] |- _ => destr x; try discriminate | |- context[if ?x then _ else _] => destr x; try discriminate - end; unfold EmptyMetricLog. + end. Ltac cost_solve := cost_unfold; cost_destr; try solve_MetricLog. Ltac cost_solve_piecewise := cost_unfold; cost_destr; try solve_MetricLog_piecewise. diff --git a/compiler/src/compiler/Spilling.v b/compiler/src/compiler/Spilling.v index 82239a7a4..b4458f648 100644 --- a/compiler/src/compiler/Spilling.v +++ b/compiler/src/compiler/Spilling.v @@ -11,6 +11,7 @@ Require Import compiler.Registers. Require Import compiler.SeparationLogic. Require Import compiler.SpillingMapGoals. Require Import bedrock2.MetricLogging. +Require Import bedrock2.MetricCosts. Require Import compiler.FlatImpConstraints. Require Import coqutil.Tactics.autoforward. Require Import coqutil.Tactics.fwd. @@ -22,7 +23,7 @@ Section Spilling. Notation stmt := (stmt Z). Notation execpre := (exec PreSpill isRegZ). Notation execpost := (exec PostSpill isRegZ). - + Definition zero := 0. Definition ra := 1. Definition sp := 2. @@ -129,9 +130,9 @@ Section Spilling. | SLit x n => SLit (ires_reg x) n;; save_ires_reg x - | SOp x op y oz => - load_iarg_reg 1 y;; - match oz with + | SOp x op y oz => + load_iarg_reg 1 y;; + match oz with | Var z => load_iarg_reg 2 z;; SOp (ires_reg x) op (iarg_reg 1 y) (Var (iarg_reg 2 z)) | Const _ => SOp (ires_reg x) op (iarg_reg 1 y) oz end;; save_ires_reg x @@ -355,7 +356,7 @@ Section Spilling. valid_vars_src m s -> valid_vars_tgt (spill_stmt s). Proof. - + unfold valid_vars_src, valid_vars_tgt. induction s; simpl; intros; repeat match goal with @@ -465,8 +466,8 @@ Section Spilling. Proof. intros. unfold load_iarg_reg, stack_loc, iarg_reg, related in *. fwd. - assert (isRegZ (9 + i) = true) by (unfold isRegZ; blia). - assert (isRegZ fp = true) by (unfold isRegZ; (assert (fp = 5) by auto); blia). + assert (isRegZ (9 + i) = true) by (unfold isRegZ; blia). + assert (isRegZ fp = true) by (unfold isRegZ; (assert (fp = 5) by auto); blia). destr (32 <=? r). - eapply exec.load. + eapply get_sep. ecancel_assumption. @@ -476,9 +477,9 @@ Section Spilling. eapply map.get_split_r. 1,3: eassumption. destr (map.get mp r); [exfalso|reflexivity]. specialize H0p2 with (1 := E0). blia. - + unfold exec.cost_SLoad. - assert (isRegZ r = false) by (unfold isRegZ; blia); rewrite H4 in H3. - unfold spill_tmp in H3. rewrite H0; rewrite H1. + + unfold cost_load. + assert (isRegZ r = false) by (unfold isRegZ; blia); rewrite H4 in H3. + unfold spill_tmp in H3. rewrite H0; rewrite H1. eapply H3. repeat match goal with | |- exists _, _ => eexists @@ -496,7 +497,7 @@ Section Spilling. destr (map.get lStack r); [exfalso|reflexivity]. specialize H0p3 with (1 := E0). blia. } - assert (isRegZ r = true) by (unfold isRegZ; blia); rewrite H4 in H3. + assert (isRegZ r = true) by (unfold isRegZ; blia); rewrite H4 in H3. eapply H3. repeat match goal with | |- exists _, _ => eexists @@ -504,7 +505,7 @@ Section Spilling. | |- _ => eassumption || reflexivity end. Qed. - + (* Lemma load_iarg_reg_correct'(i: Z): forall r e2 t1 t2 m1 m2 l1 l2 mc1 mc2 post frame maxvar v fpval, *) (* i = 1 \/ i = 2 -> *) (* related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> *) @@ -582,9 +583,9 @@ Section Spilling. | |- _ => eassumption || reflexivity end. 1: eapply put_tmp; eassumption. - unfold exec.cost_SLoad. assert (isRegZ (9+i) = true) by (unfold isRegZ; blia); rewrite H0. - assert (fp = 5) by auto; rewrite H1; cbn. - destr (isRegZ r); solve_MetricLog. + unfold cost_load. assert (isRegZ (9+i) = true) by (unfold isRegZ; blia); rewrite H0. + assert (fp = 5) by auto; rewrite H1; cbn. + destr (isRegZ r); solve_MetricLog. - eapply exec.skip. assert (l2 = map.put l2 r v) as F. { symmetry. apply map.put_idemp. @@ -617,7 +618,7 @@ Section Spilling. related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ post t1' m1' l1' mc1'' /\ (mc2'' - mc2 <= mc1'' - mc1)%metricsH). - Proof. + Proof. intros. unfold save_ires_reg, stack_loc, ires_reg, related in *. fwd. destr (32 <=? x). @@ -662,8 +663,8 @@ Section Spilling. } 1: { unfold spill_tmp. eapply put_tmp; eauto. } 1: blia. - unfold exec.cost_SStore. unfold spill_tmp; cbn. - destr (isRegZ x); solve_MetricLog. + unfold cost_store. unfold spill_tmp; cbn. + destr (isRegZ x); solve_MetricLog. } blia. - eapply exec.skip. @@ -730,11 +731,11 @@ Section Spilling. eapply get_sep. ecancel_assumption. - rewrite map.get_put_same. reflexivity. - exact St. - - unfold exec.cost_SStore. + - unfold cost_store. assert (isRegZ (spill_tmp 1) = true) by auto; rewrite H. assert (isRegZ fp = true) by auto; rewrite H0. - clear H H0. - destr (isRegZ x); try blia. + clear H H0. + destr (isRegZ x); try blia. eapply H1. repeat match goal with | |- exists _, _ => eexists @@ -771,7 +772,7 @@ Section Spilling. } blia. - eapply exec.skip. - destr (isRegZ x); try blia. + destr (isRegZ x); try blia. eapply H1. (* even though we did nothing, we have to reconstruct the `related` from the `related` that *) (* held *before* the SOp *) @@ -876,9 +877,9 @@ Section Spilling. addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 (cost_set_vars_to_reg_range args start mc))). Proof. - induction args; intros; cbn; try trivial. - destr (isRegZ a); cbn; rewrite IHargs; trivial. - Qed. + induction args; intros; cbn; try trivial. + destr (isRegZ a); cbn; rewrite IHargs; trivial. + Qed. Lemma cost_set_vars_to_reg_range_commutes': forall args start mc, @@ -887,10 +888,10 @@ Section Spilling. addMetricInstructions 1 (addMetricLoads 1 (cost_set_vars_to_reg_range args start mc)). Proof. - induction args; intros; cbn; try trivial. - destr (isRegZ a); cbn; rewrite IHargs; trivial. - Qed. - + induction args; intros; cbn; try trivial. + destr (isRegZ a); cbn; rewrite IHargs; trivial. + Qed. + Lemma set_vars_to_reg_range_correct: forall args start argvs e t1 t2 m1 m2 l1 l1' l2 mc2 maxvar frame post fpval, related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> @@ -942,7 +943,7 @@ Section Spilling. { blia. } * intros. apply H6; auto. cbn in *. destr (isRegZ start); try blia; destr (isRegZ a); try blia. - rewrite H0. rewrite cost_set_vars_to_reg_range_commutes. trivial. + rewrite H0. rewrite cost_set_vars_to_reg_range_commutes. trivial. + eapply exec.set. { eassumption. } eapply IHargs; try eassumption; try blia. 2: { @@ -969,7 +970,7 @@ Section Spilling. { assumption. } * intros. apply H6; auto. cbn in *. destr (isRegZ start); try blia; destr (isRegZ a); try blia. - rewrite H0. rewrite cost_set_vars_to_reg_range_commutes'. trivial. + rewrite H0. rewrite cost_set_vars_to_reg_range_commutes'. trivial. Qed. Fixpoint cost_set_reg_range_to_vars (start : Z) (args: list Z) (mc : MetricLog) : MetricLog := @@ -979,7 +980,7 @@ Section Spilling. addMetricInstructions 1 (addMetricLoads 1 (cost_set_reg_range_to_vars (start + 1) xs mc)) else (addMetricInstructions 1 (addMetricLoads 2 (cost_set_reg_range_to_vars (start + 1) xs mc))) end. - + Lemma set_reg_range_to_vars_correct: forall args argvs start e t1 t2 m1 m2 l1 l2 mc2 maxvar frame post fpval, related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> @@ -1088,21 +1089,21 @@ Section Spilling. Qed. Lemma factor_out_cost_external : forall phase mc, - exec.cost_SCall_external phase mc = - (mc + exec.cost_SCall_external phase EmptyMetricLog)%metricsH. + cost_call_external phase mc = + (mc + cost_call_external phase EmptyMetricLog)%metricsH. Proof. destruct phase; eauto with metric_arith. Qed. Lemma factor_out_cost_internal : forall phase mc, - exec.cost_SCall_internal phase mc = - (mc + exec.cost_SCall_internal phase EmptyMetricLog)%metricsH. + cost_call_internal phase mc = + (mc + cost_call_internal phase EmptyMetricLog)%metricsH. Proof. destruct phase; eauto with metric_arith. Qed. Lemma factor_out_cost_stackalloc : forall x mc, - exec.cost_SStackalloc isRegZ x mc = - (mc + exec.cost_SStackalloc isRegZ x EmptyMetricLog)%metricsH. + cost_stackalloc isRegZ x mc = + (mc + cost_stackalloc isRegZ x EmptyMetricLog)%metricsH. Proof. intros. - unfold exec.cost_SStackalloc, EmptyMetricLog. + unfold cost_stackalloc, EmptyMetricLog. destruct (isRegZ x); solve_MetricLog_piecewise. Qed. @@ -1339,7 +1340,7 @@ Section Spilling. (t: Semantics.trace)(m: mem)(argvals: list word)(mc: MetricLog) (post: Semantics.trace -> mem -> list word -> MetricLog -> Prop): Prop := forall l, map.of_list_zip argnames argvals = Some l -> - execpre e fbody t m l (exec.cost_SCall_internal PreSpill mc) (fun t' m' l' mc' => + execpre e fbody t m l (cost_call_internal PreSpill mc) (fun t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals mc'). @@ -1523,7 +1524,7 @@ Section Spilling. { eassumption. } { add_bounds. - unfold exec.cost_SStackalloc, exec.cost_SCall_internal in *. + unfold cost_stackalloc, cost_call_internal in *. destruct (isRegZ fp); solve_MetricLog. } Qed. @@ -1531,17 +1532,17 @@ Section Spilling. Lemma iarg_reg_isReg: forall i a, (i <= 20) -> - (isRegZ (iarg_reg i a) = true). + (isRegZ (iarg_reg i a) = true). Proof. intros. unfold isRegZ, iarg_reg. destr (32 <=? a); unfold spill_tmp; blia. Qed. - + Lemma ires_reg_isReg: forall r, - (isRegZ (ires_reg r) = true). + (isRegZ (ires_reg r) = true). Proof. intros. unfold isRegZ, ires_reg. destr (32 <=? r); unfold spill_tmp; blia. Qed. - + Ltac isReg_helper := match goal with | |- context[(isRegZ (iarg_reg _ _))] => rewrite iarg_reg_isReg by blia @@ -1550,8 +1551,10 @@ Section Spilling. | H: context[(isRegZ (ires_reg _))] |- _ => rewrite ires_reg_isReg in H end. + Ltac irs := cost_unfold; repeat isReg_helper; cost_solve. + Ltac sirs := scost_unfold; repeat isReg_helper; scost_solve. + - Lemma spilling_correct (e1 e2 : env) (Ev : spill_functions e1 = Success e2) (s1 : stmt) (t1 : Semantics.trace) @@ -1852,8 +1855,7 @@ Section Spilling. rewrite factor_out_set_reg_range_to_vars in Hmetrics. rewrite (factor_out_cost_internal PreSpill) in Hmetrics. add_bounds. - unfold exec.cost_SStackalloc, exec.cost_SCall_internal, exec.cost_SCall_external, EmptyMetricLog in *. - destruct (isRegZ fp); solve_MetricLog. + cost_solve. (* cost_SCall constraint: prespill - postspill >= (...66...) i think? *) (* TODO XXX fairly slow *) } @@ -1861,7 +1863,7 @@ Section Spilling. - (* exec.load *) eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - intros. + intros. eapply exec.seq_cps. pose proof H2 as A. unfold related in A. fwd. unfold Memory.load, Memory.load_Z, Memory.load_bytes in *. fwd. @@ -1875,8 +1877,7 @@ Section Spilling. + eassumption. + eassumption. + blia. - + unfold exec.cost_SLoad in *; repeat isReg_helper. - destr (isRegZ x); destr (isRegZ a); solve_MetricLog. + + irs. - (* exec.store *) eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. @@ -1896,11 +1897,10 @@ Section Spilling. | |- exists _, _ => eexists | |- _ /\ _ => split end. - all: try eassumption || reflexivity. + all: try eassumption || reflexivity. spec store_bytes_sep_hi2lo as A. 1: eassumption. - all: ecancel_assumption. - + unfold exec.cost_SStore; repeat isReg_helper. - destr (isRegZ v); destr (isRegZ a); solve_MetricLog. + all: ecancel_assumption. + + irs. - (* exec.inlinetable *) eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.seq_cps. @@ -1912,8 +1912,7 @@ Section Spilling. + eassumption. + eassumption. + blia. - + unfold exec.cost_SInlinetable in *; repeat isReg_helper. - destr (isRegZ x); destr (isRegZ i); solve_MetricLog. + + irs. - (* exec.stackalloc *) rename H1 into IH. eapply exec.stackalloc. 1: assumption. @@ -1931,15 +1930,14 @@ Section Spilling. | |- _ /\ _ => split end. 1,4,3,2: eassumption. - unfold exec.cost_SStackalloc in *; repeat isReg_helper. - destr (isRegZ x); solve_MetricLog. + irs. - (* exec.lit *) eapply exec.seq_cps. eapply exec.lit. eapply save_ires_reg_correct. + eassumption. + eassumption. + blia. - + unfold exec.cost_SLit; repeat isReg_helper. destr (isRegZ x); solve_MetricLog. + + irs. - (* exec.op *) unfold exec.lookup_op_locals in *. eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). @@ -1951,16 +1949,14 @@ Section Spilling. { eapply get_iarg_reg_1; eauto with zarith. } { unfold exec.lookup_op_locals in *. apply map.get_put_same. } { eapply save_ires_reg_correct; (try eassumption || blia). - unfold exec.cost_SOp; repeat isReg_helper. - destr (isRegZ v); destr (isRegZ x); destr (isRegZ y); solve_MetricLog. } + sirs. } } { eapply exec.seq_cps. eapply exec.op. { apply map.get_put_same. } { unfold exec.lookup_op_locals in *. reflexivity. } { eapply save_ires_reg_correct; (try eassumption || blia). - unfold exec.cost_SOp; repeat isReg_helper. - destr (isRegZ x); destr (isRegZ y); solve_MetricLog. } + sirs. } } - (* exec.set *) eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. @@ -1970,8 +1966,7 @@ Section Spilling. + eassumption. + eassumption. + blia. - + unfold exec.cost_SSet; repeat isReg_helper. - destr (isRegZ x); destr (isRegZ y); solve_MetricLog. + + irs. - (* exec.if_true *) unfold prepare_bcond. destr cond; cbn [ForallVars_bcond eval_bcond spill_bcond] in *; fwd. + eapply exec.seq_assoc. @@ -1980,20 +1975,18 @@ Section Spilling. eapply exec.if_true. { cbn. erewrite get_iarg_reg_1 by eauto with zarith. rewrite map.get_put_same. congruence. } - eapply exec.weaken. + eapply exec.weaken. * eapply IHexec; eassumption. * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. - unfold exec.cost_SIf in *; repeat isReg_helper. - destr (isRegZ y); destr (isRegZ x); solve_MetricLog. + sirs. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.if_true. { cbn. rewrite map.get_put_same. rewrite word.eqb_ne by assumption. reflexivity. } - eapply exec.weaken. + eapply exec.weaken. * eapply IHexec; eassumption. * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. - unfold exec.cost_SIf in *; repeat isReg_helper. - destr (isRegZ x); solve_MetricLog. + sirs. - (* exec.if_false *) unfold prepare_bcond. destr cond; cbn [ForallVars_bcond eval_bcond spill_bcond] in *; fwd. + eapply exec.seq_assoc. @@ -2002,20 +1995,18 @@ Section Spilling. eapply exec.if_false. { cbn. erewrite get_iarg_reg_1 by eauto with zarith. rewrite map.get_put_same. congruence. } - eapply exec.weaken. + eapply exec.weaken. * eapply IHexec; eassumption. * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. - unfold exec.cost_SIf in *; repeat isReg_helper. - destr (isRegZ x); destr (isRegZ y); solve_MetricLog. + sirs. + eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.if_false. { cbn. rewrite map.get_put_same. rewrite word.eqb_eq; reflexivity. } - eapply exec.weaken. + eapply exec.weaken. * eapply IHexec; eassumption. * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. - unfold exec.cost_SIf in *; repeat isReg_helper. - destr (isRegZ x); solve_MetricLog. + sirs. - (* exec.loop *) rename IHexec into IH1, H3 into IH2, H5 into IH12. eapply exec.loop_cps. @@ -2035,11 +2026,10 @@ Section Spilling. erewrite get_iarg_reg_1 by eauto with zarith. rewrite map.get_put_same. eexists. split; [reflexivity|]. split; intros. - * do 4 eexists. split. 2: split. + * do 4 eexists. split. 2: split. -- exact H3p8. -- eapply H1. 1: eassumption. cbn. rewrite E, E0. congruence. - -- unfold exec.cost_SLoop_false; repeat isReg_helper. - destr (isRegZ x); destr (isRegZ y); solve_MetricLog. + -- sirs. * eapply exec.weaken. 1: eapply IH2. -- eassumption. -- cbn. rewrite E, E0. congruence. @@ -2048,8 +2038,7 @@ Section Spilling. -- cbv beta. intros. fwd. eapply exec.weaken. ++ eapply IH12; try eassumption. repeat split; eauto; blia. ++ cbv beta; intros; fwd. exists t1'1, m1'1, l1'1, mc1'1. split. 2:split. all: try eassumption. - unfold exec.cost_SLoop_true in *; repeat isReg_helper. - destr (isRegZ x); destr (isRegZ y); solve_MetricLog. + sirs. + specialize H0 with (1 := H3p1). cbn in H0. fwd. eapply exec.weaken. { eapply load_iarg_reg_correct''; (blia || eassumption || idtac). @@ -2058,10 +2047,9 @@ Section Spilling. rewrite map.get_put_same. eexists. split; [reflexivity|]. split; intros. * do 4 eexists. split. 2: split. - -- exact H3p6. + -- exact H3p6. -- eapply H1. 1: eassumption. cbn. rewrite E. congruence. - -- unfold exec.cost_SLoop_false; repeat isReg_helper. - destr (isRegZ x); solve_MetricLog. + -- sirs. * eapply exec.weaken. 1: eapply IH2. -- eassumption. -- cbn. rewrite E. congruence. @@ -2070,19 +2058,18 @@ Section Spilling. -- cbv beta. intros. fwd. eapply exec.weaken. ++ eapply IH12; try eassumption. repeat split; eauto; blia. ++ cbv beta; intros; fwd. exists t1'1, m1'1, l1'1, mc1'1. split. 2:split. all: try eassumption. - unfold exec.cost_SLoop_true in *; repeat isReg_helper. - destr (isRegZ x); solve_MetricLog. + sirs. - (* exec.seq *) cbn in *. fwd. rename H1 into IH2, IHexec into IH1. eapply exec.seq. + eapply IH1. 1: eassumption. eauto 15. + cbn. intros. fwd. eapply exec.weaken. - * eapply IH2; eassumption. - * cbv beta. intros. fwd. exists t1'0, m1'0, l1'0, mc1'0. split. 2:split. all: try eassumption. - solve_MetricLog. + * eapply IH2; eassumption. + * cbv beta. intros. fwd. exists t1'0, m1'0, l1'0, mc1'0. split. 2:split. all: try eassumption. + solve_MetricLog. - (* exec.skip *) - eapply exec.skip. exists t, m, l, mc. repeat split; eauto; solve_MetricLog. + eapply exec.skip. exists t, m, l, mc. repeat split; eauto; solve_MetricLog. Qed. Lemma spill_fun_correct: forall e1 e2 argnames1 retnames1 body1 argnames2 retnames2 body2, From 33e328be47d6fd02f4c67f2e45b719b79dc139e4 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Sat, 15 Jun 2024 03:40:36 -0400 Subject: [PATCH 37/62] MetricCosts migration: RegAlloc it's messy and can use improvements --- compiler/src/compiler/RegAlloc.v | 187 ++++++++----------------------- 1 file changed, 46 insertions(+), 141 deletions(-) diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index 115ff99ab..b5cfe57f5 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -9,6 +9,7 @@ Require Import coqutil.Datatypes.ListSet. Require Import coqutil.Tactics.fwd. Require Import coqutil.Tactics.autoforward. Require Import compiler.Registers. +Require Import bedrock2.MetricCosts. Open Scope Z_scope. @@ -592,8 +593,8 @@ Definition assert_in(y: srcvar)(y': impvar)(m: list (srcvar * impvar)): result u match List.find (mapping_eqb (y, y')) m with | Some _ => if check_regs y y' then Success tt else error:("The register allocator found a mapping of source register variable" y - "to target stack variable" y' - "but source register variables must be in registers in the target.") + "to target stack variable" y' + "but source register variables must be in registers in the target.") | None => error:("The register allocator replaced source variable" y "by target variable" y' "but when the checker encountered this pair," @@ -631,7 +632,7 @@ Definition assert_ins(args: list srcvar)(args': list impvar)(m: list (srcvar * i "and a target variable list" args' "that are incompatible with its current mapping of source to target variables" m);; assert (List.forallb (fun '(x, x') => if check_regs x x' then true else false) - (List.combine args args')) + (List.combine args args')) error:("Register allocation checker got a source variable list" args "and a target variable list" args' "in which at least one source register variable is on the stack in the target."). @@ -655,7 +656,7 @@ Definition assignment(m: list (srcvar * impvar))(x: srcvar)(x': impvar): result error:("Register allocation checker got an assignment of source register variable" x "to target stack variable" x' "but source register variables must be in registers in the target."). - + Fixpoint assignments(m: list (srcvar * impvar))(xs: list srcvar)(xs': list impvar): result (list (srcvar * impvar)) := match xs, xs' with @@ -826,8 +827,8 @@ Lemma extends_cons_r: forall a l1 l2, Proof. unfold extends, assert_in. simpl. intros. fwd. destruct_one_match_hyp. 2: { - specialize (H0 x x'). rewrite E in H0. rewrite E0 in H0. - eapply H0. reflexivity. + specialize (H0 x x'). rewrite E in H0. rewrite E0 in H0. + eapply H0. reflexivity. } destruct_one_match. 1: reflexivity. eapply find_none in E2. 2: eassumption. @@ -1082,7 +1083,7 @@ Section CheckerCorrect. destr (String.eqb x k). + exfalso. congruence. + eapply H0. 2: eassumption. unfold assert_in. - destruct_one_match. 1: rewrite E0; trivial. + destruct_one_match. 1: rewrite E0; trivial. eapply find_none in E3. 2: eassumption. simpl in E3. rewrite String.eqb_refl, Z.eqb_refl in E3. discriminate. Qed. @@ -1113,7 +1114,7 @@ Section CheckerCorrect. induction xs; intros; destruct xs'; try discriminate. - reflexivity. - simpl in *. f_equal. destruct_one_match_hyp; try discriminate. - specialize (IHxs xs' a0 m2). eauto. + specialize (IHxs xs' a0 m2). eauto. Qed. Lemma assert_ins_same_length: forall xs xs' m u, @@ -1122,6 +1123,7 @@ Section CheckerCorrect. unfold assert_ins, assert. intros. fwd. assumption. Qed. + (* TODO XXX this should be obsolete *) Ltac unfold_metrics := unfold MetricLogging.withInstructions, MetricLogging.withLoads, @@ -1143,103 +1145,13 @@ Section CheckerCorrect. unfold check_regs_op, assert_in_op, assert_in; intros; destruct x'; fwd; reflexivity. Qed. - Lemma check_regs_cost_SLoad: forall x x' a a' mc mc', - check_regs a a' = true -> - check_regs x x' = true -> - (MetricLogging.metricsLeq - (MetricLogging.metricsSub (exec.cost_SLoad isRegZ x' a' mc') mc') - (MetricLogging.metricsSub (exec.cost_SLoad isRegStr x a mc) mc)). - Proof. - intros; unfold check_regs in *; cbn in *; - destr (isRegStr x); destr (isRegStr a); destr (isRegZ x'); destr (isRegZ a'); - try discriminate; - unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; blia. - Qed. - - Lemma check_regs_cost_SStore: forall x x' a a' mc mc', - check_regs a a' = true -> - check_regs x x' = true -> - (MetricLogging.metricsLeq - (MetricLogging.metricsSub (exec.cost_SStore isRegZ x' a' mc') mc') - (MetricLogging.metricsSub (exec.cost_SStore isRegStr x a mc) mc)). - Proof. - intros; unfold check_regs in *; cbn in *; - destr (isRegStr x); destr (isRegStr a); destr (isRegZ x'); destr (isRegZ a'); - try discriminate; - unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; blia. - Qed. - - Lemma check_regs_cost_SInlinetable: forall x x' a a' mc mc', - check_regs a a' = true -> - check_regs x x' = true -> - (MetricLogging.metricsLeq - (MetricLogging.metricsSub (exec.cost_SInlinetable isRegZ x' a' mc') mc') - (MetricLogging.metricsSub (exec.cost_SInlinetable isRegStr x a mc) mc)). - Proof. - intros; unfold check_regs in *; cbn in *; - destr (isRegStr x); destr (isRegStr a); destr (isRegZ x'); destr (isRegZ a'); - try discriminate; - unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; blia. - Qed. - - Lemma check_regs_cost_SStackalloc: forall x x' mcL mcL' mcH mcH', - check_regs x x' = true -> - (MetricLogging.metricsLeq (MetricLogging.metricsSub mcL' mcL) - (MetricLogging.metricsSub mcH' mcH)) -> - (MetricLogging.metricsLeq - (MetricLogging.metricsSub (exec.cost_SStackalloc isRegZ x' mcL') mcL) - (MetricLogging.metricsSub (exec.cost_SStackalloc isRegStr x mcH') mcH)). - Proof. - intros; unfold check_regs in *; cbn in *; unfold exec.cost_SStackalloc in *; - destr (isRegStr x); destr (isRegZ x'); - try discriminate; - unfold_metrics; cbn; repeat split; destruct mcL; destruct mcH; destruct mcL'; destruct mcH'; cbn in *; unfold_metrics; cbn in *; try blia. - Qed. - - Lemma check_regs_cost_SLit: forall x x' mc mc', - check_regs x x' = true -> - (MetricLogging.metricsLeq - (MetricLogging.metricsSub (exec.cost_SLit isRegZ x' mc') mc') - (MetricLogging.metricsSub (exec.cost_SLit isRegStr x mc) mc)). - Proof. - intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLit in *; - destr (isRegStr x); destr (isRegZ x'); - try discriminate; - unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; try blia. - Qed. - - Lemma check_regs_cost_SOp: forall x x' y y' z z' mc mc', - check_regs x x' = true -> - check_regs y y' = true -> - check_regs_op z z' = true -> - (MetricLogging.metricsLeq - (MetricLogging.metricsSub (exec.cost_SOp isRegZ x' y' z' mc') mc') - (MetricLogging.metricsSub (exec.cost_SOp isRegStr x y z mc) mc)). - Proof. - intros; unfold check_regs_op, check_regs in *; cbn in *; - destr (isRegStr x); destr (isRegZ x'); destr (isRegStr y); destr (isRegZ y'); destruct z as [z|z]; destruct z' as [z'|z']; try destr (isRegStr z); try destr (isRegZ z'); - try discriminate; - unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; blia. - Qed. - - Lemma check_regs_cost_SSet: forall x x' y y' mc mc', - check_regs x x' = true -> - check_regs y y' = true -> - (MetricLogging.metricsLeq - (MetricLogging.metricsSub (exec.cost_SSet isRegZ x' y' mc') mc') - (MetricLogging.metricsSub (exec.cost_SSet isRegStr x y mc) mc)). - Proof. - intros; unfold check_regs in *; cbn in *; - destr (isRegStr x); destr (isRegStr y); destr (isRegZ x'); destr (isRegZ y'); - try discriminate; - unfold_metrics; cbn; repeat split; destruct mc; cbn; unfold_metrics; cbn; blia. - Qed. - Ltac discr_match_success := match goal with | H: match ?expr with _ => _ end = Success _ |- _ => destr expr; discriminate end. - + + (* TODO XXX these lemms are unnecessarily slow i think *) + Lemma check_regs_cost_SIf: forall (cond : bcond) (bThen bElse : stmt) (corresp : list (string * Z)) (cond0 : bcond') (s'1 s'2 : stmt') (a0 a1 : list (string * Z)) @@ -1261,12 +1173,12 @@ Section CheckerCorrect. try discriminate; unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; fwd; destr mcL; destr mc'; destr mc; destr mcH'; try blia. - all: try discr_match_success. + all: try discr_match_success. all: cbn in *; fwd; try blia. Qed. Lemma check_regs_cost_SLoop_false: - forall (cond : bcond) (body1 body2 : stmt) (corresp' : list (string * Z)) (s'1 : stmt') + forall (cond : bcond) (body1 body2 : stmt) (corresp' : list (string * Z)) (s'1 : stmt') (cond0 : bcond') (s'2 : stmt') (mc mcL : MetricLogging.MetricLog) (a : list (string * Z)), check a body1 s'1 = Success corresp' -> check_bcond corresp' cond cond0 = Success tt -> @@ -1284,7 +1196,7 @@ Section CheckerCorrect. try discriminate; unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; fwd; destr mcL; destr mc'; destr mc; destr mcH'; try blia. - all: try discr_match_success. + all: try discr_match_success. all: cbn in *; fwd; try blia. Qed. @@ -1292,12 +1204,18 @@ Section CheckerCorrect. Hint Resolve states_compat_get : checker_hints. Hint Resolve states_compat_put : checker_hints. Hint Resolve states_compat_get_op : checker_hints. - Hint Resolve states_compat_then_op : checker_hints. - Hint Resolve - check_regs_cost_SLoad check_regs_cost_SStore check_regs_cost_SInlinetable - check_regs_cost_SStackalloc check_regs_cost_SLit (*check_regs_cost_SOp*) - check_regs_cost_SSet check_regs_cost_SIf check_regs_cost_SLoop_false - : checker_hints. + Hint Resolve states_compat_then_op : checker_hints. + (*Hint Resolve*) + (* check_regs_cost_SLoad check_regs_cost_SStore check_regs_cost_SInlinetable*) + (* check_regs_cost_SStackalloc check_regs_cost_SLit (*check_regs_cost_SOp*)*) + (* check_regs_cost_SSet check_regs_cost_SIf check_regs_cost_SLoop_false*) + (* : checker_hints.*) + + Ltac a := + repeat match goal with | |- MetricLogging.metricsLeq _ _ => fail 1 | _ => econstructor end; + eauto 10 with checker_hints. + + Ltac b := unfold assert_in, assignment, check_regs in *; cost_hammer. Lemma checker_correct: forall (e: srcEnv) (e': impEnv) s t m lH mcH post, check_funcs e e' = Success tt -> @@ -1334,7 +1252,7 @@ Section CheckerCorrect. eapply putmany_of_list_zip_states_compat in P. 2-3: eassumption. destruct P as (lL' & P & SC). eexists. split. 1: eassumption. intros. eauto. repeat eexists; eauto. - all: repeat (unfold_metrics; cbn; try blia). + all: repeat (unfold_metrics; cbn; try blia). - (* Case exec.call *) rename binds0 into binds', args0 into args'. unfold check_funcs in H. @@ -1360,41 +1278,29 @@ Section CheckerCorrect. * eapply states_compat_getmany; eassumption. * exact L4. * repeat eexists. 6: eassumption. 1: exact SC. - all: repeat (unfold_metrics; cbn in *; try blia). + all: repeat (unfold_metrics; cbn in *; try blia). - (* Case exec.load *) - repeat econstructor; eauto 10 with checker_hints. - all: unfold assert_in, assignment in *; fwd. - all: eapply check_regs_cost_SLoad; eauto. + a. b. - (* Case exec.store *) - repeat econstructor; eauto 10 with checker_hints. - all: unfold assert_in, assignment in *; fwd. - all: eapply check_regs_cost_SStore; eauto. + a. b. - (* Case exec.inlinetable *) - repeat econstructor; eauto 10 with checker_hints. - all: unfold assert_in, assignment in *; fwd. - all: eapply check_regs_cost_SInlinetable; eauto. + a. b. - (* Case exec.stackalloc *) eapply exec.stackalloc. 1: assumption. intros. eapply exec.weaken. + eapply H2; try eassumption. - eapply states_compat_precond. eapply states_compat_put. all: eassumption. + eapply states_compat_precond. eapply states_compat_put; eassumption. + cbv beta. intros. fwd. - eexists. eexists. repeat (split; try eassumption). - eexists. eexists. repeat (split; try eassumption). - all: unfold assert_in, assignment in *; fwd. - all: eapply check_regs_cost_SStackalloc; eauto. + eexists. eexists. do 2 (split; try eassumption). + eexists. eexists. do 2 (split; try eassumption). + b. - (* Case exec.lit *) - repeat econstructor; eauto 10 with checker_hints. - all: unfold assert_in, assignment in *; fwd. - all: eapply check_regs_cost_SLit; eauto. + a. b. - (* Case exec.op *) - repeat econstructor; eauto 10 with checker_hints. - all: unfold assert_in, assignment in *; fwd. - all: eapply check_regs_cost_SOp; eauto; eapply assert_in_then_check_regs_op; eauto. + a. + unfold assert_in_op, assert_in, assignment, check_regs in *; scost_hammer. (* TODO XXX could be faster *) - (* Case exec.set *) - repeat econstructor; eauto 10 with checker_hints. - all: unfold assert_in, assignment in *; fwd. - all: eapply check_regs_cost_SSet; eauto. + a. b. - (* Case exec.if_true *) eapply exec.if_true. 1: eauto using states_compat_eval_bcond. eapply exec.weaken. @@ -1436,7 +1342,7 @@ Section CheckerCorrect. 1: eapply states_compat_precond; eassumption. cbv beta. intros. fwd. eexists. eexists. split. 2: split. 1,3: eauto. exists mcH'. exists mc'. - split; eauto. + split; eauto. + cbv beta. intros. fwd. eapply exec.weaken. 1: eapply IH12. 1: eassumption. 1: eassumption. * eapply states_compat_extends. 2: eassumption. pose proof defuel_loop_inv as P. @@ -1451,7 +1357,7 @@ Section CheckerCorrect. * cbv beta. intros. fwd. eexists. eexists. split. 2: split. 1: eauto. 2: eauto. intros. repeat (unfold check_bcond, assert_in, assignment in *; fwd). - clear -E0 E1 E2 H4p1p0 H4p1p1 H4p1 H4p3. + clear -E0 E1 E2 H4p1p0 H4p1p1 H4p1 H4p3. intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLoop_true in *; try discr_match_success; destr cond; destr cond0; destr (isRegStr x); destr (isRegZ x0); try (destr (isRegStr y)); try (destr (isRegZ y0)); try discriminate; try discr_match_success. @@ -1465,13 +1371,12 @@ Section CheckerCorrect. + eapply IH1. 1: eassumption. eapply states_compat_precond. assumption. + cbv beta. intros t' m' l' mcblah ?. fwd. - eapply IH2 in H2p2. 2,3: eauto using states_compat_precond. + eapply IH2 in H2p2. 2,3: eauto using states_compat_precond. eapply exec.weaken; eauto. cbv beta. intros. fwd. exists lH'0. exists mcH'0. split. 2:split. 1,3: eauto. - destr mc; destr mcL; unfold_metrics; cbn in *; unfold_metrics; cbn in *; blia. + b. - (* case exec.skip *) - eapply exec.skip. eexists. exists mc. split. 2: split. 1,3: eauto. - destr mc; destr mcL; unfold_metrics; cbn in *; unfold_metrics; cbn in *; blia. + a. b. Qed. End CheckerCorrect. From 25004bf9ec55a34f4fce041b1e540cc0c639d989 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Sat, 15 Jun 2024 03:43:53 -0400 Subject: [PATCH 38/62] MetricCosts migration: UseImmediate --- compiler/src/compiler/UseImmediate.v | 21 ++++++--------------- 1 file changed, 6 insertions(+), 15 deletions(-) diff --git a/compiler/src/compiler/UseImmediate.v b/compiler/src/compiler/UseImmediate.v index 8a197bf2b..7d9114c2e 100644 --- a/compiler/src/compiler/UseImmediate.v +++ b/compiler/src/compiler/UseImmediate.v @@ -6,6 +6,7 @@ Require Import coqutil.Tactics.fwd. Require Import String. Require Import compiler.UseImmediateDef. Require Import bedrock2.MetricLogging. +Require Import bedrock2.MetricCosts. Local Notation var := String.string (only parsing). @@ -49,31 +50,21 @@ Section WithArguments. | H : _ <= _ |- _ => revert H end; clear; - unfold exec.cost_SInteract, exec.cost_SCall_internal, - exec.cost_SCall_external, exec.cost_SLoad, exec.cost_SStore, - exec.cost_SInlinetable, exec.cost_SStackalloc, exec.cost_SLit, - exec.cost_SOp, exec.cost_SSet, exec.cost_SIf, exec.cost_SLoop_true, - exec.cost_SLoop_false; - repeat match goal with - | x : operand |- _ => destr x - | x : bcond _ |- _ => destr x - | |- context[isRegStr ?x] => destr (isRegStr x) - end; - try solve_MetricLog + FlatImp.scost_hammer ). (* TODO these two lemmas are somewhat slow *) Lemma op_cost_y : forall x0 y v0 mcH' mc v mcL lit, exec.cost_SOp isRegStr x0 y (Var v0) EmptyMetricLog - EmptyMetricLog <= - mcH' - exec.cost_SLit isRegStr lit mc -> - exec.cost_SOp isRegStr x0 y (Const v) (exec.cost_SLit isRegStr lit mcL) - mcL <= + mcH' - cost_lit isRegStr lit mc -> + exec.cost_SOp isRegStr x0 y (Const v) (cost_lit isRegStr lit mcL) - mcL <= mcH' - mc. Proof. finish. Qed. Lemma op_cost_v0 : forall x0 y v0 mcH' mc v mcL lit, exec.cost_SOp isRegStr x0 y (Var v0) EmptyMetricLog - EmptyMetricLog <= - mcH' - exec.cost_SLit isRegStr lit mc -> - exec.cost_SOp isRegStr x0 v0 (Const v) (exec.cost_SLit isRegStr lit mcL) - mcL <= + mcH' - cost_lit isRegStr lit mc -> + exec.cost_SOp isRegStr x0 v0 (Const v) (cost_lit isRegStr lit mcL) - mcL <= mcH' - mc. Proof. finish. Qed. From 5851d6d41587a2668f2ff72964401a21e5219aef Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Sat, 15 Jun 2024 03:45:08 -0400 Subject: [PATCH 39/62] MetricCosts migration: DeadCodeElim --- compiler/src/compiler/DeadCodeElim.v | 24 ++++++++++++------------ 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/compiler/src/compiler/DeadCodeElim.v b/compiler/src/compiler/DeadCodeElim.v index 305ad6af4..eeb9c86ec 100644 --- a/compiler/src/compiler/DeadCodeElim.v +++ b/compiler/src/compiler/DeadCodeElim.v @@ -10,7 +10,7 @@ Require Import coqutil.Datatypes.ListSet. Local Notation var := String.string (only parsing). Require Import compiler.util.Common. Require Import bedrock2.MetricLogging. -Require Import compiler.MetricTactics. +Require Import bedrock2.MetricCosts. (* below only for of_list_list_diff *) Require Import compiler.DeadCodeElimDef. @@ -24,7 +24,7 @@ Section WithArguments1. Context {mem: map.map word (Init.Byte.byte : Type) } {mem_ok : map.ok mem } . Context {locals: map.map string word } {locals_ok : map.ok locals }. Context {ext_spec : Semantics.ExtSpec } {ext_spec_ok: Semantics.ext_spec.ok ext_spec } . - + Lemma agree_on_put_existsb_false: forall used_after x (l: locals) lL, map.agree_on (diff (of_list used_after) (singleton_set x)) l lL @@ -38,7 +38,7 @@ Section WithArguments1. intros. propositional idtac. eapply existsb_of_list in H1. rewrite H1 in H0. - discriminate. + discriminate. Qed. Ltac subset_union_solve := @@ -73,7 +73,7 @@ Section WithArguments1. [ idtac | eapply H ]; subset_union_solve | H: map.agree_on ?s ?x ?y |- map.agree_on _ ?y ?x => - eapply agree_on_comm; agree_on_solve + eapply agree_on_comm; agree_on_solve | H: map.agree_on ?s ?mH ?mL, H1: map.putmany_of_list_zip ?lk ?lv ?mH = Some ?mH', H2: map.putmany_of_list_zip ?lk ?lv ?mL = Some ?mL' @@ -109,7 +109,7 @@ Section WithArguments1. rewrite ListSet.of_list_removeb end. - Ltac mcsolve := eexists; split; [|split; cycle 1; [eauto|exec_cost_hammer]]; try assumption. + Ltac mcsolve := eexists; split; [|split; cycle 1; [eauto|FlatImp.scost_hammer]]; try assumption. Lemma dce_correct_aux : forall eH eL, @@ -153,7 +153,7 @@ Section WithArguments1. fwd. eassumption. + erewrite agree_on_getmany. * eapply H1. - * listset_to_set. agree_on_solve. + * listset_to_set. agree_on_solve. + eapply IHexec. eapply agree_on_refl. + intros. @@ -172,10 +172,10 @@ Section WithArguments1. repeat listset_to_set. subset_union_solve. - intros. - eapply agree_on_find in H3; fwd. + eapply agree_on_find in H3; fwd. destr (existsb (eqb x) used_after); fwd. + eapply @exec.load. - * rewrite <- H3p1. eassumption. + * rewrite <- H3p1. eassumption. * eauto. * unfold compile_post. exists (map.put l x v); mcsolve. @@ -200,7 +200,7 @@ Section WithArguments1. eapply agree_on_find in H4; fwd. destr (existsb (eqb x) used_after); fwd. + eapply @exec.inlinetable; eauto. - * rewrite <- H4p1. eassumption. + * rewrite <- H4p1. eassumption. * unfold compile_post; eexists; mcsolve. repeat listset_to_set; agree_on_solve. + eapply @exec.skip; eauto. @@ -255,10 +255,10 @@ Section WithArguments1. eexists; mcsolve. agree_on_solve. + intros. - eapply agree_on_find in H3; fwd. + eapply agree_on_find in H3; fwd. destr (existsb (eqb x) used_after). * eapply @exec.op. - -- rewrite <- H3p1. eassumption. + -- rewrite <- H3p1. eassumption. -- simpl. constructor. -- unfold compile_post. simpl in *. inversion H1. fwd. eexists; mcsolve. repeat listset_to_set. @@ -272,7 +272,7 @@ Section WithArguments1. repeat listset_to_set. destr (existsb (eqb x) used_after). { eapply @exec.set. - - rewrite <- H2p1; eassumption. + - rewrite <- H2p1; eassumption. - unfold compile_post. eexists; mcsolve. agree_on_solve. } From a9d7747eaa483dc89fbb7735a004fea4f90317c7 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Sat, 15 Jun 2024 03:47:49 -0400 Subject: [PATCH 40/62] MetricCosts migration: Pipeline (finished!) the compiler has zero admits :D --- compiler/src/compiler/Pipeline.v | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index e2350f161..420e058c8 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -40,6 +40,7 @@ Require Import compiler.RegAlloc. Require Import compiler.RiscvEventLoop. Require Import compiler.MetricsToRiscv. Require Import bedrock2.MetricLogging. +Require Import bedrock2.MetricCosts. Require Import compiler.FlatToRiscvCommon. Require Import compiler.FlatToRiscvFunctions. Require Import compiler.DivisibleBy4. @@ -163,7 +164,7 @@ Section WithWordAndMem. exists argnames retnames fbody l, map.get e f = Some (argnames, retnames, fbody) /\ map.of_list_zip argnames argvals = Some l /\ - Exec e fbody t m l (exec.cost_SCall_internal PreSpill mc) (fun t' m' l' mc' => + Exec e fbody t m l (cost_call_internal PreSpill mc) (fun t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals mc'). @@ -375,7 +376,7 @@ Section WithWordAndMem. exists retvals. split. + erewrite MapEauto.agree_on_getmany; [ eauto | eapply MapEauto.agree_on_comm; [ eassumption ] ]. - + eexists; split; eauto. unfold exec.cost_SCall_internal in *. solve_MetricLog. + + eexists; split; eauto. unfold cost_call_internal in *. solve_MetricLog. Qed. Lemma regalloc_functions_NoDup: forall funs funs', @@ -590,7 +591,7 @@ Section WithWordAndMem. Some (argnames, retnames, fbody) /\ map.of_list_zip argnames argvals = Some l /\ MetricSemantics.exec (map.of_list functions) fbody t mH l - (exec.cost_SCall_internal PreSpill mc) + (cost_call_internal PreSpill mc) (fun t' m' l' mc' => exists retvals: list word, map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals mc')) -> @@ -665,7 +666,7 @@ Section WithWordAndMem. NoDup (map fst fs) -> compile fs = Success (instrs, finfo, req_stack_size) -> MetricWeakestPrecondition.call (map.of_list fs) fname t mH argvals - (exec.cost_SCall_internal PreSpill mc) post -> + (cost_call_internal PreSpill mc) post -> forall mcL, map.get (map.of_list finfo) fname = Some f_rel_pos -> req_stack_size <= word.unsigned (word.sub stack_hi stack_lo) / bytes_per_word -> From 681aeadf72f3500454e1637e30ff85d5498509a9 Mon Sep 17 00:00:00 2001 From: Andres Erbsen Date: Fri, 21 Jun 2024 21:51:00 +0000 Subject: [PATCH 41/62] fix MetricWeakestPrecondition and ...Properties --- .../src/bedrock2/MetricWeakestPrecondition.v | 6 +- .../MetricWeakestPreconditionProperties.v | 82 ++++++++++++++++--- 2 files changed, 75 insertions(+), 13 deletions(-) diff --git a/bedrock2/src/bedrock2/MetricWeakestPrecondition.v b/bedrock2/src/bedrock2/MetricWeakestPrecondition.v index 2e2e906e6..46a079819 100644 --- a/bedrock2/src/bedrock2/MetricWeakestPrecondition.v +++ b/bedrock2/src/bedrock2/MetricWeakestPrecondition.v @@ -84,7 +84,7 @@ Section WeakestPrecondition. post t m l mc | cmd.store sz ea ev => exists a mc', dexpr m l ea mc (a, mc') /\ - exists v mc'', dexpr m l ev mc (v, mc'') /\ + exists v mc'', dexpr m l ev mc' (v, mc'') /\ store sz m a v (fun m => post t m l (cost_store isRegStr UNK UNK mc'')) | cmd.stackalloc x n c => @@ -107,8 +107,8 @@ Section WeakestPrecondition. | cmd.while _ _ => MetricSemantics.exec e c t m l mc post | cmd.call binds fname arges => exists args mc', dexprs m l arges mc (args, mc') /\ - MetricSemantics.call e fname t m args mc (fun t m rets mc'' => - exists l', map.putmany_of_list_zip binds rets l = Some l' /\ post t m l' (cost_call_internal PreSpill mc'')) + MetricSemantics.call e fname t m args (cost_call_internal PreSpill mc') (fun t m rets mc'' => + exists l', map.putmany_of_list_zip binds rets l = Some l' /\ post t m l' (cost_call_external PreSpill mc'')) | cmd.interact binds action arges => exists args mc', dexprs m l arges mc (args, mc') /\ exists mKeep mGive, map.split m mKeep mGive /\ diff --git a/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v b/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v index 9248a2973..53e89e338 100644 --- a/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v +++ b/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v @@ -157,8 +157,10 @@ Global Instance Proper_program : try solve [typeclasses eauto with core]. Qed. + From coqutil Require Import Datatypes.Prod. Ltac t := repeat match goal with + | _ => progress inversion_prod | |- forall _, _ => progress intros | H: exists _, _ |- _ => destruct H | H: and _ _ |- _ => destruct H @@ -252,34 +254,92 @@ width word word_ok. Lemma sound_cmd e c t m l mc post (H: MetricWeakestPrecondition.cmd e c t m l mc post) : MetricSemantics.exec e c t m l mc post. Proof. ind_on Syntax.cmd; repeat (t; try match reverse goal with H: MetricWeakestPrecondition.expr _ _ _ _ _ |- _ => eapply expr_sound in H end). - { destruct (BinInt.Z.eq_dec (Interface.word.unsigned x) (BinNums.Z0)) as [Hb|Hb]; cycle 1; apply pair_equal_spec in H1; destruct H1; econstructor; t. } - { inversion H0. inversion H. - apply pair_equal_spec in H3; destruct H3. - apply pair_equal_spec in H4; destruct H4. subst x0. - econstructor; t. - Admitted. + { destruct (BinInt.Z.eq_dec (word.unsigned x) 0) as [|]; t. } + { inversion H0. t. eapply sound_args in H; t. } + { eapply sound_args in H; t. } + Qed. Lemma weaken_cmd: forall e c t m l mc (post1 post2: _->_->_->_->Prop), MetricWeakestPrecondition.cmd e c t m l mc post1 -> (forall t m l mc, post1 t m l mc -> post2 t m l mc) -> MetricWeakestPrecondition.cmd e c t m l mc post2. - Proof. Admitted. + Proof. + intros. + eapply Proper_cmd. 2: eassumption. + cbv [RelationClasses.Reflexive Morphisms.pointwise_relation + Morphisms.respectful Basics.impl]. + assumption. + Qed. Lemma complete_args : forall m l args mc vs, MetricSemantics.eval_call_args m l args mc = Some vs -> MetricWeakestPrecondition.dexprs m l args mc vs. - Proof using word_ok. Admitted. + Proof using word_ok. + induction args; cbn; repeat (subst; t). + 1: inversion H; reflexivity. + destruct_one_match_hyp. 2: discriminate. + destruct_one_match_hyp. + destruct_one_match_hyp. 2: discriminate. + case p in *; inversion_clear H. + eapply Proper_expr. 2: eapply expr_complete. 2: eassumption. + intros x ?. subst x. + eapply Proper_list_map. 3: { eapply IHargs. eassumption. } + { eapply Proper_expr. } + { intros ? ?. subst. reflexivity. } + Qed. Lemma complete_cmd: forall e c t m l mc post, MetricSemantics.exec e c t m l mc post -> MetricWeakestPrecondition.cmd e c t m l mc post. - Proof. Admitted. + Proof. + induction 1. + { eassumption. } + { eapply expr_complete in H. eexists _, _. split. 1: exact H. + eassumption. } + { eauto. } + { eapply expr_complete in H. + eapply expr_complete in H0. + eexists _, _. split. 1: eassumption. + eexists _, _. split. 1: eassumption. + eexists. eauto. } + { split. 1: assumption. + intros * HA HSp. specialize H1 with (1 := HA) (2 := HSp). + unfold dlet.dlet. eapply weaken_cmd. 1: eapply H1. cbv beta. + clear. intros * (? & ? & ? & ? & ?). eauto 8. } + { eexists _, _; cbv[dlet.dlet]; ssplit; intros; eauto using expr_complete; congruence. } + { eexists _, _; cbv[dlet.dlet]; ssplit; intros; eauto using expr_complete; congruence. } + { cbn. eapply weaken_cmd. + { eapply IHexec. } + cbv beta. intros. + eapply H1. eassumption. } + { cbn. eapply MetricSemantics.exec.while_false; eauto. } + { rename IHexec into IH1, H3 into IH2. + cbn. eapply MetricSemantics.exec.while_true; eassumption. } + { cbn. eexists _, _. split. + { eapply complete_args. eassumption. } + unfold MetricSemantics.call. do 4 eexists. 1: eassumption. do 2 eexists. 1: eassumption. + eapply MetricSemantics.exec.weaken. + { eassumption. } + cbv beta. intros. + specialize H3 with (1 := H4). destruct H3 as (retvs & G & ? & ? & ?). eauto 8. } + { cbn. eexists _, _. split. + { eapply complete_args. eassumption. } + eexists _, _. split. 1: eassumption. + eapply Semantics.ext_spec.weaken. 2: eassumption. + intros m0 args0 Hmid. specialize H2 with (1 := Hmid). destruct H2 as (? & ? & ?). + eauto 8. } + Qed. Lemma start_func: forall e fname fimpl t m args mc post, map.get e fname = Some fimpl -> MetricWeakestPrecondition.func e fimpl t m args mc post -> MetricWeakestPrecondition.call e fname t m args mc post. - Proof. Admitted. + Proof. + intros * G. destruct fimpl as [[argnames retnames] body]. intros (? & ? & ?). + do 4 eexists. 1: eassumption. do 2 eexists. 1: eassumption. eapply sound_cmd. + eapply weaken_cmd. 1: eassumption. cbv beta. intros. + edestruct sound_getmany as (?&?&?&?); eauto. + Qed. (** Ad-hoc lemmas here? *) @@ -303,6 +363,7 @@ width word word_ok. intros. eapply Properties.map.split_empty_r in H. subst. assumption. Qed.*) + (* Lemma intersect_expr: forall m l e mc (post1 post2: word * MetricLog -> Prop), MetricWeakestPrecondition.expr m l e mc post1 -> MetricWeakestPrecondition.expr m l e mc post2 -> @@ -313,5 +374,6 @@ width word word_ok. (H : MetricWeakestPrecondition.expr m l e mc P) : exists v, MetricWeakestPrecondition.dexpr m l e mc v /\ P v. Proof using word_ok. Admitted. + *) End MetricWeakestPrecondition. From e5ba50c2148b2cd0bdb12ce6bdabd8eb76285d02 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Mon, 8 Jul 2024 00:21:32 +0100 Subject: [PATCH 42/62] account for isReg in regalloc heuristic --- compiler/src/compiler/RegAlloc.v | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index b5cfe57f5..b60073cc0 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -318,9 +318,16 @@ Fixpoint corresp_of_impvars(start_impvar: Z)(o: occupancy): list (srcvar * impva | nil => nil end. +Definition isRegInterval (l: lifetime) : bool := + match l with + | (srcvarname, _) => isRegStr srcvarname + end. + Definition events_to_corresp(events: list event): list (srcvar * impvar) := let sorted_intervals := sort compare_interval_length (events_to_intervals 0 events) in - let occ := List.fold_left assign_srcvar sorted_intervals [] in + let (reg_intervals, nonreg_intervals) := List.partition isRegInterval sorted_intervals in + let regfirst_intervals := reg_intervals ++ nonreg_intervals in + let occ := List.fold_left assign_srcvar regfirst_intervals [] in corresp_of_impvars first_available_impvar occ. From b38989916a93764417a99627c242cfa635ca2ddc Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Mon, 8 Jul 2024 16:09:41 +0100 Subject: [PATCH 43/62] keep old code in a comment for easier diffing --- compiler/src/compiler/RegAlloc.v | 1 + 1 file changed, 1 insertion(+) diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index b60073cc0..0bf8d8564 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -325,6 +325,7 @@ Definition isRegInterval (l: lifetime) : bool := Definition events_to_corresp(events: list event): list (srcvar * impvar) := let sorted_intervals := sort compare_interval_length (events_to_intervals 0 events) in + (* let occ := List.fold_left assign_srcvar sorted_intervals [] in *) let (reg_intervals, nonreg_intervals) := List.partition isRegInterval sorted_intervals in let regfirst_intervals := reg_intervals ++ nonreg_intervals in let occ := List.fold_left assign_srcvar regfirst_intervals [] in From 44380e0e183c7ee7b75f50803176194c68d007ad Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Mon, 8 Jul 2024 22:13:59 +0100 Subject: [PATCH 44/62] add some simple tests for the new heuristic --- compiler/src/compilerExamples/RegAllocTests.v | 568 ++++++++++++++++++ 1 file changed, 568 insertions(+) create mode 100644 compiler/src/compilerExamples/RegAllocTests.v diff --git a/compiler/src/compilerExamples/RegAllocTests.v b/compiler/src/compilerExamples/RegAllocTests.v new file mode 100644 index 000000000..307bcaf14 --- /dev/null +++ b/compiler/src/compilerExamples/RegAllocTests.v @@ -0,0 +1,568 @@ +Require Import bedrock2.NotationsCustomEntry. +Require Import Coq.Lists.List. +Import ListNotations. +Require bedrock2Examples.Demos. +Require Import coqutil.Decidable. +Require Import bedrock2.NotationsCustomEntry coqutil.Macros.WithBaseName. +Require Import compiler.ExprImp. +Require Import compiler.NameGen. +Require Import compiler.Pipeline. +Require Import riscv.Spec.Decode. +Require Import riscv.Utility.Words32Naive. +Require Import riscv.Utility.DefaultMemImpl32. +Require Import riscv.Utility.Monads. +Require Import compiler.util.Common. +Require Import coqutil.Decidable. +Require riscv.Utility.InstructionNotations. +Require Import riscv.Platform.MinimalLogging. +Require Import bedrock2.MetricLogging. +Require Import riscv.Platform.MetricMinimal. +Require Import riscv.Utility.Utility. +Require Import riscv.Utility.Encode. +Require Import coqutil.Map.SortedList. +Require Import compiler.MemoryLayout. +Require Import compiler.StringNameGen. +Require Import riscv.Utility.InstructionCoercions. +Require Import riscv.Platform.MetricRiscvMachine. +Require bedrock2.Hexdump. + +Open Scope Z_scope. +Open Scope string_scope. +Open Scope ilist_scope. + +(* ********************************************************* *) +(* Preliminaries *) +(* ********************************************************* *) + +Definition var: Set := Z. +Definition Reg: Set := Z. + + +Local Existing Instance DefaultRiscvState. + +Axiom TODO: forall {T: Type}, T. + +Local Instance funpos_env: map.map string Z := SortedListString.map _. + +Definition compile_ext_call(posenv: funpos_env)(mypos stackoffset: Z)(s: FlatImp.stmt Z) := + match s with + | FlatImp.SInteract _ fname _ => + if string_dec fname "nop" then + [[Addi Register0 Register0 0]] + else + nil + | _ => [] + end. + +Notation RiscvMachine := MetricRiscvMachine. + +Local Existing Instance coqutil.Map.SortedListString.map. +Local Existing Instance coqutil.Map.SortedListString.ok. + +Definition main_stackalloc := func! { + stackalloc 4 as x; stackalloc 4 as y; swap_swap(x, y) }. + + +(* stack grows from high addreses to low addresses, first stack word will be written to + (stack_pastend-8), next stack word to (stack_pastend-16) etc *) +Definition stack_pastend: Z := 2048. + +Lemma f_equal2: forall {A B: Type} {f1 f2: A -> B} {a1 a2: A}, + f1 = f2 -> a1 = a2 -> f1 a1 = f2 a2. +Proof. intros. congruence. Qed. + +Lemma f_equal3: forall {A B C: Type} {f1 f2: A -> B -> C} {a1 a2: A} {b1 b2: B}, + f1 = f2 -> a1 = a2 -> b1 = b2 -> f1 a1 b1 = f2 a2 b2. +Proof. intros. congruence. Qed. + +Lemma f_equal3_dep: forall {A B C: Type} {f1 f2: A -> B -> C} {a1 a2: A} {b1 b2: B}, + f1 = f2 -> a1 = a2 -> b1 = b2 -> f1 a1 b1 = f2 a2 b2. +Proof. intros. congruence. Qed. + + +Local Instance RV32I_bitwidth: FlatToRiscvCommon.bitwidth_iset 32 RV32I. +Proof. reflexivity. Qed. + + + + +(* ********************************************************* *) +(* Tests *) +(* ********************************************************* *) + +Definition swap := func! (a, b) { + t = load(b); + store(b, load(a)); + store(a, t) +}. + +(* Based on `long1` from compilerExamples/SpillingTests.v *) +(* Expected to fail, since too many `reg` variables are live simultaneously *) +Definition too_many_regs := func! (reg_a0, reg_b0) ~> (reg_a19, reg_b19) { + (* TODO once spilling is improved, test that it still works if these two lines are removed + (currently the test that all resnames are <32 fails if these two lines are removed) *) + reg_a19 = $0; + reg_b19 = $0; + + swap(reg_a0, reg_b0); + reg_a1 = reg_a0 + reg_b0 * reg_b0; + swap(reg_a1, reg_b0); + reg_a12 = reg_a1 - reg_a0 - reg_b0; + reg_b12 = reg_a1 + reg_a12 * reg_b0; + reg_a5 = reg_a12 * reg_b12; + reg_a9 = reg_a5 - reg_a0 - reg_b0; + reg_a4 = reg_a1; + reg_b3 = reg_a4 + reg_a12 * reg_b0; + reg_b4 = reg_a4; + reg_a10 = reg_a1 - reg_a0 - reg_b4; + reg_b12 = reg_a1 + reg_a12 * reg_b0; + reg_a3 = reg_a1; + reg_a6 = reg_a10 + reg_b12; + reg_a15 = reg_a3 - reg_a0 - reg_b0; + reg_a2 = reg_a1; + reg_b12 = reg_a2 + reg_a12 * reg_b0; + reg_a11 = reg_a10 - reg_a0 - reg_b0; + reg_b5 = reg_a1; + reg_b4 = reg_a6 + reg_a12 * reg_b5; + reg_a2 = reg_a11 - reg_a0 - reg_b0; + reg_b12 = reg_a1 + reg_a12 * reg_b0; + reg_a16 = reg_a12 - reg_a0 - reg_b0; + reg_a7 = reg_a2; + reg_b14 = reg_a7 + reg_a12 * reg_b0; + reg_a13 = reg_b4; + reg_a17 = $2; + reg_b6 = $33; + reg_a12 = reg_a13 - reg_a17 - reg_b6; + reg_b12 = reg_a1 + reg_a12 * reg_b0; + reg_a18 = $18; + reg_b19 = $19; + reg_a17 = reg_a1 - reg_a18 - reg_b19; + reg_b12 = reg_a1 + reg_a9 * reg_b0; + reg_a14 = $14; + reg_a12 = reg_a14 - reg_a0 - reg_b0; + reg_a8 = reg_b12; + reg_b13 = reg_a8 + reg_a12 * reg_b0; + reg_a12 = reg_a1 - reg_a0 - reg_b0; + reg_b2 = reg_a1 + reg_a12 * reg_b0; + reg_a18 = reg_a1 - reg_a0 - reg_b0; + reg_b16 = $23; + reg_b12 = reg_a15 + reg_a8 * reg_b16; + reg_a12 = reg_a1 - reg_a0 - reg_b3; + reg_b12 = reg_a1 + reg_a12 * reg_b0; + + reg_a19 = reg_a1 - reg_a0 - reg_b0; + reg_b19 = reg_a9 + reg_a7 * reg_b0; + + reg_b0 = reg_b0 + reg_a1; + reg_b0 = reg_b0 + reg_a12; + reg_b0 = reg_b0 + reg_b12; + reg_b0 = reg_b0 + reg_a5; + reg_b0 = reg_b0 + reg_a9; + reg_b0 = reg_b0 + reg_a4; + reg_b0 = reg_b0 + reg_b3; + reg_b0 = reg_b0 + reg_b4; + reg_b0 = reg_b0 + reg_a10; + reg_b0 = reg_b0 + reg_b12; + reg_b0 = reg_b0 + reg_a3; + reg_b0 = reg_b0 + reg_a6; + reg_b0 = reg_b0 + reg_a15; + reg_b0 = reg_b0 + reg_a2; + reg_b0 = reg_b0 + reg_b12; + reg_b0 = reg_b0 + reg_a11; + reg_b0 = reg_b0 + reg_b5; + reg_b0 = reg_b0 + reg_b4; + reg_b0 = reg_b0 + reg_a2; + reg_b0 = reg_b0 + reg_b12; + reg_b0 = reg_b0 + reg_a16; + reg_b0 = reg_b0 + reg_a7; + reg_b0 = reg_b0 + reg_b14; + reg_b0 = reg_b0 + reg_a13; + reg_b0 = reg_b0 + reg_a17; + reg_b0 = reg_b0 + reg_b6; + reg_b0 = reg_b0 + reg_a12; + reg_b0 = reg_b0 + reg_b12; + reg_b0 = reg_b0 + reg_a18; + reg_b0 = reg_b0 + reg_b19; + reg_b0 = reg_b0 + reg_a17; + reg_b0 = reg_b0 + reg_b12; + reg_b0 = reg_b0 + reg_a14; + reg_b0 = reg_b0 + reg_a12; + reg_b0 = reg_b0 + reg_a8; + reg_b0 = reg_b0 + reg_b13; + reg_b0 = reg_b0 + reg_a12; + reg_b0 = reg_b0 + reg_b2; + reg_b0 = reg_b0 + reg_a18; + reg_b0 = reg_b0 + reg_b16; + reg_b0 = reg_b0 + reg_b12; + reg_b0 = reg_b0 + reg_a12; + reg_b0 = reg_b0 + reg_b12; + reg_b0 = reg_b0 + reg_a19; + reg_b0 = reg_b0 + reg_b19; + + reg_a19 = reg_b0; + reg_b19 = reg_b0 +}. + +Remark too_many_regs_fails: + exists x, compile compile_ext_call &[,too_many_regs] = Failure x. +Proof. + (* let r := eval cbv in (compile compile_ext_call &[,too_many_regs]) in set (res := r). *) + eexists; cbv; reflexivity. +Qed. + + +Definition some_regs := + func! (reg_a0, b0) ~> (reg_a0, b0) { + reg_a1 = reg_a0 + $1; + reg_a2 = reg_a0 + $2; + reg_a3 = reg_a0 + $3; + reg_a4 = reg_a0 + $4; + reg_a5 = reg_a0 + $5; + reg_a6 = reg_a0 + $6; + reg_a7 = reg_a0 + $7; + reg_a8 = reg_a0 + $8; + reg_a9 = reg_a0 + $9; + reg_a10 = reg_a0 + $10; + reg_a11 = reg_a0 + $11; + reg_a12 = reg_a0 + $12; + reg_a13 = reg_a0 + $13; + reg_a14 = reg_a0 + $14; + reg_a15 = reg_a0 + $15; + reg_a16 = reg_a0 + $16; + + b0 = b0 + reg_a0; + b0 = b0 + reg_a1; + b0 = b0 + reg_a2; + b0 = b0 + reg_a3; + b0 = b0 + reg_a4; + b0 = b0 + reg_a5; + b0 = b0 + reg_a6; + b0 = b0 + reg_a7; + b0 = b0 + reg_a8; + b0 = b0 + reg_a9; + b0 = b0 + reg_a10; + b0 = b0 + reg_a11; + b0 = b0 + reg_a12; + b0 = b0 + reg_a13; + b0 = b0 + reg_a14; + b0 = b0 + reg_a15; + b0 = b0 + reg_a16; + + reg_a0 = reg_a0 + b0; + b0 = b0 + reg_a0 +}. + +Definition some_regs_asm: list Instruction. + let r := eval cbv in (compile compile_ext_call &[,some_regs]) in set (res := r). + match goal with + | res := Success (?x, _, _) |- _ => exact x + end. +Defined. + +Module PrintSomeRegsAsm. + Import riscv.Utility.InstructionNotations. + Goal True. let r := eval unfold some_regs_asm in some_regs_asm in idtac r. Abort. +End PrintSomeRegsAsm. + +Definition long_some_regs := func! (reg_a0, b0) ~> (reg_a19, b19) { + reg_a19 = $0; + b19 = $0; + + swap(reg_a0, b0); + reg_a1 = reg_a0 + b0 * b0; + swap(reg_a1, b0); + reg_a12 = reg_a1 - reg_a0 - b0; + b12 = reg_a1 + reg_a12 * b0; + reg_a5 = reg_a12 * b12; + reg_a9 = reg_a5 - reg_a0 - b0; + reg_a4 = reg_a1; + b3 = reg_a4 + reg_a12 * b0; + b4 = reg_a4; + reg_a10 = reg_a1 - reg_a0 - b4; + b12 = reg_a1 + reg_a12 * b0; + reg_a3 = reg_a1; + reg_a6 = reg_a10 + b12; + reg_a15 = reg_a3 - reg_a0 - b0; + a2 = reg_a1; + b12 = a2 + reg_a12 * b0; + reg_a11 = reg_a10 - reg_a0 - b0; + b5 = reg_a1; + b4 = reg_a6 + reg_a12 * b5; + a2 = reg_a11 - reg_a0 - b0; + b12 = reg_a1 + reg_a12 * b0; + reg_a16 = reg_a12 - reg_a0 - b0; + reg_a7 = a2; + b14 = reg_a7 + reg_a12 * b0; + reg_a13 = b4; + reg_a17 = $2; + b6 = $33; + reg_a12 = reg_a13 - reg_a17 - b6; + b12 = reg_a1 + reg_a12 * b0; + reg_a18 = $18; + b19 = $19; + reg_a17 = reg_a1 - reg_a18 - b19; + b12 = reg_a1 + reg_a9 * b0; + reg_a14 = $14; + reg_a12 = reg_a14 - reg_a0 - b0; + reg_a8 = b12; + b13 = reg_a8 + reg_a12 * b0; + reg_a12 = reg_a1 - reg_a0 - b0; + b2 = reg_a1 + reg_a12 * b0; + reg_a18 = reg_a1 - reg_a0 - b0; + b16 = $23; + b12 = reg_a15 + reg_a8 * b16; + reg_a12 = reg_a1 - reg_a0 - b3; + b12 = reg_a1 + reg_a12 * b0; + + reg_a19 = reg_a1 - reg_a0 - b0; + b19 = reg_a9 + reg_a7 * b0; + + b0 = b0 + reg_a1; + b0 = b0 + reg_a12; + b0 = b0 + b12; + b0 = b0 + reg_a5; + b0 = b0 + reg_a9; + b0 = b0 + reg_a4; + b0 = b0 + b3; + b0 = b0 + b4; + b0 = b0 + reg_a10; + b0 = b0 + b12; + b0 = b0 + reg_a3; + b0 = b0 + reg_a6; + b0 = b0 + reg_a15; + b0 = b0 + b12; + b0 = b0 + reg_a11; + b0 = b0 + b5; + b0 = b0 + b4; + b0 = b0 + b12; + b0 = b0 + reg_a16; + b0 = b0 + reg_a7; + b0 = b0 + b14; + b0 = b0 + reg_a13; + b0 = b0 + reg_a17; + b0 = b0 + b6; + b0 = b0 + reg_a12; + b0 = b0 + b12; + b0 = b0 + reg_a18; + b0 = b0 + b19; + b0 = b0 + reg_a17; + b0 = b0 + b12; + b0 = b0 + reg_a14; + b0 = b0 + reg_a12; + b0 = b0 + reg_a8; + b0 = b0 + b13; + b0 = b0 + reg_a12; + b0 = b0 + b2; + b0 = b0 + reg_a18; + b0 = b0 + b16; + b0 = b0 + b12; + b0 = b0 + reg_a12; + b0 = b0 + b12; + b0 = b0 + reg_a19; + b0 = b0 + b19; + + reg_a19 = b0; + b19 = b0 +}. + +Definition long_some_regs_asm: list Instruction. + let r := eval cbv in (compile compile_ext_call &[,swap;long_some_regs]) in set (res := r). + match goal with + | res := Success (?x, _, _) |- _ => exact x + end. +Defined. + +Module PrintLongSomeRegsAsm. + Import riscv.Utility.InstructionNotations. + Goal True. let r := eval unfold long_some_regs_asm in long_some_regs_asm in idtac r. Abort. +End PrintLongSomeRegsAsm. + + +Definition long_no_regs := func! (a0, b0) ~> (a19, b19) { + a19 = $0; + b19 = $0; + + swap(a0, b0); + a1 = a0 + b0 * b0; + swap(a1, b0); + a12 = a1 - a0 - b0; + b12 = a1 + a12 * b0; + a5 = a12 * b12; + a9 = a5 - a0 - b0; + a4 = a1; + b3 = a4 + a12 * b0; + b4 = a4; + a10 = a1 - a0 - b4; + b12 = a1 + a12 * b0; + a3 = a1; + a6 = a10 + b12; + a15 = a3 - a0 - b0; + a2 = a1; + b12 = a2 + a12 * b0; + a11 = a10 - a0 - b0; + b5 = a1; + b4 = a6 + a12 * b5; + a2 = a11 - a0 - b0; + b12 = a1 + a12 * b0; + a16 = a12 - a0 - b0; + a7 = a2; + b14 = a7 + a12 * b0; + a13 = b4; + a17 = $2; + b6 = $33; + a12 = a13 - a17 - b6; + b12 = a1 + a12 * b0; + a18 = $18; + b19 = $19; + a17 = a1 - a18 - b19; + b12 = a1 + a9 * b0; + a14 = $14; + a12 = a14 - a0 - b0; + a8 = b12; + b13 = a8 + a12 * b0; + a12 = a1 - a0 - b0; + b2 = a1 + a12 * b0; + a18 = a1 - a0 - b0; + b16 = $23; + b12 = a15 + a8 * b16; + a12 = a1 - a0 - b3; + b12 = a1 + a12 * b0; + + a19 = a1 - a0 - b0; + b19 = a9 + a7 * b0; + + b0 = b0 + a1; + b0 = b0 + a12; + b0 = b0 + b12; + b0 = b0 + a5; + b0 = b0 + a9; + b0 = b0 + a4; + b0 = b0 + b3; + b0 = b0 + b4; + b0 = b0 + a10; + b0 = b0 + b12; + b0 = b0 + a3; + b0 = b0 + a6; + b0 = b0 + a15; + b0 = b0 + b12; + b0 = b0 + a11; + b0 = b0 + b5; + b0 = b0 + b4; + b0 = b0 + b12; + b0 = b0 + a16; + b0 = b0 + a7; + b0 = b0 + b14; + b0 = b0 + a13; + b0 = b0 + a17; + b0 = b0 + b6; + b0 = b0 + a12; + b0 = b0 + b12; + b0 = b0 + a18; + b0 = b0 + b19; + b0 = b0 + a17; + b0 = b0 + b12; + b0 = b0 + a14; + b0 = b0 + a12; + b0 = b0 + a8; + b0 = b0 + b13; + b0 = b0 + a12; + b0 = b0 + b2; + b0 = b0 + a18; + b0 = b0 + b16; + b0 = b0 + b12; + b0 = b0 + a12; + b0 = b0 + b12; + b0 = b0 + a19; + b0 = b0 + b19; + + a19 = b0; + b19 = b0 +}. + +Definition long_no_regs_asm: list Instruction. + let r := eval cbv in (compile compile_ext_call &[,swap;long_no_regs]) in set (res := r). + match goal with + | res := Success (?x, _, _) |- _ => exact x + end. +Defined. + +Module PrintLongNoRegsAsm. + Import riscv.Utility.InstructionNotations. + Goal True. let r := eval unfold long_no_regs_asm in long_no_regs_asm in idtac r. Abort. +End PrintLongNoRegsAsm. + + + + +(* Definition long_lived_in_regs := *) +(* func! (reg_a0) ~> (reg_a0) { *) +(* reg_a1 = reg_a0 + $1; *) +(* reg_a2 = reg_a0 + $2; *) +(* reg_a3 = reg_a0 + $3; *) +(* reg_a4 = reg_a0 + $4; *) +(* reg_a5 = reg_a0 + $5; *) +(* reg_a6 = reg_a0 + $6; *) +(* reg_a7 = reg_a0 + $7; *) +(* reg_a8 = reg_a0 + $8; *) +(* reg_a9 = reg_a0 + $9; *) +(* reg_a10 = reg_a0 + $10; *) +(* reg_a11 = reg_a0 + $11; *) +(* reg_a12 = reg_a0 + $12; *) +(* reg_a13 = reg_a0 + $13; *) +(* reg_a14 = reg_a0 + $14; *) +(* reg_a15 = reg_a0 + $15; *) +(* reg_a16 = reg_a0 + $16; *) + +(* b1 = b1 + reg_a1; *) +(* reg_a1 = reg_a1 + b1; *) +(* b2 = b2 + reg_a2; *) +(* reg_a2 = reg_a2 + b2; *) +(* b3 = b3 + reg_a3; *) +(* reg_a3 = reg_a3 + b3; *) +(* b4 = b4 + reg_a4; *) +(* reg_a4 = reg_a4 + b4; *) +(* b5 = b5 + reg_a5; *) +(* reg_a5 = reg_a5 + b5; *) +(* b6 = b6 + reg_a6; *) +(* reg_a6 = reg_a6 + b6; *) +(* b7 = b7 + reg_a7; *) +(* reg_a7 = reg_a7 + b7; *) +(* b8 = b8 + reg_a8; *) +(* reg_a8 = reg_a8 + b8; *) +(* b9 = b9 + reg_a9; *) +(* reg_a9 = reg_a9 + b9; *) +(* b10 = b10 + reg_a10; *) +(* reg_a10 = reg_a10 + b10; *) +(* b11 = b11 + reg_a11; *) +(* reg_a11 = reg_a11 + b11; *) +(* b12 = b12 + reg_a12; *) +(* reg_a12 = reg_a12 + b12; *) +(* b13 = b13 + reg_a13; *) +(* reg_a13 = reg_a13 + b13; *) +(* b14 = b14 + reg_a14; *) +(* reg_a14 = reg_a14 + b14; *) +(* b15 = b15 + reg_a15; *) +(* reg_a15 = reg_a15 + b15; *) +(* b16 = b16 + reg_a16; *) +(* reg_a16 = reg_a16 + b16; *) + +(* reg_a0 = reg_a16 *) +(* }. *) + +(* Definition long_lived_in_regs_asm: list Instruction. *) +(* let r := eval cbv in (compile compile_ext_call &[,long_lived_in_regs]) in set (res := r). *) +(* match goal with *) +(* | res := Success (?x, _, _) |- _ => exact x *) +(* end. *) +(* Defined. *) + +(* Module PrintLongLivedInRegsAsm. *) +(* Import riscv.Utility.InstructionNotations. *) +(* Goal True. let r := eval unfold long_lived_in_regs_asm in long_lived_in_regs_asm in idtac r. Abort. *) +(* End PrintLongLivedInRegsAsm. *) + + + + + + From c4213c436aed49d9f6ab3dcfe9183b8d9471b096 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Mon, 8 Jul 2024 22:16:55 +0100 Subject: [PATCH 45/62] Revert "add some simple tests for the new heuristic" This reverts commit 44380e0e183c7ee7b75f50803176194c68d007ad. --- compiler/src/compilerExamples/RegAllocTests.v | 568 ------------------ 1 file changed, 568 deletions(-) delete mode 100644 compiler/src/compilerExamples/RegAllocTests.v diff --git a/compiler/src/compilerExamples/RegAllocTests.v b/compiler/src/compilerExamples/RegAllocTests.v deleted file mode 100644 index 307bcaf14..000000000 --- a/compiler/src/compilerExamples/RegAllocTests.v +++ /dev/null @@ -1,568 +0,0 @@ -Require Import bedrock2.NotationsCustomEntry. -Require Import Coq.Lists.List. -Import ListNotations. -Require bedrock2Examples.Demos. -Require Import coqutil.Decidable. -Require Import bedrock2.NotationsCustomEntry coqutil.Macros.WithBaseName. -Require Import compiler.ExprImp. -Require Import compiler.NameGen. -Require Import compiler.Pipeline. -Require Import riscv.Spec.Decode. -Require Import riscv.Utility.Words32Naive. -Require Import riscv.Utility.DefaultMemImpl32. -Require Import riscv.Utility.Monads. -Require Import compiler.util.Common. -Require Import coqutil.Decidable. -Require riscv.Utility.InstructionNotations. -Require Import riscv.Platform.MinimalLogging. -Require Import bedrock2.MetricLogging. -Require Import riscv.Platform.MetricMinimal. -Require Import riscv.Utility.Utility. -Require Import riscv.Utility.Encode. -Require Import coqutil.Map.SortedList. -Require Import compiler.MemoryLayout. -Require Import compiler.StringNameGen. -Require Import riscv.Utility.InstructionCoercions. -Require Import riscv.Platform.MetricRiscvMachine. -Require bedrock2.Hexdump. - -Open Scope Z_scope. -Open Scope string_scope. -Open Scope ilist_scope. - -(* ********************************************************* *) -(* Preliminaries *) -(* ********************************************************* *) - -Definition var: Set := Z. -Definition Reg: Set := Z. - - -Local Existing Instance DefaultRiscvState. - -Axiom TODO: forall {T: Type}, T. - -Local Instance funpos_env: map.map string Z := SortedListString.map _. - -Definition compile_ext_call(posenv: funpos_env)(mypos stackoffset: Z)(s: FlatImp.stmt Z) := - match s with - | FlatImp.SInteract _ fname _ => - if string_dec fname "nop" then - [[Addi Register0 Register0 0]] - else - nil - | _ => [] - end. - -Notation RiscvMachine := MetricRiscvMachine. - -Local Existing Instance coqutil.Map.SortedListString.map. -Local Existing Instance coqutil.Map.SortedListString.ok. - -Definition main_stackalloc := func! { - stackalloc 4 as x; stackalloc 4 as y; swap_swap(x, y) }. - - -(* stack grows from high addreses to low addresses, first stack word will be written to - (stack_pastend-8), next stack word to (stack_pastend-16) etc *) -Definition stack_pastend: Z := 2048. - -Lemma f_equal2: forall {A B: Type} {f1 f2: A -> B} {a1 a2: A}, - f1 = f2 -> a1 = a2 -> f1 a1 = f2 a2. -Proof. intros. congruence. Qed. - -Lemma f_equal3: forall {A B C: Type} {f1 f2: A -> B -> C} {a1 a2: A} {b1 b2: B}, - f1 = f2 -> a1 = a2 -> b1 = b2 -> f1 a1 b1 = f2 a2 b2. -Proof. intros. congruence. Qed. - -Lemma f_equal3_dep: forall {A B C: Type} {f1 f2: A -> B -> C} {a1 a2: A} {b1 b2: B}, - f1 = f2 -> a1 = a2 -> b1 = b2 -> f1 a1 b1 = f2 a2 b2. -Proof. intros. congruence. Qed. - - -Local Instance RV32I_bitwidth: FlatToRiscvCommon.bitwidth_iset 32 RV32I. -Proof. reflexivity. Qed. - - - - -(* ********************************************************* *) -(* Tests *) -(* ********************************************************* *) - -Definition swap := func! (a, b) { - t = load(b); - store(b, load(a)); - store(a, t) -}. - -(* Based on `long1` from compilerExamples/SpillingTests.v *) -(* Expected to fail, since too many `reg` variables are live simultaneously *) -Definition too_many_regs := func! (reg_a0, reg_b0) ~> (reg_a19, reg_b19) { - (* TODO once spilling is improved, test that it still works if these two lines are removed - (currently the test that all resnames are <32 fails if these two lines are removed) *) - reg_a19 = $0; - reg_b19 = $0; - - swap(reg_a0, reg_b0); - reg_a1 = reg_a0 + reg_b0 * reg_b0; - swap(reg_a1, reg_b0); - reg_a12 = reg_a1 - reg_a0 - reg_b0; - reg_b12 = reg_a1 + reg_a12 * reg_b0; - reg_a5 = reg_a12 * reg_b12; - reg_a9 = reg_a5 - reg_a0 - reg_b0; - reg_a4 = reg_a1; - reg_b3 = reg_a4 + reg_a12 * reg_b0; - reg_b4 = reg_a4; - reg_a10 = reg_a1 - reg_a0 - reg_b4; - reg_b12 = reg_a1 + reg_a12 * reg_b0; - reg_a3 = reg_a1; - reg_a6 = reg_a10 + reg_b12; - reg_a15 = reg_a3 - reg_a0 - reg_b0; - reg_a2 = reg_a1; - reg_b12 = reg_a2 + reg_a12 * reg_b0; - reg_a11 = reg_a10 - reg_a0 - reg_b0; - reg_b5 = reg_a1; - reg_b4 = reg_a6 + reg_a12 * reg_b5; - reg_a2 = reg_a11 - reg_a0 - reg_b0; - reg_b12 = reg_a1 + reg_a12 * reg_b0; - reg_a16 = reg_a12 - reg_a0 - reg_b0; - reg_a7 = reg_a2; - reg_b14 = reg_a7 + reg_a12 * reg_b0; - reg_a13 = reg_b4; - reg_a17 = $2; - reg_b6 = $33; - reg_a12 = reg_a13 - reg_a17 - reg_b6; - reg_b12 = reg_a1 + reg_a12 * reg_b0; - reg_a18 = $18; - reg_b19 = $19; - reg_a17 = reg_a1 - reg_a18 - reg_b19; - reg_b12 = reg_a1 + reg_a9 * reg_b0; - reg_a14 = $14; - reg_a12 = reg_a14 - reg_a0 - reg_b0; - reg_a8 = reg_b12; - reg_b13 = reg_a8 + reg_a12 * reg_b0; - reg_a12 = reg_a1 - reg_a0 - reg_b0; - reg_b2 = reg_a1 + reg_a12 * reg_b0; - reg_a18 = reg_a1 - reg_a0 - reg_b0; - reg_b16 = $23; - reg_b12 = reg_a15 + reg_a8 * reg_b16; - reg_a12 = reg_a1 - reg_a0 - reg_b3; - reg_b12 = reg_a1 + reg_a12 * reg_b0; - - reg_a19 = reg_a1 - reg_a0 - reg_b0; - reg_b19 = reg_a9 + reg_a7 * reg_b0; - - reg_b0 = reg_b0 + reg_a1; - reg_b0 = reg_b0 + reg_a12; - reg_b0 = reg_b0 + reg_b12; - reg_b0 = reg_b0 + reg_a5; - reg_b0 = reg_b0 + reg_a9; - reg_b0 = reg_b0 + reg_a4; - reg_b0 = reg_b0 + reg_b3; - reg_b0 = reg_b0 + reg_b4; - reg_b0 = reg_b0 + reg_a10; - reg_b0 = reg_b0 + reg_b12; - reg_b0 = reg_b0 + reg_a3; - reg_b0 = reg_b0 + reg_a6; - reg_b0 = reg_b0 + reg_a15; - reg_b0 = reg_b0 + reg_a2; - reg_b0 = reg_b0 + reg_b12; - reg_b0 = reg_b0 + reg_a11; - reg_b0 = reg_b0 + reg_b5; - reg_b0 = reg_b0 + reg_b4; - reg_b0 = reg_b0 + reg_a2; - reg_b0 = reg_b0 + reg_b12; - reg_b0 = reg_b0 + reg_a16; - reg_b0 = reg_b0 + reg_a7; - reg_b0 = reg_b0 + reg_b14; - reg_b0 = reg_b0 + reg_a13; - reg_b0 = reg_b0 + reg_a17; - reg_b0 = reg_b0 + reg_b6; - reg_b0 = reg_b0 + reg_a12; - reg_b0 = reg_b0 + reg_b12; - reg_b0 = reg_b0 + reg_a18; - reg_b0 = reg_b0 + reg_b19; - reg_b0 = reg_b0 + reg_a17; - reg_b0 = reg_b0 + reg_b12; - reg_b0 = reg_b0 + reg_a14; - reg_b0 = reg_b0 + reg_a12; - reg_b0 = reg_b0 + reg_a8; - reg_b0 = reg_b0 + reg_b13; - reg_b0 = reg_b0 + reg_a12; - reg_b0 = reg_b0 + reg_b2; - reg_b0 = reg_b0 + reg_a18; - reg_b0 = reg_b0 + reg_b16; - reg_b0 = reg_b0 + reg_b12; - reg_b0 = reg_b0 + reg_a12; - reg_b0 = reg_b0 + reg_b12; - reg_b0 = reg_b0 + reg_a19; - reg_b0 = reg_b0 + reg_b19; - - reg_a19 = reg_b0; - reg_b19 = reg_b0 -}. - -Remark too_many_regs_fails: - exists x, compile compile_ext_call &[,too_many_regs] = Failure x. -Proof. - (* let r := eval cbv in (compile compile_ext_call &[,too_many_regs]) in set (res := r). *) - eexists; cbv; reflexivity. -Qed. - - -Definition some_regs := - func! (reg_a0, b0) ~> (reg_a0, b0) { - reg_a1 = reg_a0 + $1; - reg_a2 = reg_a0 + $2; - reg_a3 = reg_a0 + $3; - reg_a4 = reg_a0 + $4; - reg_a5 = reg_a0 + $5; - reg_a6 = reg_a0 + $6; - reg_a7 = reg_a0 + $7; - reg_a8 = reg_a0 + $8; - reg_a9 = reg_a0 + $9; - reg_a10 = reg_a0 + $10; - reg_a11 = reg_a0 + $11; - reg_a12 = reg_a0 + $12; - reg_a13 = reg_a0 + $13; - reg_a14 = reg_a0 + $14; - reg_a15 = reg_a0 + $15; - reg_a16 = reg_a0 + $16; - - b0 = b0 + reg_a0; - b0 = b0 + reg_a1; - b0 = b0 + reg_a2; - b0 = b0 + reg_a3; - b0 = b0 + reg_a4; - b0 = b0 + reg_a5; - b0 = b0 + reg_a6; - b0 = b0 + reg_a7; - b0 = b0 + reg_a8; - b0 = b0 + reg_a9; - b0 = b0 + reg_a10; - b0 = b0 + reg_a11; - b0 = b0 + reg_a12; - b0 = b0 + reg_a13; - b0 = b0 + reg_a14; - b0 = b0 + reg_a15; - b0 = b0 + reg_a16; - - reg_a0 = reg_a0 + b0; - b0 = b0 + reg_a0 -}. - -Definition some_regs_asm: list Instruction. - let r := eval cbv in (compile compile_ext_call &[,some_regs]) in set (res := r). - match goal with - | res := Success (?x, _, _) |- _ => exact x - end. -Defined. - -Module PrintSomeRegsAsm. - Import riscv.Utility.InstructionNotations. - Goal True. let r := eval unfold some_regs_asm in some_regs_asm in idtac r. Abort. -End PrintSomeRegsAsm. - -Definition long_some_regs := func! (reg_a0, b0) ~> (reg_a19, b19) { - reg_a19 = $0; - b19 = $0; - - swap(reg_a0, b0); - reg_a1 = reg_a0 + b0 * b0; - swap(reg_a1, b0); - reg_a12 = reg_a1 - reg_a0 - b0; - b12 = reg_a1 + reg_a12 * b0; - reg_a5 = reg_a12 * b12; - reg_a9 = reg_a5 - reg_a0 - b0; - reg_a4 = reg_a1; - b3 = reg_a4 + reg_a12 * b0; - b4 = reg_a4; - reg_a10 = reg_a1 - reg_a0 - b4; - b12 = reg_a1 + reg_a12 * b0; - reg_a3 = reg_a1; - reg_a6 = reg_a10 + b12; - reg_a15 = reg_a3 - reg_a0 - b0; - a2 = reg_a1; - b12 = a2 + reg_a12 * b0; - reg_a11 = reg_a10 - reg_a0 - b0; - b5 = reg_a1; - b4 = reg_a6 + reg_a12 * b5; - a2 = reg_a11 - reg_a0 - b0; - b12 = reg_a1 + reg_a12 * b0; - reg_a16 = reg_a12 - reg_a0 - b0; - reg_a7 = a2; - b14 = reg_a7 + reg_a12 * b0; - reg_a13 = b4; - reg_a17 = $2; - b6 = $33; - reg_a12 = reg_a13 - reg_a17 - b6; - b12 = reg_a1 + reg_a12 * b0; - reg_a18 = $18; - b19 = $19; - reg_a17 = reg_a1 - reg_a18 - b19; - b12 = reg_a1 + reg_a9 * b0; - reg_a14 = $14; - reg_a12 = reg_a14 - reg_a0 - b0; - reg_a8 = b12; - b13 = reg_a8 + reg_a12 * b0; - reg_a12 = reg_a1 - reg_a0 - b0; - b2 = reg_a1 + reg_a12 * b0; - reg_a18 = reg_a1 - reg_a0 - b0; - b16 = $23; - b12 = reg_a15 + reg_a8 * b16; - reg_a12 = reg_a1 - reg_a0 - b3; - b12 = reg_a1 + reg_a12 * b0; - - reg_a19 = reg_a1 - reg_a0 - b0; - b19 = reg_a9 + reg_a7 * b0; - - b0 = b0 + reg_a1; - b0 = b0 + reg_a12; - b0 = b0 + b12; - b0 = b0 + reg_a5; - b0 = b0 + reg_a9; - b0 = b0 + reg_a4; - b0 = b0 + b3; - b0 = b0 + b4; - b0 = b0 + reg_a10; - b0 = b0 + b12; - b0 = b0 + reg_a3; - b0 = b0 + reg_a6; - b0 = b0 + reg_a15; - b0 = b0 + b12; - b0 = b0 + reg_a11; - b0 = b0 + b5; - b0 = b0 + b4; - b0 = b0 + b12; - b0 = b0 + reg_a16; - b0 = b0 + reg_a7; - b0 = b0 + b14; - b0 = b0 + reg_a13; - b0 = b0 + reg_a17; - b0 = b0 + b6; - b0 = b0 + reg_a12; - b0 = b0 + b12; - b0 = b0 + reg_a18; - b0 = b0 + b19; - b0 = b0 + reg_a17; - b0 = b0 + b12; - b0 = b0 + reg_a14; - b0 = b0 + reg_a12; - b0 = b0 + reg_a8; - b0 = b0 + b13; - b0 = b0 + reg_a12; - b0 = b0 + b2; - b0 = b0 + reg_a18; - b0 = b0 + b16; - b0 = b0 + b12; - b0 = b0 + reg_a12; - b0 = b0 + b12; - b0 = b0 + reg_a19; - b0 = b0 + b19; - - reg_a19 = b0; - b19 = b0 -}. - -Definition long_some_regs_asm: list Instruction. - let r := eval cbv in (compile compile_ext_call &[,swap;long_some_regs]) in set (res := r). - match goal with - | res := Success (?x, _, _) |- _ => exact x - end. -Defined. - -Module PrintLongSomeRegsAsm. - Import riscv.Utility.InstructionNotations. - Goal True. let r := eval unfold long_some_regs_asm in long_some_regs_asm in idtac r. Abort. -End PrintLongSomeRegsAsm. - - -Definition long_no_regs := func! (a0, b0) ~> (a19, b19) { - a19 = $0; - b19 = $0; - - swap(a0, b0); - a1 = a0 + b0 * b0; - swap(a1, b0); - a12 = a1 - a0 - b0; - b12 = a1 + a12 * b0; - a5 = a12 * b12; - a9 = a5 - a0 - b0; - a4 = a1; - b3 = a4 + a12 * b0; - b4 = a4; - a10 = a1 - a0 - b4; - b12 = a1 + a12 * b0; - a3 = a1; - a6 = a10 + b12; - a15 = a3 - a0 - b0; - a2 = a1; - b12 = a2 + a12 * b0; - a11 = a10 - a0 - b0; - b5 = a1; - b4 = a6 + a12 * b5; - a2 = a11 - a0 - b0; - b12 = a1 + a12 * b0; - a16 = a12 - a0 - b0; - a7 = a2; - b14 = a7 + a12 * b0; - a13 = b4; - a17 = $2; - b6 = $33; - a12 = a13 - a17 - b6; - b12 = a1 + a12 * b0; - a18 = $18; - b19 = $19; - a17 = a1 - a18 - b19; - b12 = a1 + a9 * b0; - a14 = $14; - a12 = a14 - a0 - b0; - a8 = b12; - b13 = a8 + a12 * b0; - a12 = a1 - a0 - b0; - b2 = a1 + a12 * b0; - a18 = a1 - a0 - b0; - b16 = $23; - b12 = a15 + a8 * b16; - a12 = a1 - a0 - b3; - b12 = a1 + a12 * b0; - - a19 = a1 - a0 - b0; - b19 = a9 + a7 * b0; - - b0 = b0 + a1; - b0 = b0 + a12; - b0 = b0 + b12; - b0 = b0 + a5; - b0 = b0 + a9; - b0 = b0 + a4; - b0 = b0 + b3; - b0 = b0 + b4; - b0 = b0 + a10; - b0 = b0 + b12; - b0 = b0 + a3; - b0 = b0 + a6; - b0 = b0 + a15; - b0 = b0 + b12; - b0 = b0 + a11; - b0 = b0 + b5; - b0 = b0 + b4; - b0 = b0 + b12; - b0 = b0 + a16; - b0 = b0 + a7; - b0 = b0 + b14; - b0 = b0 + a13; - b0 = b0 + a17; - b0 = b0 + b6; - b0 = b0 + a12; - b0 = b0 + b12; - b0 = b0 + a18; - b0 = b0 + b19; - b0 = b0 + a17; - b0 = b0 + b12; - b0 = b0 + a14; - b0 = b0 + a12; - b0 = b0 + a8; - b0 = b0 + b13; - b0 = b0 + a12; - b0 = b0 + b2; - b0 = b0 + a18; - b0 = b0 + b16; - b0 = b0 + b12; - b0 = b0 + a12; - b0 = b0 + b12; - b0 = b0 + a19; - b0 = b0 + b19; - - a19 = b0; - b19 = b0 -}. - -Definition long_no_regs_asm: list Instruction. - let r := eval cbv in (compile compile_ext_call &[,swap;long_no_regs]) in set (res := r). - match goal with - | res := Success (?x, _, _) |- _ => exact x - end. -Defined. - -Module PrintLongNoRegsAsm. - Import riscv.Utility.InstructionNotations. - Goal True. let r := eval unfold long_no_regs_asm in long_no_regs_asm in idtac r. Abort. -End PrintLongNoRegsAsm. - - - - -(* Definition long_lived_in_regs := *) -(* func! (reg_a0) ~> (reg_a0) { *) -(* reg_a1 = reg_a0 + $1; *) -(* reg_a2 = reg_a0 + $2; *) -(* reg_a3 = reg_a0 + $3; *) -(* reg_a4 = reg_a0 + $4; *) -(* reg_a5 = reg_a0 + $5; *) -(* reg_a6 = reg_a0 + $6; *) -(* reg_a7 = reg_a0 + $7; *) -(* reg_a8 = reg_a0 + $8; *) -(* reg_a9 = reg_a0 + $9; *) -(* reg_a10 = reg_a0 + $10; *) -(* reg_a11 = reg_a0 + $11; *) -(* reg_a12 = reg_a0 + $12; *) -(* reg_a13 = reg_a0 + $13; *) -(* reg_a14 = reg_a0 + $14; *) -(* reg_a15 = reg_a0 + $15; *) -(* reg_a16 = reg_a0 + $16; *) - -(* b1 = b1 + reg_a1; *) -(* reg_a1 = reg_a1 + b1; *) -(* b2 = b2 + reg_a2; *) -(* reg_a2 = reg_a2 + b2; *) -(* b3 = b3 + reg_a3; *) -(* reg_a3 = reg_a3 + b3; *) -(* b4 = b4 + reg_a4; *) -(* reg_a4 = reg_a4 + b4; *) -(* b5 = b5 + reg_a5; *) -(* reg_a5 = reg_a5 + b5; *) -(* b6 = b6 + reg_a6; *) -(* reg_a6 = reg_a6 + b6; *) -(* b7 = b7 + reg_a7; *) -(* reg_a7 = reg_a7 + b7; *) -(* b8 = b8 + reg_a8; *) -(* reg_a8 = reg_a8 + b8; *) -(* b9 = b9 + reg_a9; *) -(* reg_a9 = reg_a9 + b9; *) -(* b10 = b10 + reg_a10; *) -(* reg_a10 = reg_a10 + b10; *) -(* b11 = b11 + reg_a11; *) -(* reg_a11 = reg_a11 + b11; *) -(* b12 = b12 + reg_a12; *) -(* reg_a12 = reg_a12 + b12; *) -(* b13 = b13 + reg_a13; *) -(* reg_a13 = reg_a13 + b13; *) -(* b14 = b14 + reg_a14; *) -(* reg_a14 = reg_a14 + b14; *) -(* b15 = b15 + reg_a15; *) -(* reg_a15 = reg_a15 + b15; *) -(* b16 = b16 + reg_a16; *) -(* reg_a16 = reg_a16 + b16; *) - -(* reg_a0 = reg_a16 *) -(* }. *) - -(* Definition long_lived_in_regs_asm: list Instruction. *) -(* let r := eval cbv in (compile compile_ext_call &[,long_lived_in_regs]) in set (res := r). *) -(* match goal with *) -(* | res := Success (?x, _, _) |- _ => exact x *) -(* end. *) -(* Defined. *) - -(* Module PrintLongLivedInRegsAsm. *) -(* Import riscv.Utility.InstructionNotations. *) -(* Goal True. let r := eval unfold long_lived_in_regs_asm in long_lived_in_regs_asm in idtac r. Abort. *) -(* End PrintLongLivedInRegsAsm. *) - - - - - - From 105c5ea0399252be0299ace16648dc722d34c1e4 Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Mon, 8 Jul 2024 22:17:12 +0100 Subject: [PATCH 46/62] Revert "keep old code in a comment for easier diffing" This reverts commit b38989916a93764417a99627c242cfa635ca2ddc. --- compiler/src/compiler/RegAlloc.v | 1 - 1 file changed, 1 deletion(-) diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index 0bf8d8564..b60073cc0 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -325,7 +325,6 @@ Definition isRegInterval (l: lifetime) : bool := Definition events_to_corresp(events: list event): list (srcvar * impvar) := let sorted_intervals := sort compare_interval_length (events_to_intervals 0 events) in - (* let occ := List.fold_left assign_srcvar sorted_intervals [] in *) let (reg_intervals, nonreg_intervals) := List.partition isRegInterval sorted_intervals in let regfirst_intervals := reg_intervals ++ nonreg_intervals in let occ := List.fold_left assign_srcvar regfirst_intervals [] in From 2c29dcefbd286f9b1e7f2e5458d6a871f411a9fc Mon Sep 17 00:00:00 2001 From: Pratap Singh Date: Mon, 8 Jul 2024 22:17:31 +0100 Subject: [PATCH 47/62] Revert "account for isReg in regalloc heuristic" This reverts commit e5ba50c2148b2cd0bdb12ce6bdabd8eb76285d02. --- compiler/src/compiler/RegAlloc.v | 9 +-------- 1 file changed, 1 insertion(+), 8 deletions(-) diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index b60073cc0..b5cfe57f5 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -318,16 +318,9 @@ Fixpoint corresp_of_impvars(start_impvar: Z)(o: occupancy): list (srcvar * impva | nil => nil end. -Definition isRegInterval (l: lifetime) : bool := - match l with - | (srcvarname, _) => isRegStr srcvarname - end. - Definition events_to_corresp(events: list event): list (srcvar * impvar) := let sorted_intervals := sort compare_interval_length (events_to_intervals 0 events) in - let (reg_intervals, nonreg_intervals) := List.partition isRegInterval sorted_intervals in - let regfirst_intervals := reg_intervals ++ nonreg_intervals in - let occ := List.fold_left assign_srcvar regfirst_intervals [] in + let occ := List.fold_left assign_srcvar sorted_intervals [] in corresp_of_impvars first_available_impvar occ. From 7d7417982e56ea3b9d5e625a183514c9b50d8f97 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Fri, 5 Jul 2024 06:45:44 -0400 Subject: [PATCH 48/62] fix ipow metrics proof --- bedrock2/src/bedrock2Examples/metric_ipow.v | 29 +++++++++++---------- 1 file changed, 15 insertions(+), 14 deletions(-) diff --git a/bedrock2/src/bedrock2Examples/metric_ipow.v b/bedrock2/src/bedrock2Examples/metric_ipow.v index 7312585f1..639b98b8d 100644 --- a/bedrock2/src/bedrock2Examples/metric_ipow.v +++ b/bedrock2/src/bedrock2Examples/metric_ipow.v @@ -1,6 +1,7 @@ Require Import Coq.ZArith.ZArith coqutil.Z.div_mod_to_equations. Require Import bedrock2.NotationsCustomEntry. Require Import bedrock2.MetricLogging. +Require Import bedrock2.MetricCosts. Import Syntax BinInt String List.ListNotations ZArith. Require Import coqutil.Z.Lia. Local Open Scope string_scope. Local Open Scope Z_scope. Local Open Scope list_scope. @@ -19,9 +20,9 @@ From bedrock2 Require Import MetricLoops. From coqutil Require Import Word.Properties Word.Interface Tactics.letexists. Import Interface.word. -Definition initCost := {| instructions := 9; stores := 0; loads := 9; jumps := 0 |}. -Definition iterCost := {| instructions := 38; stores := 0; loads := 45; jumps := 2 |}. -Definition endCost := {| instructions := 2; stores := 0; loads := 3; jumps := 1 |}. +Definition initCost := {| instructions := 12; stores := 2; loads := 13; jumps := 0 |}. +Definition iterCost := {| instructions := 76; stores := 16; loads := 98; jumps := 2 |}. +Definition endCost := {| instructions := 6; stores := 1; loads := 9; jumps := 1 |}. Definition msb z := match z with | Zpos _ => Z.log2 z + 1 @@ -151,21 +152,18 @@ Proof. rewrite Z.mul_mod_idemp_r by discriminate. f_equal; ring. } { (* metrics correct *) + cbn in H4. rewrite <- (MetricLog_eq mc0) in H4. applyAddMetrics H4. + cbn in H4. rewrite msb_shift in H4 by blia. unfold iterCost in *. + rewrite MetricArith.mul_sub_distr_r in H4. + rewrite <- MetricArith.add_sub_swap in H4. + rewrite <- MetricArith.le_add_le_sub_r in H4. + eapply MetricArith.le_trans. + 2: exact H4. solve_MetricLog. - (*destruct mc0.*) - (*applyAddMetrics H4.*) - (*rewrite msb_shift in H4 by blia.*) - (*rewrite MetricArith.mul_sub_distr_r in H4.*) - (*rewrite <- MetricArith.add_sub_swap in H4.*) - (*rewrite <- MetricArith.le_add_le_sub_r in H4.*) - (*eapply MetricArith.le_trans.*) - (*2: exact H4.*) - (*unfold iterCost.*) - (*solve_MetricLog.*) } } { @@ -186,8 +184,10 @@ Proof. rewrite Z.mul_mod_idemp_r by discriminate. f_equal; ring. } { (* metrics correct *) + cbn in H4. rewrite <- (MetricLog_eq mc0) in H4. applyAddMetrics H4. + cbn in H4. rewrite msb_shift in H4 by blia. unfold iterCost in *. solve_MetricLog. @@ -199,6 +199,7 @@ Proof. split; [reflexivity|]. unfold msb, iterCost, endCost. subst brmc. + cbn. solve_MetricLog. } } @@ -207,5 +208,5 @@ Proof. repeat (split || letexists || t || trivial). { setoid_rewrite H1; setoid_rewrite Z.mul_1_l; trivial. } - all: unfold initCost, iterCost, endCost in *; solve_MetricLog. + all: cbn in H2; unfold initCost, iterCost, endCost in *; solve_MetricLog. Qed. From 74337b80ca5ba3a4f5473190451a67f39aaf8091 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Wed, 10 Jul 2024 09:17:58 -0400 Subject: [PATCH 49/62] additive style --- bedrock2/src/bedrock2/MetricCosts.v | 150 +++++++----------- bedrock2/src/bedrock2/MetricLogging.v | 8 +- compiler/src/compiler/FlatImp.v | 34 ++-- compiler/src/compiler/FlatToRiscvFunctions.v | 2 +- compiler/src/compiler/MMIO.v | 4 +- compiler/src/compiler/Pipeline.v | 6 +- compiler/src/compiler/RegAlloc.v | 53 ++----- compiler/src/compiler/Spilling.v | 119 ++++---------- .../memory_mapped_ext_calls_compiler.v | 4 +- 9 files changed, 126 insertions(+), 254 deletions(-) diff --git a/bedrock2/src/bedrock2/MetricCosts.v b/bedrock2/src/bedrock2/MetricCosts.v index b6f56d6b3..d6ec0d5e4 100644 --- a/bedrock2/src/bedrock2/MetricCosts.v +++ b/bedrock2/src/bedrock2/MetricCosts.v @@ -17,98 +17,98 @@ Section FlatImpExec. Definition cost_interact mc := match phase with - | PreSpill => (addMetricInstructions 100 (addMetricJumps 100 (addMetricStores 100 (addMetricLoads 100 mc)))) - | PostSpill => (addMetricInstructions 50 (addMetricJumps 50 (addMetricStores 50 (addMetricLoads 50 mc)))) - end. + | PreSpill => mkMetricLog 100 100 100 100 + | PostSpill => mkMetricLog 50 50 50 50 + end + mc. Definition cost_call_internal mc := match phase with - | PreSpill => addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc))) - | PostSpill => addMetricInstructions 50 (addMetricJumps 50 (addMetricLoads 50 (addMetricStores 50 mc))) - end. + | PreSpill => mkMetricLog 100 100 100 100 + | PostSpill => mkMetricLog 50 50 50 50 + end + mc. Definition cost_call_external mc := match phase with - | PreSpill => addMetricInstructions 100 (addMetricJumps 100 (addMetricLoads 100 (addMetricStores 100 mc))) - | PostSpill => addMetricInstructions 50 (addMetricJumps 50 (addMetricLoads 50 (addMetricStores 50 mc))) - end. + | PreSpill => mkMetricLog 100 100 100 100 + | PostSpill => mkMetricLog 50 50 50 50 + end + mc. (* TODO think about a non-fixed bound on the cost of function preamble and postamble *) Definition cost_load x a mc := match (isReg x, isReg a) with - | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricStores 1 mc))) - | (false, true) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) - | ( true, false) => (addMetricInstructions 2 (addMetricLoads 4 mc)) - | ( true, true) => (addMetricInstructions 1 (addMetricLoads 2 mc)) - end. + | (false, false) => mkMetricLog 3 1 5 0 + | (false, true) => mkMetricLog 2 1 3 0 + | ( true, false) => mkMetricLog 2 0 4 0 + | ( true, true) => mkMetricLog 1 0 2 0 + end + mc. Definition cost_store a v mc := match (isReg a, isReg v) with - | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricStores 1 mc))) - | (false, true) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) - | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricStores 1 mc))) - | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc))) - end. + | (false, false) => mkMetricLog 3 1 5 0 + | (false, true) => mkMetricLog 2 1 3 0 + | ( true, false) => mkMetricLog 2 1 3 0 + | ( true, true) => mkMetricLog 1 1 1 0 + end + mc. Definition cost_inlinetable x i mc := match (isReg x, isReg i) with - | (false, false) => (addMetricInstructions 5 (addMetricLoads 7 (addMetricStores 1 (addMetricJumps 1 mc)))) - | (false, true) => (addMetricInstructions 4 (addMetricLoads 5 (addMetricStores 1 (addMetricJumps 1 mc)))) - | ( true, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) - | ( true, true) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) - end. + | (false, false) => mkMetricLog 5 1 7 1 + | (false, true) => mkMetricLog 4 1 5 1 + | ( true, false) => mkMetricLog 4 0 6 1 + | ( true, true) => mkMetricLog 3 0 4 1 + end + mc. Definition cost_stackalloc x mc := match isReg x with - | false => (addMetricInstructions 2 (addMetricLoads 2 (addMetricStores 1 mc))) - | true => (addMetricInstructions 1 (addMetricLoads 1 mc)) - end. + | false => mkMetricLog 2 1 2 0 + | true => mkMetricLog 1 0 1 0 + end + mc. Definition cost_lit x mc := match isReg x with - | false => (addMetricInstructions 9 (addMetricLoads 9 (addMetricStores 1 mc))) - | true => (addMetricInstructions 8 (addMetricLoads 8 mc)) - end. + | false => mkMetricLog 9 1 9 0 + | true => mkMetricLog 8 0 8 0 + end + mc. Definition cost_op x y z mc := match (isReg x, isReg y, isReg z) with - | (false, false, false) => (addMetricInstructions 5 (addMetricLoads 7 (addMetricStores 1 mc))) - | (false, false, true) | (false, true, false) => (addMetricInstructions 4 (addMetricLoads 5 (addMetricStores 1 mc))) - | (false, true, true) => (addMetricInstructions 3 (addMetricLoads 3 (addMetricStores 1 mc))) - | ( true, false, false) => (addMetricInstructions 4 (addMetricLoads 6 mc)) - | ( true, false, true) | ( true, true, false) => (addMetricInstructions 3 (addMetricLoads 4 mc)) - | ( true, true, true) => (addMetricInstructions 2 (addMetricLoads 2 mc)) - end. + | (false, false, false) => mkMetricLog 5 1 7 0 + | (false, false, true) | (false, true, false) => mkMetricLog 4 1 5 0 + | (false, true, true) => mkMetricLog 3 1 3 0 + | ( true, false, false) => mkMetricLog 4 0 6 0 + | ( true, false, true) | ( true, true, false) => mkMetricLog 3 0 4 0 + | ( true, true, true) => mkMetricLog 2 0 2 0 + end + mc. Definition cost_set x y mc := match (isReg x, isReg y) with - | (false, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricStores 1 mc))) - | (false, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricStores 1 mc))) - | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 mc)) - | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 mc)) - end. + | (false, false) => mkMetricLog 3 1 4 0 + | (false, true) => mkMetricLog 2 1 2 0 + | ( true, false) => mkMetricLog 2 0 3 0 + | ( true, true) => mkMetricLog 1 0 1 0 + end + mc. Definition cost_if x y mc := match (isReg x, match y with | Some y' => isReg y' | None => true end) with - | (false, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) - | (false, true) | ( true, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) - | ( true, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) - end. + | (false, false) => mkMetricLog 4 0 6 1 + | (false, true) | ( true, false) => mkMetricLog 3 0 4 1 + | ( true, true) => mkMetricLog 2 0 2 1 + end + mc. Definition cost_loop_true x y mc := match (isReg x, match y with | Some y' => isReg y' | None => true end) with - | (false, false) => (addMetricInstructions 4 (addMetricLoads 6 (addMetricJumps 1 mc))) - | (false, true) | ( true, false) => (addMetricInstructions 3 (addMetricLoads 4 (addMetricJumps 1 mc))) - | ( true, true) => (addMetricInstructions 2 (addMetricLoads 2 (addMetricJumps 1 mc))) - end. + | (false, false) => mkMetricLog 4 0 6 1 + | (false, true) | ( true, false) => mkMetricLog 3 0 4 1 + | ( true, true) => mkMetricLog 2 0 2 1 + end + mc. Definition cost_loop_false x y mc := match (isReg x, match y with | Some y' => isReg y' | None => true end) with - | (false, false) => (addMetricInstructions 3 (addMetricLoads 5 (addMetricJumps 1 mc))) - | (false, true) | ( true, false) => (addMetricInstructions 2 (addMetricLoads 3 (addMetricJumps 1 mc))) - | ( true, true) => (addMetricInstructions 1 (addMetricLoads 1 (addMetricJumps 1 mc))) - end. + | (false, false) => mkMetricLog 3 0 5 1 + | (false, true) | ( true, false) => mkMetricLog 2 0 3 1 + | ( true, true) => mkMetricLog 1 0 1 1 + end + mc. End FlatImpExec. @@ -132,44 +132,4 @@ Ltac cost_destr := Ltac cost_solve := cost_unfold; cost_destr; try solve_MetricLog. Ltac cost_solve_piecewise := cost_unfold; cost_destr; try solve_MetricLog_piecewise. - -Definition metrics_additive f := forall mc, f mc - mc = f EmptyMetricLog. -#[global] Transparent metrics_additive. - -Ltac t := unfold metrics_additive; intros; cost_solve_piecewise. -Lemma cost_additive_interact : forall phase, metrics_additive (cost_interact phase). Proof. t. Qed. -Lemma cost_additive_call_internal : forall phase, metrics_additive (cost_call_internal phase). Proof. t. Qed. -Lemma cost_additive_call_external : forall phase, metrics_additive (cost_call_external phase). Proof. t. Qed. -Lemma cost_additive_load : forall varname isReg x a, metrics_additive (@cost_load varname isReg x a). Proof. t. Qed. -Lemma cost_additive_store : forall varname isReg x a, metrics_additive (@cost_store varname isReg x a). Proof. t. Qed. -Lemma cost_additive_inlinetable : forall varname isReg x a, metrics_additive (@cost_inlinetable varname isReg x a). Proof. t. Qed. -Lemma cost_additive_stackalloc : forall varname isReg x, metrics_additive (@cost_stackalloc varname isReg x). Proof. t. Qed. -Lemma cost_additive_lit : forall varname isReg x, metrics_additive (@cost_lit varname isReg x). Proof. t. Qed. -Lemma cost_additive_op : forall varname isReg x y z, metrics_additive (@cost_op varname isReg x y z). Proof. t. Qed. -Lemma cost_additive_set : forall varname isReg x y, metrics_additive (@cost_set varname isReg x y). Proof. t. Qed. -Lemma cost_additive_if : forall varname isReg x y, metrics_additive (@cost_if varname isReg x y). Proof. t. Qed. -Lemma cost_additive_loop_true : forall varname isReg x y, metrics_additive (@cost_loop_true varname isReg x y). Proof. t. Qed. -Lemma cost_additive_loop_false : forall varname isReg x y, metrics_additive (@cost_loop_false varname isReg x y). Proof. t. Qed. - -(* COQBUG: - these cannot have "in *", because this causes a crash if the environment contains a Semantics.ExtSpec. - Anomaly "Unable to handle arbitrary u+k <= v constraints." - Please report at http://coq.inria.fr/bugs/. - to use these in hypotheses, revert them first *) -Ltac cost_additive := repeat ( - rewrite cost_additive_interact || - rewrite cost_additive_call_internal || - rewrite cost_additive_call_external || - rewrite cost_additive_load || - rewrite cost_additive_store || - rewrite cost_additive_inlinetable || - rewrite cost_additive_stackalloc || - rewrite cost_additive_lit || - rewrite cost_additive_op || - rewrite cost_additive_set || - rewrite cost_additive_if || - rewrite cost_additive_loop_true || - rewrite cost_additive_loop_false - ). - -Ltac cost_hammer := cost_additive; try solve [eauto 3 with metric_arith | cost_solve]. +Ltac cost_hammer := try solve [eauto 3 with metric_arith | cost_solve]. diff --git a/bedrock2/src/bedrock2/MetricLogging.v b/bedrock2/src/bedrock2/MetricLogging.v index 2468f050a..cc7147055 100644 --- a/bedrock2/src/bedrock2/MetricLogging.v +++ b/bedrock2/src/bedrock2/MetricLogging.v @@ -151,10 +151,16 @@ Module MetricArith. Lemma sub_0_r : forall mc, (mc - EmptyMetricLog)%metricsH = mc. Proof. destruct mc. unfold EmptyMetricLog. solve_MetricLog_piecewise. Qed. + Lemma add_comm : forall n m, (n + m = m + n)%metricsH. + Proof. intros. unfold_MetricLog. f_equal; apply Z.add_comm. Qed. + + Lemma add_assoc : forall n m p, (n + (m + p) = n + m + p)%metricsH. + Proof. intros. unfold_MetricLog. f_equal; apply Z.add_assoc. Qed. + End MetricArith. Create HintDb metric_arith. -#[export] Hint Resolve MetricArith.le_trans MetricArith.le_refl MetricArith.add_0_r MetricArith.sub_0_r : metric_arith. +#[export] Hint Resolve MetricArith.le_trans MetricArith.le_refl MetricArith.add_0_r MetricArith.sub_0_r MetricArith.add_comm MetricArith.add_assoc : metric_arith. #[export] Hint Resolve <- MetricArith.le_sub_mono : metric_arith. #[export] Hint Resolve -> MetricArith.le_sub_mono : metric_arith. diff --git a/compiler/src/compiler/FlatImp.v b/compiler/src/compiler/FlatImp.v index d5c558b76..829eac6c0 100644 --- a/compiler/src/compiler/FlatImp.v +++ b/compiler/src/compiler/FlatImp.v @@ -305,21 +305,21 @@ Module exec. Definition cost_SIf bcond mc := match bcond with - | CondBinary _ x y => cost_if isReg x (Some y) mc - | CondNez x => cost_if isReg x None mc - end. + | CondBinary _ x y => cost_if isReg x (Some y) + | CondNez x => cost_if isReg x None + end mc. Definition cost_SLoop_true bcond mc := match bcond with - | CondBinary _ x y => cost_loop_true isReg x (Some y) mc - | CondNez x => cost_loop_true isReg x None mc - end. + | CondBinary _ x y => cost_loop_true isReg x (Some y) + | CondNez x => cost_loop_true isReg x None + end mc. Definition cost_SLoop_false bcond mc := match bcond with - | CondBinary _ x y => cost_loop_false isReg x (Some y) mc - | CondNez x => cost_loop_false isReg x None mc - end. + | CondBinary _ x y => cost_loop_false isReg x (Some y) + | CondNez x => cost_loop_false isReg x None + end mc. (* alternative semantics which allow non-determinism *) Inductive exec: @@ -708,18 +708,4 @@ Ltac scost_destr := Ltac scost_solve := scost_unfold; scost_destr; try solve_MetricLog. Ltac scost_solve_piecewise := scost_unfold; scost_destr; try solve_MetricLog_piecewise. - -Ltac t := unfold metrics_additive; intros; scost_solve_piecewise. -Lemma scost_additive_op : forall varname isReg x y z, metrics_additive (@exec.cost_SOp varname isReg x y z). Proof. t. Qed. -Lemma scost_additive_if : forall varname isReg x, metrics_additive (@exec.cost_SIf varname isReg x). Proof. t. Qed. -Lemma scost_additive_loop_true : forall varname isReg x, metrics_additive (@exec.cost_SLoop_true varname isReg x). Proof. t. Qed. -Lemma scost_additive_loop_false : forall varname isReg x, metrics_additive (@exec.cost_SLoop_false varname isReg x). Proof. t. Qed. - -Ltac scost_additive := repeat ( - rewrite scost_additive_op || - rewrite scost_additive_if || - rewrite scost_additive_loop_true || - rewrite scost_additive_loop_false - ); cost_additive. - -Ltac scost_hammer := scost_additive; try solve [eauto 3 with metric_arith | scost_solve]. +Ltac scost_hammer := try solve [eauto 3 with metric_arith | scost_solve]. diff --git a/compiler/src/compiler/FlatToRiscvFunctions.v b/compiler/src/compiler/FlatToRiscvFunctions.v index 7900c1edc..6defbb0ae 100644 --- a/compiler/src/compiler/FlatToRiscvFunctions.v +++ b/compiler/src/compiler/FlatToRiscvFunctions.v @@ -1634,7 +1634,7 @@ Section Proofs. split; eauto 8 with map_hints. split; eauto 8 with map_hints. (* cost_SCall constraint: cost_SCall_L + (1,1,1,0) <= cost_SCall_internal + cost_SCall_external *) - unfold cost_SCall_L in *. + unfold cost_SCall_L, cost_call_internal, cost_call_external in *. MetricsToRiscv.solve_MetricLog. - idtac "Case compile_stmt_correct/SLoad". diff --git a/compiler/src/compiler/MMIO.v b/compiler/src/compiler/MMIO.v index 52bfedd47..df348f136 100644 --- a/compiler/src/compiler/MMIO.v +++ b/compiler/src/compiler/MMIO.v @@ -369,7 +369,7 @@ Section MMIO1. split; eauto. split; [unfold map.only_differ; eauto|]. split. { - unfold id. MetricsToRiscv.solve_MetricLog. + unfold id, MetricCosts.cost_interact. MetricsToRiscv.solve_MetricLog. } split; eauto. split; eauto. @@ -522,7 +522,7 @@ Section MMIO1. destruct_one_match; auto. } split. { - unfold id. MetricsToRiscv.solve_MetricLog. + unfold id, MetricCosts.cost_interact. MetricsToRiscv.solve_MetricLog. } split. { eapply map.put_extends. eassumption. diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index 420e058c8..bc2ceb4ff 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -293,7 +293,7 @@ Section WithWordAndMem. + eapply @freshNameGenState_disjoint_fbody. - simpl. intros. fwd. eexists. split. + eauto using map.getmany_of_list_extends. - + eexists. split; [|eassumption]. solve_MetricLog. + + eexists. split; [|eassumption]. cost_unfold; solve_MetricLog. Qed. Lemma useimmediate_functions_NoDup: forall funs funs', @@ -334,7 +334,7 @@ Section WithWordAndMem. - eapply useImmediate_correct_aux; eauto. - simpl. destruct 1 as (?&?&?&?&?). repeat (eexists; split; try eassumption). - solve_MetricLog. + cost_unfold; solve_MetricLog. Qed. @@ -432,7 +432,7 @@ Section WithWordAndMem. rewrite H1 in P'. inversion P'. exact Cp. - simpl. intros. fwd. eexists. split. + eauto using states_compat_getmany. - + eexists. split; [|eassumption]. solve_MetricLog. + + eexists. split; [|eassumption]. cost_unfold; solve_MetricLog. Qed. Ltac debool := diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index b5cfe57f5..18c8a9551 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -1123,19 +1123,6 @@ Section CheckerCorrect. unfold assert_ins, assert. intros. fwd. assumption. Qed. - (* TODO XXX this should be obsolete *) - Ltac unfold_metrics := - unfold - MetricLogging.withInstructions, MetricLogging.withLoads, - MetricLogging.withStores, MetricLogging.withJumps, - MetricLogging.addMetricInstructions, MetricLogging.addMetricLoads, - MetricLogging.addMetricStores, MetricLogging.addMetricJumps, - MetricLogging.subMetricInstructions, MetricLogging.subMetricLoads, - MetricLogging.subMetricStores, MetricLogging.subMetricJumps, - MetricLogging.metricsOp, MetricLogging.metricSub, MetricLogging.metricsSub, - MetricLogging.metricLeq, MetricLogging.metricsLeq - in *. - Opaque isRegStr. Opaque isRegZ. @@ -1150,8 +1137,6 @@ Section CheckerCorrect. | H: match ?expr with _ => _ end = Success _ |- _ => destr expr; discriminate end. - (* TODO XXX these lemms are unnecessarily slow i think *) - Lemma check_regs_cost_SIf: forall (cond : bcond) (bThen bElse : stmt) (corresp : list (string * Z)) (cond0 : bcond') (s'1 s'2 : stmt') (a0 a1 : list (string * Z)) @@ -1168,13 +1153,12 @@ Section CheckerCorrect. Proof. intros. repeat (unfold check_bcond, assert_in, assignment in *; fwd). - intros; unfold check_regs in *; cbn in *; unfold exec.cost_SIf in *; - destr cond; destr cond0; destr (isRegStr x); destr (isRegZ x0); try (destr (isRegStr y)); try (destr (isRegZ y0)); - try discriminate; - unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; fwd; - destr mcL; destr mc'; destr mc; destr mcH'; try blia. - all: try discr_match_success. - all: cbn in *; fwd; try blia. + destr cond; destr cond0; + repeat match goal with + | H: match ?expr with _ => _ end = Success _ |- _ => destr expr; try discriminate + end; + unfold check_regs, exec.cost_SIf in *; + cost_solve. Qed. Lemma check_regs_cost_SLoop_false: @@ -1191,13 +1175,12 @@ Section CheckerCorrect. Proof. intros. repeat (unfold check_bcond, assert_in, assignment in *; fwd). - intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLoop_false in *; - destr cond; destr cond0; destr (isRegStr x); destr (isRegZ x0); try (destr (isRegStr y)); try (destr (isRegZ y0)); - try discriminate; - unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; fwd; - destr mcL; destr mc'; destr mc; destr mcH'; try blia. - all: try discr_match_success. - all: cbn in *; fwd; try blia. + destr cond; destr cond0; + repeat match goal with + | H: match ?expr with _ => _ end = Success _ |- _ => destr expr; try discriminate + end; + unfold check_regs, exec.cost_SLoop_false in *; + cost_solve. Qed. Hint Constructors exec.exec : checker_hints. @@ -1251,8 +1234,7 @@ Section CheckerCorrect. intros. edestruct H3 as (l' & P & F). 1: eassumption. eapply putmany_of_list_zip_states_compat in P. 2-3: eassumption. destruct P as (lL' & P & SC). eexists. split. 1: eassumption. intros. eauto. - repeat eexists; eauto. - all: repeat (unfold_metrics; cbn; try blia). + repeat eexists; eauto; cost_solve. - (* Case exec.call *) rename binds0 into binds', args0 into args'. unfold check_funcs in H. @@ -1278,7 +1260,7 @@ Section CheckerCorrect. * eapply states_compat_getmany; eassumption. * exact L4. * repeat eexists. 6: eassumption. 1: exact SC. - all: repeat (unfold_metrics; cbn in *; try blia). + all: cost_solve. - (* Case exec.load *) a. b. - (* Case exec.store *) @@ -1298,7 +1280,7 @@ Section CheckerCorrect. a. b. - (* Case exec.op *) a. - unfold assert_in_op, assert_in, assignment, check_regs in *; scost_hammer. (* TODO XXX could be faster *) + unfold assert_in_op, assert_in, assignment, check_regs in *; scost_hammer. (* could be faster *) - (* Case exec.set *) a. b. - (* Case exec.if_true *) @@ -1361,10 +1343,7 @@ Section CheckerCorrect. intros; unfold check_regs in *; cbn in *; unfold exec.cost_SLoop_true in *; try discr_match_success; destr cond; destr cond0; destr (isRegStr x); destr (isRegZ x0); try (destr (isRegStr y)); try (destr (isRegZ y0)); try discriminate; try discr_match_success. - all: unfold_metrics; cbn in *; repeat split; cbn in *; unfold_metrics; cbn in *; fwd. - all: destr mc''; destr mcH'; destr mcHmid; destr mcLmid; destr mc'; destr mcH'0; destr mc; destr mcL; - try discr_match_success. - all: cbn in *; try blia. + all: cost_solve. - (* Case exec.seq *) rename H2 into IH2, IHexec into IH1. eapply exec.seq. diff --git a/compiler/src/compiler/Spilling.v b/compiler/src/compiler/Spilling.v index b4458f648..01acb0903 100644 --- a/compiler/src/compiler/Spilling.v +++ b/compiler/src/compiler/Spilling.v @@ -461,7 +461,7 @@ Section Spilling. map.get l1 r = Some v -> (related maxvar frame fpval t1 m1 l1 t2 m2 (map.put l2 (iarg_reg i r) v) -> post t2 m2 (map.put l2 (iarg_reg i r) v) - (if isRegZ r then mc2 else (addMetricInstructions 1 (addMetricLoads 2 mc2)))) -> + (if isRegZ r then mc2 else (mkMetricLog 1 0 2 0 + mc2)%metricsH)) -> execpost e2 (load_iarg_reg i r) t2 m2 l2 mc2 post. Proof. intros. @@ -564,7 +564,7 @@ Section Spilling. execpost e2 (load_iarg_reg i r) t2 m2 l2 mc2 (fun t2' m2' l2' mc2' => t2' = t2 /\ m2' = m2 /\ l2' = map.put l2 (iarg_reg i r) v /\ related maxvar frame fpval t1 m1 l1 t2' m2' l2' /\ - (mc2' <= (if isRegZ r then mc2 else (addMetricInstructions 1 (addMetricLoads 2 mc2))))%metricsH). + (mc2' <= (if isRegZ r then mc2 else mkMetricLog 1 0 2 0 + mc2))%metricsH). Proof. intros. unfold load_iarg_reg, stack_loc, iarg_reg, related in *. fwd. @@ -612,7 +612,7 @@ Section Spilling. post t1 m1 (map.put l1 x v) mc1' -> related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> fp < x <= maxvar /\ (x < a0 \/ a7 < x) -> - (mc2' - mc2 <= mc1' - (if isRegZ x then mc1 else (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc1)))))%metricsH -> + (mc2' - mc2 <= mc1' - (if isRegZ x then mc1 else mkMetricLog 1 1 1 0 + mc1))%metricsH -> execpost e (save_ires_reg x) t2 m2 (map.put l2 (ires_reg x) v) mc2' (fun t2' m2' l2' mc2'' => exists t1' m1' l1' mc1'', related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ @@ -717,7 +717,7 @@ Section Spilling. fp < x <= maxvar /\ (x < a0 \/ a7 < x) -> (forall t2' m2' l2', related maxvar frame fpval t1 m1 (map.put l1 x v) t2' m2' l2' -> - post t2' m2' l2' (if isRegZ x then mc2 else (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc2))))) -> + post t2' m2' l2' (if isRegZ x then mc2 else mkMetricLog 1 1 1 0 + mc2)%metricsH) -> execpost e (save_ires_reg x) t2 m2 (map.put l2 (ires_reg x) v) mc2 post. Proof. intros. @@ -860,36 +860,17 @@ Section Spilling. Fixpoint cost_set_vars_to_reg_range (args: list Z) (start : Z) (mc : MetricLog) : MetricLog := match args with | [] => mc - | x :: xs => if isRegZ x then - addMetricInstructions 1 - (addMetricLoads 1 - (cost_set_vars_to_reg_range xs (start + 1) mc)) - else (addMetricInstructions 1 - (addMetricLoads 1 - (addMetricStores 1 - (cost_set_vars_to_reg_range xs (start + 1) mc)))) + | x :: xs => (if isRegZ x then mkMetricLog 1 0 1 0 else mkMetricLog 1 1 1 0) + + cost_set_vars_to_reg_range xs (start + 1) mc end. Lemma cost_set_vars_to_reg_range_commutes: - forall args start mc, - cost_set_vars_to_reg_range args start - (addMetricInstructions 1 (addMetricLoads 1 (addMetricStores 1 mc))) = - addMetricInstructions 1 - (addMetricLoads 1 (addMetricStores 1 (cost_set_vars_to_reg_range args start mc))). + forall args start n m, + (cost_set_vars_to_reg_range args start (n + m) = n + cost_set_vars_to_reg_range args start m)%metricsH. Proof. - induction args; intros; cbn; try trivial. - destr (isRegZ a); cbn; rewrite IHargs; trivial. - Qed. - - Lemma cost_set_vars_to_reg_range_commutes': - forall args start mc, - cost_set_vars_to_reg_range args start - (addMetricInstructions 1 (addMetricLoads 1 mc)) = - addMetricInstructions 1 - (addMetricLoads 1 (cost_set_vars_to_reg_range args start mc)). - Proof. - induction args; intros; cbn; try trivial. - destr (isRegZ a); cbn; rewrite IHargs; trivial. + induction args; trivial. + intros; cbn; destr (isRegZ a); cbn; rewrite IHargs; + do 2 rewrite MetricArith.add_assoc; rewrite (MetricArith.add_comm n); reflexivity. Qed. Lemma set_vars_to_reg_range_correct: @@ -943,7 +924,10 @@ Section Spilling. { blia. } * intros. apply H6; auto. cbn in *. destr (isRegZ start); try blia; destr (isRegZ a); try blia. - rewrite H0. rewrite cost_set_vars_to_reg_range_commutes. trivial. + rewrite H0. unfold cost_store, isRegZ. cbn. + rewrite cost_set_vars_to_reg_range_commutes. + rewrite (proj2 (Z.leb_le start 31)) by assumption. + reflexivity. + eapply exec.set. { eassumption. } eapply IHargs; try eassumption; try blia. 2: { @@ -970,15 +954,18 @@ Section Spilling. { assumption. } * intros. apply H6; auto. cbn in *. destr (isRegZ start); try blia; destr (isRegZ a); try blia. - rewrite H0. rewrite cost_set_vars_to_reg_range_commutes'. trivial. + rewrite H0. unfold cost_set, isRegZ. cbn. + rewrite cost_set_vars_to_reg_range_commutes. + rewrite (proj2 (Z.leb_le a 31)) by assumption. + rewrite (proj2 (Z.leb_le start 31)) by assumption. + reflexivity. Qed. Fixpoint cost_set_reg_range_to_vars (start : Z) (args: list Z) (mc : MetricLog) : MetricLog := match args with | [] => mc - | x :: xs => if isRegZ x then - addMetricInstructions 1 (addMetricLoads 1 (cost_set_reg_range_to_vars (start + 1) xs mc)) - else (addMetricInstructions 1 (addMetricLoads 2 (cost_set_reg_range_to_vars (start + 1) xs mc))) + | x :: xs => (if isRegZ x then mkMetricLog 1 0 1 0 else mkMetricLog 1 0 2 0) + + cost_set_reg_range_to_vars (start + 1) xs mc end. Lemma set_reg_range_to_vars_correct: @@ -1035,7 +1022,9 @@ Section Spilling. eapply map.getmany_of_list_put_diff. 2: eassumption. eauto using List.not_In_Z_seq with zarith. -- cbn. destr (isRegZ start); destr (isRegZ a); cbn in *; try blia. - rewrite H6; trivial. + rewrite H6; unfold cost_load, isRegZ; cbn. + rewrite (proj2 (Z.leb_le start 31)) by assumption. + reflexivity. + eapply exec.seq_cps. eapply IHargs; try eassumption; try blia. intros. @@ -1063,48 +1052,10 @@ Section Spilling. end. eapply put_arg_reg; try eassumption. blia. -- cbn. destr (isRegZ start); destr (isRegZ a); cbn in *; try blia. - rewrite H6; trivial. - Qed. - - (* XXX this section seems silly *) - - Lemma factor_out_set_reg_range_to_vars : forall args start mc, - cost_set_reg_range_to_vars start args mc = - (mc + cost_set_reg_range_to_vars start args EmptyMetricLog)%metricsH. - Proof. - induction args; eauto with metric_arith. - simpl; intros. - rewrite IHargs. - destruct (isRegZ a); solve_MetricLog_piecewise. - Qed. - - Lemma factor_out_set_vars_to_reg_range : forall args start mc, - cost_set_vars_to_reg_range args start mc = - (mc + cost_set_vars_to_reg_range args start EmptyMetricLog)%metricsH. - Proof. - induction args; eauto with metric_arith. - simpl; intros. - rewrite IHargs. - destruct (isRegZ a); solve_MetricLog_piecewise. - Qed. - - Lemma factor_out_cost_external : forall phase mc, - cost_call_external phase mc = - (mc + cost_call_external phase EmptyMetricLog)%metricsH. - Proof. destruct phase; eauto with metric_arith. Qed. - - Lemma factor_out_cost_internal : forall phase mc, - cost_call_internal phase mc = - (mc + cost_call_internal phase EmptyMetricLog)%metricsH. - Proof. destruct phase; eauto with metric_arith. Qed. - - Lemma factor_out_cost_stackalloc : forall x mc, - cost_stackalloc isRegZ x mc = - (mc + cost_stackalloc isRegZ x EmptyMetricLog)%metricsH. - Proof. - intros. - unfold cost_stackalloc, EmptyMetricLog. - destruct (isRegZ x); solve_MetricLog_piecewise. + rewrite H6; unfold cost_set, isRegZ; cbn. + rewrite (proj2 (Z.leb_le a 31)) by assumption. + rewrite (proj2 (Z.leb_le start 31)) by assumption. + reflexivity. Qed. Lemma cost_set_reg_range_to_vars_bound : forall args start mc len, @@ -1625,7 +1576,7 @@ Section Spilling. eapply map.shrink_disjoint_l; eassumption. } cbn in *. subst. add_bounds. - solve_MetricLog. + cost_solve. (* cost_SInteract constraint: prespill - postspill >= (...32...) i think? *) } (* related for set_vars_to_reg_range_correct: *) @@ -1845,19 +1796,9 @@ Section Spilling. split; try eassumption. subst. move Hmetrics at bottom. - rewrite factor_out_set_vars_to_reg_range. - rewrite factor_out_cost_external. - rewrite factor_out_cost_stackalloc. - rewrite factor_out_set_reg_range_to_vars. - rewrite (factor_out_cost_external PreSpill). - rewrite factor_out_set_vars_to_reg_range in Hmetrics. - rewrite factor_out_cost_internal in Hmetrics. - rewrite factor_out_set_reg_range_to_vars in Hmetrics. - rewrite (factor_out_cost_internal PreSpill) in Hmetrics. add_bounds. cost_solve. (* cost_SCall constraint: prespill - postspill >= (...66...) i think? *) - (* TODO XXX fairly slow *) } - (* exec.load *) diff --git a/compiler/src/compiler/memory_mapped_ext_calls_compiler.v b/compiler/src/compiler/memory_mapped_ext_calls_compiler.v index 5a421ec6a..e9cb05b57 100644 --- a/compiler/src/compiler/memory_mapped_ext_calls_compiler.v +++ b/compiler/src/compiler/memory_mapped_ext_calls_compiler.v @@ -422,7 +422,7 @@ Section MMIO. { eapply MapEauto.only_differ_union_l. eapply MapEauto.only_differ_put. cbn. left. reflexivity. } - { MetricsToRiscv.solve_MetricLog. } + { unfold MetricCosts.cost_interact in *; MetricsToRiscv.solve_MetricLog. } { eapply map.put_extends. eassumption. } { eapply map.forall_keys_put; assumption. } { rewrite map.get_put_diff; eauto. unfold RegisterNames.sp. @@ -537,7 +537,7 @@ Section MMIO. eapply map.split_empty_r. reflexivity. } { reflexivity. } { eapply MapEauto.only_differ_refl. } - { MetricsToRiscv.solve_MetricLog. } + { unfold MetricCosts.cost_interact in *; MetricsToRiscv.solve_MetricLog. } { eassumption. } { assumption. } { assumption. } From 63eeb7e0acea09d6d418668937618a3c9277ea46 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Wed, 10 Jul 2024 10:24:17 -0400 Subject: [PATCH 50/62] fix ipow again --- bedrock2/src/bedrock2Examples/metric_ipow.v | 36 ++++++++++----------- 1 file changed, 17 insertions(+), 19 deletions(-) diff --git a/bedrock2/src/bedrock2Examples/metric_ipow.v b/bedrock2/src/bedrock2Examples/metric_ipow.v index 639b98b8d..d8ebc57b7 100644 --- a/bedrock2/src/bedrock2Examples/metric_ipow.v +++ b/bedrock2/src/bedrock2Examples/metric_ipow.v @@ -100,6 +100,17 @@ Proof. blia. Qed. +Lemma metriclit : forall a b c d a' b' c' d', + metricsAdd (mkMetricLog a b c d) (mkMetricLog a' b' c' d') = + mkMetricLog (a+a') (b+b') (c+c') (d+d'). +Proof. reflexivity. Qed. + +Ltac s := unfold initCost, iterCost, endCost in *; + cost_unfold; + repeat rewrite MetricArith.add_assoc in *; cbn in *; + repeat rewrite metriclit in *; cbn in *; + solve_MetricLog. + Lemma ipow_ok : program_logic_goal_for_function! ipow. Proof. repeat straightline. @@ -152,18 +163,12 @@ Proof. rewrite Z.mul_mod_idemp_r by discriminate. f_equal; ring. } { (* metrics correct *) - cbn in H4. - rewrite <- (MetricLog_eq mc0) in H4. - applyAddMetrics H4. - cbn in H4. rewrite msb_shift in H4 by blia. - unfold iterCost in *. rewrite MetricArith.mul_sub_distr_r in H4. rewrite <- MetricArith.add_sub_swap in H4. rewrite <- MetricArith.le_add_le_sub_r in H4. - eapply MetricArith.le_trans. - 2: exact H4. - solve_MetricLog. + eapply MetricArith.le_trans with (2 := H4). + s. } } { @@ -184,23 +189,16 @@ Proof. rewrite Z.mul_mod_idemp_r by discriminate. f_equal; ring. } { (* metrics correct *) - cbn in H4. - rewrite <- (MetricLog_eq mc0) in H4. - applyAddMetrics H4. - cbn in H4. rewrite msb_shift in H4 by blia. - unfold iterCost in *. - solve_MetricLog. + s. } } } { (* postcondition *) rewrite H, Z.pow_0_r, Z.mul_1_r, word.wrap_unsigned. split; [reflexivity|]. - unfold msb, iterCost, endCost. - subst brmc. - cbn. - solve_MetricLog. + unfold msb; subst brmc. + s. } } @@ -208,5 +206,5 @@ Proof. repeat (split || letexists || t || trivial). { setoid_rewrite H1; setoid_rewrite Z.mul_1_l; trivial. } - all: cbn in H2; unfold initCost, iterCost, endCost in *; solve_MetricLog. + all: s. Qed. From ce31e90eb3e7de26b2129a6d140674e67a21b25c Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Wed, 10 Jul 2024 10:25:22 -0400 Subject: [PATCH 51/62] loop/wp properties lemmas needed --- bedrock2/src/bedrock2/MetricLoops.v | 110 +++++------------- .../MetricWeakestPreconditionProperties.v | 13 +-- 2 files changed, 37 insertions(+), 86 deletions(-) diff --git a/bedrock2/src/bedrock2/MetricLoops.v b/bedrock2/src/bedrock2/MetricLoops.v index 28417652e..52210674e 100644 --- a/bedrock2/src/bedrock2/MetricLoops.v +++ b/bedrock2/src/bedrock2/MetricLoops.v @@ -359,88 +359,49 @@ Section Loops. constructor. intros [] []. Qed. - (*METRICSTODO - Lemma atleastonce_localsmap {e c t} {m : mem} {l mc} {post : _->_->_->_-> Prop} {measure : Type} (invariant:_->_->_->_->_->Prop) lt (Hwf : well_founded lt) - (Henter : exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ (word.unsigned br = 0%Z -> post t m l (addMetricInstructions 1 - (addMetricLoads 1 - (addMetricJumps 1 mc')))) - /\ (word.unsigned br <> 0 -> exists v', invariant v' t m l (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc')))) + (Henter : exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) + /\ (word.unsigned br = 0%Z -> post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc)) + /\ (word.unsigned br <> 0 -> exists v', invariant v' t m l (cost_loop_true isRegStr UNK (Some UNK) brmc)) ) (v0 : measure) (Hpre : invariant v0 t m l mc) (Hbody : forall v t m l mc, invariant v t m l mc -> cmd call c t m l mc (fun t m l mc => - exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ - (word.unsigned br <> 0 -> exists v', invariant v' t m l (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc'))) /\ lt v' v) /\ - (word.unsigned br = 0 -> post t m l (addMetricInstructions 1 - (addMetricLoads 1 - (addMetricJumps 1 mc')))))) + exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) + /\ (word.unsigned br <> 0 -> exists v', invariant v' t m l (cost_loop_true isRegStr UNK (Some UNK) brmc) /\ lt v' v) + /\ (word.unsigned br = 0 -> post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc)))) : cmd call (cmd.while e c) t m l mc post. Proof. -<<<<<<< HEAD eapply wp_while. - eexists (option measure), (with_bottom lt), (fun ov t m l => - exists br, expr m l e (eq br) /\ - ((word.unsigned br <> 0 -> exists v, ov = Some v /\ invariant v t m l) /\ - (word.unsigned br = 0 -> ov = None /\ post t m l))). -======= eexists (option measure), (with_bottom lt), (fun ov t m l mc => - exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ - ((word.unsigned br <> 0 -> exists v, ov = Some v /\ invariant v t m l (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc')))) /\ - (word.unsigned br = 0 -> ov = None /\ post t m l (addMetricInstructions 1 - (addMetricLoads 1 - (addMetricJumps 1 mc')))))). ->>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) + exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) + /\ ((word.unsigned br <> 0 -> exists v, ov = Some v /\ invariant v t m l (cost_loop_true isRegStr UNK (Some UNK) brmc)) + /\ (word.unsigned br = 0 -> ov = None /\ post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc)))). split; auto using well_founded_with_bottom; []. split. - { destruct Henter as [br [mc' [He [Henter0 Henterm]]]]. + { destruct Henter as [br [brmc [He [Henter0 Henterm]]]]. destruct (BinInt.Z.eq_dec (word.unsigned br) 0). - { exists None, br, mc'; split; trivial. + { exists None, br, brmc; split; trivial. split; intros; try contradiction; split; eauto. } - { - assert (Hnonzero := n). + { assert (Hnonzero := n). apply Henterm in Hnonzero. destruct Hnonzero as [v Hinv]. - exists (Some v), br, mc'. - + exists (Some v), br, brmc. split; trivial; []; split; try contradiction. -<<<<<<< HEAD - exists v0; split; trivial. } } - intros vi ti mi li (br&Ebr&Hcontinue&Hexit). - eexists; split; [eassumption|]; split. + exists v; split; trivial. } } + intros vi ti mi li mci (br&brmc&Ebr&Hcontinue&Hexit). + eexists; eexists; split; [eassumption|]; split. + { admit. } + (* { intros Hc; destruct (Hcontinue Hc) as (v&?&Hinv); subst. - eapply Proper_cmd; [ |eapply Hbody; eassumption]. -======= - - exists v; split; trivial. - - - } } - intros vi ti mi li mci (br&mcnew&Ebr&Hcontinue&Hexit). - eexists; eexists; split. - - eassumption. - - split. - - - (*eexists; eexists; split; [eassumption|]; split.*) - - { - intros Hc; destruct (Hcontinue Hc) as (v&?&Hinv); subst. eapply Proper_cmd. - 1: { eapply Proper_call. } - 2: { (* eapply Hbody. eassumption. } - - [ | | eapply Hbody; eassumption]. [eapply Proper_call|]. ->>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) - intros t' m' l' (br'&Ebr'&Hinv'&Hpost'). + 2: { + eapply Hbody. + } + eapply Proper_cmd; [ |eapply Hbody; eassumption]. + intros t' m' l' mc' (br'&brmc'&Ebr'&Hinv'&Hpost'). destruct (BinInt.Z.eq_dec (word.unsigned br') 0). { exists None; split; try constructor. exists br'; split; trivial; []. @@ -451,10 +412,11 @@ Section Loops. exists br'; split; trivial. split; intros; try contradiction. eexists; split; eauto. } } + *) eapply Hexit. - Qed. - *) Admitted. + (*Qed.*) + Admitted. Lemma atleastonce {e c t l mc} {m : mem} @@ -464,12 +426,8 @@ Section Loops. {measure : Type} (invariant:_->_->_-> ufunc word (length variables) (MetricLog -> Prop)) lt (Hwf : well_founded lt) {post : _->_->_->_->Prop} - (Henter : exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ (word.unsigned br = 0%Z -> post t m l (addMetricInstructions 1 - (addMetricLoads 1 - (addMetricJumps 1 mc')))) /\ - (word.unsigned br <> 0 -> exists v', tuple.apply (invariant v' t m) localstuple (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc')))) + (Henter : exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ (word.unsigned br = 0%Z -> post t m l (cost_loop_false isRegStr UNK (Some UNK) mc')) /\ + (word.unsigned br <> 0 -> exists v', tuple.apply (invariant v' t m) localstuple (cost_loop_true isRegStr UNK (Some UNK) mc')) ) (v0 : measure) (Hpre : tuple.apply (invariant v0 t m) localstuple mc) @@ -477,12 +435,8 @@ Section Loops. tuple.apply (invariant v t m) localstuple mc -> cmd call c t m (reconstruct variables localstuple) mc (fun t m l mc => exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ - (word.unsigned br <> 0 -> Markers.unique (Markers.left (tuple.existss (fun localstuple => enforce variables localstuple l /\ Markers.right (Markers.unique (exists v', tuple.apply (invariant v' t m) localstuple (addMetricInstructions 2 - (addMetricLoads 2 - (addMetricJumps 1 mc'))) /\ lt v' v)))))) /\ - (word.unsigned br = 0 -> post t m l (addMetricInstructions 1 - (addMetricLoads 1 - (addMetricJumps 1 mc'))))))) + (word.unsigned br <> 0 -> Markers.unique (Markers.left (tuple.existss (fun localstuple => enforce variables localstuple l /\ Markers.right (Markers.unique (exists v', tuple.apply (invariant v' t m) localstuple (cost_loop_true isRegStr UNK (Some UNK) mc') /\ lt v' v)))))) /\ + (word.unsigned br = 0 -> post t m l (cost_loop_false isRegStr UNK (Some UNK) mc'))))) : cmd call (cmd.while e c) t m l mc post. Proof. eapply (atleastonce_localsmap (fun v t m l mc => exists localstuple, Logic.and (enforce variables localstuple l) (tuple.apply (invariant v t m) localstuple mc))); eauto. @@ -498,7 +452,7 @@ Section Loops. eapply hlist.foralls_forall in Hbody. specialize (Hbody Y). rewrite <-(reconstruct_enforce _ _ _ X) in Hbody. - eapply Proper_cmd; [eapply Proper_call| |eapply Hbody]. + eapply Proper_cmd; [|eapply Hbody]. intros t' m' l' mc' (?&?&?&HH). eexists; eexists; split; eauto. destruct HH as [HH1 HH2]. split; intros; eauto. @@ -508,8 +462,6 @@ Section Loops. } Qed. - METRICSTODO*) - (* Lemma tailrec_earlyout_localsmap {e c t} {m : mem} {l mc} {post : _-> _->_->_-> Prop} diff --git a/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v b/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v index 53e89e338..b0e16f67e 100644 --- a/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v +++ b/bedrock2/src/bedrock2/MetricWeakestPreconditionProperties.v @@ -345,23 +345,22 @@ width word word_ok. Import bedrock2.Syntax bedrock2.MetricSemantics bedrock2.MetricWeakestPrecondition. - (* Lemma interact_nomem call action binds arges t m l post - args (Hargs : dexprs m l arges args) + mc args mc' (Hargs : dexprs m l arges mc (args, mc')) (Hext : ext_spec t map.empty binds args (fun mReceive (rets : list word) => mReceive = map.empty /\ exists l0 : locals, map.putmany_of_list_zip action rets l = Some l0 /\ - post (cons (map.empty, binds, args, (map.empty, rets)) t) m l0)) - : WeakestPrecondition.cmd call (cmd.interact action binds arges) t m l post. + post (cons (map.empty, binds, args, (map.empty, rets)) t) m l0 (MetricCosts.cost_interact MetricCosts.PreSpill mc'))) + : MetricWeakestPrecondition.cmd call (cmd.interact action binds arges) t m l mc post. Proof using word_ok mem_ok ext_spec_ok. - exists args; split; [exact Hargs|]. + exists args, mc'; split; [exact Hargs|]. exists m. exists map.empty. split; [eapply Properties.map.split_empty_r; exact eq_refl|]. - eapply ext_spec.weaken; [|eapply Hext]; intros ? ? [? [? []]]. subst a; subst. + eapply Semantics.ext_spec.weaken; [|eapply Hext]; intros ? ? [? [? []]]. subst a; subst. eexists; split; [eassumption|]. intros. eapply Properties.map.split_empty_r in H. subst. assumption. - Qed.*) + Qed. (* Lemma intersect_expr: forall m l e mc (post1 post2: word * MetricLog -> Prop), From 496c5b7980e79869ef14a59b4745c15eab883b90 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Wed, 10 Jul 2024 10:38:41 -0400 Subject: [PATCH 52/62] very messy SPI wip --- bedrock2/src/bedrock2Examples/metric_SPI.v | 466 +++++++++++++++++++++ 1 file changed, 466 insertions(+) create mode 100644 bedrock2/src/bedrock2Examples/metric_SPI.v diff --git a/bedrock2/src/bedrock2Examples/metric_SPI.v b/bedrock2/src/bedrock2Examples/metric_SPI.v new file mode 100644 index 000000000..9acfc1e2e --- /dev/null +++ b/bedrock2/src/bedrock2Examples/metric_SPI.v @@ -0,0 +1,466 @@ +Require Import bedrock2.Syntax bedrock2.NotationsCustomEntry Coq.Strings.String. +Require Import coqutil.Z.div_mod_to_equations. +Require Import coqutil.Z.Lia. +Require Import coqutil.Word.Interface. +Require Import coqutil.Byte. + +Import BinInt String List.ListNotations ZArith. +Local Open Scope Z_scope. Local Open Scope string_scope. Local Open Scope list_scope. + +Require bedrock2Examples.lightbulb_spec. +Local Notation patience := lightbulb_spec.patience. + +Definition spi_write := func! (b) ~> busy { + busy = $-1; + i = $patience; while i { i = i - $1; + io! busy = MMIOREAD($0x10024048); + if !(busy >> $31) { i = i^i } + }; + if !(busy >> $31) { + output! MMIOWRITE($0x10024048, b); + busy = (busy ^ busy) + } + }. + +Definition spi_read := func! () ~> (b, busy) { + busy = $-1; + b = $0x5a; + i = $patience; while i { i = i - $1; + io! busy = MMIOREAD($0x1002404c); + if !(busy >> $31) { + b = busy & $0xff; + i = i^i; + busy = i + } + } + }. + +Definition spi_xchg := func! (b) ~> (b, busy) { + unpack! busy = spi_write(b); + require !busy; + unpack! b, busy = spi_read() + }. + +Require Import bedrock2.MetricProgramLogic. +Require Import bedrock2.MetricWeakestPreconditionProperties. +Require Import bedrock2.FE310CSemantics bedrock2.MetricSemantics. +Require Import Coq.Lists.List. Import ListNotations. +Require Import bedrock2.TracePredicate. Import TracePredicateNotations. +Require Import bedrock2.ZnWords. + +Import coqutil.Map.Interface. +Import ReversedListNotations. + +Section WithParameters. + Context {word: word.word 32} {mem: map.map word Byte.byte}. + Context {word_ok: word.ok word} {mem_ok: map.ok mem}. + + Definition mmio_event_abstraction_relation + (h : lightbulb_spec.OP word) + (l : mem * string * list word * (mem * list word)) := + Logic.or + (exists a v, h = ("st", a, v) /\ l = (map.empty, "MMIOWRITE", [a; v], (map.empty, []))) + (exists a v, h = ("ld", a, v) /\ l = (map.empty, "MMIOREAD", [a], (map.empty, [v]))). + Definition mmio_trace_abstraction_relation := + List.Forall2 mmio_event_abstraction_relation. + Definition only_mmio_satisfying P t := + exists mmios, mmio_trace_abstraction_relation mmios t /\ P mmios. + + Global Instance spec_of_spi_write : spec_of "spi_write" := fun functions => forall t m b mc, + word.unsigned b < 2 ^ 8 -> + MetricWeakestPrecondition.call functions "spi_write" t m [b] mc (fun T M RETS MC => + M = m /\ exists iol, T = t ;++ iol /\ exists ioh, mmio_trace_abstraction_relation ioh iol /\ exists err, RETS = [err] /\ Logic.or + (((word.unsigned err <> 0) /\ lightbulb_spec.spi_write_full _ ^* ioh /\ Z.of_nat (length ioh) = patience)) + (word.unsigned err = 0 /\ lightbulb_spec.spi_write word (byte.of_Z (word.unsigned b)) ioh)). + + Global Instance spec_of_spi_read : spec_of "spi_read" := fun functions => forall t m mc, + MetricWeakestPrecondition.call functions "spi_read" t m [] mc (fun T M RETS MC => + M = m /\ exists iol, T = t ;++ iol /\ exists ioh, mmio_trace_abstraction_relation ioh iol /\ exists (b: byte) (err : word), RETS = [word.of_Z (byte.unsigned b); err] /\ Logic.or + (word.unsigned err <> 0 /\ lightbulb_spec.spi_read_empty _ ^* ioh /\ Z.of_nat (length ioh) = patience) + (word.unsigned err = 0 /\ lightbulb_spec.spi_read word b ioh)). + + Lemma nonzero_because_high_bit_set (x : word) (H : word.unsigned (word.sru x (word.of_Z 31)) <> 0) + : word.unsigned x <> 0. + Proof. ZnWords. Qed. + + Add Ring wring : (Properties.word.ring_theory (word := word)) + (preprocess [autorewrite with rew_word_morphism], + morphism (Properties.word.ring_morph (word := word)), + constants [Properties.word_cst]). + + Import coqutil.Tactics.letexists. + Import MetricLoops. + Lemma spi_write_ok : program_logic_goal_for_function! spi_write. + Proof. + repeat straightline. + rename H into Hb. + + (* WHY do theese parentheses matter? *) + refine ((atleastonce ["b"; "busy"; "i"] (fun v T M B BUSY I MC => + b = B /\ v = word.unsigned I /\ word.unsigned I <> 0 /\ M = m /\ + exists tl, T = tl++t /\ + exists th, mmio_trace_abstraction_relation th tl /\ + lightbulb_spec.spi_write_full _ ^* th /\ + Z.of_nat (length th) + word.unsigned I = patience + )) _ _ _ _ _ _ _); + cbn [reconstruct map.putmany_of_list HList.tuple.to_list + HList.hlist.foralls HList.tuple.foralls + HList.hlist.existss HList.tuple.existss + HList.hlist.apply HList.tuple.apply + HList.hlist + List.repeat Datatypes.length + HList.polymorphic_list.repeat HList.polymorphic_list.length + PrimitivePair.pair._1 PrimitivePair.pair._2] in *. + { repeat straightline. } + { eapply (Z.lt_wf 0). } + { eexists; eexists; split; repeat straightline. + 2: { + (* TODO why do i need to do all this manually now *) + repeat split; eauto. + exists []; split; eauto. + exists []; repeat split; eauto. + - unfold mmio_trace_abstraction_relation; eauto. + - apply kleene_empty. + - rewrite word.unsigned_of_Z; reflexivity. + } + exfalso. ZnWords. } + + { admit. } + (*{ repeat (split; trivial; []).*) + (* subst i. rewrite word.unsigned_of_Z.*) + (* split.*) + (* { discriminate. }*) + (* split; trivial.*) + (* eexists; split.*) + (* { rewrite app_nil_l; trivial. }*) + (* eexists; split.*) + (* { constructor. }*) + (* split.*) + (* { constructor. }*) + (* exact eq_refl. }*) + + repeat straightline. + eapply MetricWeakestPreconditionProperties.interact_nomem; repeat straightline. + letexists; split; [exact eq_refl|]; split; [split; trivial|]. + { + cbv [isMMIOAddr addr]. + ZnWords. } + repeat straightline. + (* The hnf inside this letexists used to substitute + + c := cmd.cond (expr.op bopname.sru "busy" 31) cmd.skip + (cmd.set "i" (expr.op bopname.xor "i" "i")) : cmd.cmd + + and also unfolded WeakestPrecondition.cmd of c into + + WeakestPrecondition.dexpr m l0 cond letboundEvar /\ + ThenCorrectness /\ + ElseCorrectness + *) + letexists. letexists. split. + { repeat straightline. } + split; intros. + { (* CASE if-condition was true (word.unsigned v0 <> 0), i.e. NOP, loop exit depends on whether timeout *) + repeat straightline. (* <-- does split on a postcondition of the form + (word.unsigned br <> 0 -> loop invariant still holds) /\ + (word.unsigned br = 0 -> code after loop is fine) + which corresponds to case distinction over whether loop was exited *) + + (* fails to generate any cases now *) + admit. + + (*{ (* SUBCASE loop condition was true (do loop again) *)*) + (* eexists; split.*) + (* { repeat (split; trivial; []). subst t0.*) + (* eexists (_ ;++ cons _ nil); split; [exact eq_refl|].*) + (* eexists; split.*) + (* { refine (List.Forall2_app _ _); try eassumption.*) + (* econstructor; [|constructor].*) + (* right; eexists _, _; repeat split. }*) + (* split.*) + (* { eapply kleene_app; eauto.*) + (* refine (kleene_step _ _ nil _ (kleene_empty _)).*) + (* repeat econstructor.*) + (* ZnWords. }*) + (* { ZnWordsL. } }*) + (* { ZnWords. } }*) + (*{ (* SUBCASE loop condition was false (exit loop because of timeout *)*) + (* letexists; split; [solve[repeat straightline]|split]; repeat straightline; try contradiction.*) + (* subst t0.*) + (* eexists (_ ;++ cons _ nil); split.*) + (* { rewrite <-app_assoc; cbn [app]; f_equal. }*) + (* eexists. split.*) + (* { eapply Forall2_app; eauto.*) + (* constructor; [|constructor].*) + (* right; eauto. }*) + (* eexists. split; trivial.*) + (* { left; repeat split; eauto using nonzero_because_high_bit_set.*) + (* { (* copied from above -- trace element for "fifo full" *)*) + (* eapply kleene_app; eauto.*) + (* refine (kleene_step _ _ nil _ (kleene_empty _)).*) + (* repeat econstructor.*) + (* ZnWords. }*) + (* { ZnWordsL. } } }*) + } + + (* CASE if-condition was false (word.unsigned v0 = 0), i.e. we'll set i=i^i and exit loop *) + repeat straightline. + (* fails to generate any cases now *) + admit. + + (*{ subst i.*) + (* rewrite Properties.word.unsigned_xor_nowrap in *; rewrite Z.lxor_nilpotent in *; contradiction. }*) + (*(* evaluate condition then split if *) letexists; split; [solve[repeat straightline]|split].*) + (*1:contradiction.*) + (*repeat straightline.*) + (*eapply WeakestPreconditionProperties.interact_nomem; repeat straightline.*) + (*letexists; letexists; split; [exact eq_refl|]; split; [split; trivial|].*) + (*{ cbv [isMMIOAddr]. ZnWords. }*) + (*repeat straightline.*) + (*subst t0.*) + (*eexists (_ ;++ cons _ (cons _ nil)). split.*) + (*{ rewrite <-app_assoc. cbn [app]. f_equal. }*) + (*eexists. split.*) + (*{ eapply List.Forall2_app; eauto.*) + (* { constructor.*) + (* { left. eexists _, _; repeat split. }*) + (* { right; [|constructor].*) + (* right; eexists _, _; repeat split. } } }*) + (*eexists; split; trivial.*) + (*right.*) + (*subst busy.*) + (*split.*) + (*{ f_equal. rewrite Properties.word.unsigned_xor_nowrap; rewrite Z.lxor_nilpotent; reflexivity. }*) + (*cbv [lightbulb_spec.spi_write].*) + (*eexists _, _; split; eauto; []; split; eauto.*) + (*eexists (cons _ nil), (cons _ nil); split; cbn [app]; eauto.*) + (*split; repeat econstructor.*) + (*{ ZnWords. }*) + (*{ cbv [lightbulb_spec.spi_write_enqueue one].*) + (* repeat f_equal.*) + (* eapply Properties.word.unsigned_inj.*) + (* rewrite byte.unsigned_of_Z; cbv [byte.wrap]; rewrite Z.mod_small; ZnWords. }*) + Admitted. + + Local Ltac split_if := + lazymatch goal with + |- WeakestPrecondition.cmd _ ?c _ _ _ ?post => + let c := eval hnf in c in + lazymatch c with + | cmd.cond _ _ _ => letexists; split; [solve[repeat straightline]|split] + end + end. + + Lemma spi_read_ok : program_logic_goal_for_function! spi_read. + repeat straightline. + refine ((atleastonce ["b"; "busy"; "i"] (fun v T M B BUSY I => + v = word.unsigned I /\ word.unsigned I <> 0 /\ M = m /\ + B = word.of_Z (byte.unsigned (byte.of_Z (word.unsigned B))) /\ + exists tl, T = tl++t /\ + exists th, mmio_trace_abstraction_relation th tl /\ + lightbulb_spec.spi_read_empty _ ^* th /\ + Z.of_nat (length th) + word.unsigned I = patience + )) + _ _ _ _ _ _ _); + cbn [reconstruct map.putmany_of_list HList.tuple.to_list + HList.hlist.foralls HList.tuple.foralls + HList.hlist.existss HList.tuple.existss + HList.hlist.apply HList.tuple.apply + HList.hlist + List.repeat Datatypes.length + HList.polymorphic_list.repeat HList.polymorphic_list.length + PrimitivePair.pair._1 PrimitivePair.pair._2] in *; repeat straightline. + { exact (Z.lt_wf 0). } + { exfalso. ZnWords. } + { subst i. rewrite word.unsigned_of_Z. + split; [inversion 1|]. + split; trivial. + subst b; rewrite byte.unsigned_of_Z; cbv [byte.wrap]; + rewrite Z.mod_small; rewrite word.unsigned_of_Z. + 2: { cbv. split; congruence. } + split; trivial. + eexists nil; split; trivial. + eexists nil; split; try split; solve [constructor]. } + { eapply WeakestPreconditionProperties.interact_nomem; repeat straightline. + letexists; split; [exact eq_refl|]; split; [split; trivial|]. + { cbv [isMMIOAddr]. ZnWords. } + repeat ((split; trivial; []) || straightline || split_if). + { + letexists. split; split. + { subst v'; exact eq_refl. } + { split; trivial. + split; trivial. + split; trivial. + eexists (x2 ;++ cons _ nil); split; cbn [app]; eauto. + eexists. split. + { econstructor; try eassumption; right; eauto. } + split. + { + refine (kleene_app _ (cons _ nil) _ x3 _); eauto. + refine (kleene_step _ (cons _ nil) nil _ (kleene_empty _)). + eexists; split. + { exact eq_refl. } + { ZnWords. } } + { ZnWordsL. } } + { ZnWords. } + { ZnWords. } } + { eexists (x2 ;++ cons _ nil); split; cbn [app]; eauto. + eexists. split. + { econstructor; try eassumption; right; eauto. } + eexists (byte.of_Z (word.unsigned x)), _; split. + { f_equal. eassumption. } + left; repeat split; eauto using nonzero_because_high_bit_set. + { refine (kleene_app _ (cons _ nil) _ x3 _); eauto. + refine (kleene_step _ (cons _ nil) nil _ (kleene_empty _)). + eexists; split. + { exact eq_refl. } + { ZnWords. } } + { ZnWordsL. } } + { repeat straightline. + repeat letexists; split. + 1: split. + { repeat straightline. } + 2: { + subst v'. + subst v. + subst i. + rewrite Properties.word.unsigned_xor_nowrap, Z.lxor_nilpotent. + ZnWords. } + repeat straightline. + repeat (split; trivial; []). + split. + { subst b. + (* automatable: multi-word bitwise *) + change (255) with (Z.ones 8). + pose proof Properties.word.unsigned_range v0. + eapply Properties.word.unsigned_inj. + repeat ( + cbv [byte.wrap word.wrap]; + rewrite ?byte.unsigned_of_Z, ?word.unsigned_of_Z, ?Properties.word.unsigned_and_nowrap, + ?Z.land_ones, ?Z.mod_mod, ?Z.mod_small + by blia; + change (Z.ones 8 mod 2 ^ 32) with (Z.ones 8)). + symmetry; eapply Z.mod_small. + pose proof Z.mod_pos_bound (word.unsigned v0) (2^8) eq_refl. + clear. Z.div_mod_to_equations. blia. } + { (* copy-paste from above, trace manipulation *) + eexists (x2 ;++ cons _ nil); split; cbn [app]; eauto. + eexists. split. + { econstructor; try eassumption; right; eauto. } + subst i. + rewrite Properties.word.unsigned_xor_nowrap, Z.lxor_nilpotent in H1; contradiction. } } + { (* copy-paste from above, trace manipulation *) + eexists (x2 ;++ cons _ nil); split; cbn [app]; eauto. + eexists. split. + { econstructor; try eassumption; right; eauto. } + eexists (byte.of_Z (word.unsigned b)), _; split. + { subst b; f_equal. + (* tag:bitwise *) + (* automatable: multi-word bitwise *) + change (255) with (Z.ones 8). + pose proof Properties.word.unsigned_range v0. + eapply Properties.word.unsigned_inj. + repeat ( + cbv [byte.wrap word.wrap]; + rewrite ?byte.unsigned_of_Z, ?word.unsigned_of_Z, ?Properties.word.unsigned_and_nowrap, + ?Z.land_ones, ?Z.mod_mod, ?Z.mod_small + by blia; + change (Z.ones 8 mod 2 ^ 32) with (Z.ones 8)). + symmetry; eapply Z.mod_small. + pose proof Z.mod_pos_bound (word.unsigned v0) (2^8) eq_refl. + clear. Z.div_mod_to_equations. blia. } + (* tag:symex *) + { right; split. + { subst_words. rewrite Properties.word.unsigned_xor_nowrap, Z.lxor_nilpotent; exact eq_refl. } + eexists x3, (cons _ nil); split; cbn [app]; eauto. + split; eauto. + eexists; split; cbv [one]; trivial. + split. + (* tag:bitwise *) + { ZnWords. } + subst b. + (* automatable: multi-word bitwise *) + change (255) with (Z.ones 8). + pose proof Properties.word.unsigned_range v0. + eapply byte.unsigned_inj. + repeat ( + cbv [byte.wrap word.wrap]; + rewrite ?byte.unsigned_of_Z, ?word.unsigned_of_Z, ?Properties.word.unsigned_and_nowrap, + ?Z.land_ones, ?Z.mod_mod, ?Z.mod_small + by blia; + change (Z.ones 8 mod 2 ^ 32) with (Z.ones 8)). + trivial. } } } + Qed. + + Global Instance spec_of_spi_xchg : spec_of "spi_xchg" := fun functions => forall t m b_out, + word.unsigned b_out < 2 ^ 8 -> + WeakestPrecondition.call functions "spi_xchg" t m [b_out] (fun T M RETS => + M = m /\ exists iol, T = t ;++ iol /\ exists ioh, mmio_trace_abstraction_relation ioh iol /\ exists (b_in:byte) (err : word), RETS = [word.of_Z (byte.unsigned b_in); err] /\ Logic.or + (word.unsigned err <> 0 /\ (any +++ lightbulb_spec.spi_timeout _) ioh) + (word.unsigned err = 0 /\ lightbulb_spec.spi_xchg word (byte.of_Z (word.unsigned b_out)) b_in ioh)). + + Lemma spi_xchg_ok : program_logic_goal_for_function! spi_xchg. + Proof. + repeat ( + match goal with + | |- ?F ?a ?b ?c => + match F with WeakestPrecondition.get => idtac end; + let f := (eval cbv beta delta [WeakestPrecondition.get] in F) in + change (f a b c); cbv beta + | H : _ /\ _ \/ ?Y /\ _, G : not ?X |- _ => + constr_eq X Y; let Z := fresh in destruct H as [|[Z ?]]; [|case (G Z)] + | H : not ?Y /\ _ \/ _ /\ _, G : ?X |- _ => + constr_eq X Y; let Z := fresh in destruct H as [[Z ?]|]; [case (Z G)|] + end || + + straightline || straightline_call || split_if || refine (conj _ _) || eauto). + + { eexists. split. + { exact eq_refl. } + eexists. split. + { eauto. } + eexists. eexists. split. + { repeat f_equal. + instantiate (1 := byte.of_Z (word.unsigned b_out)). + (* automatable: multi-word bitwise *) + change (255) with (Z.ones 8). + pose proof Properties.word.unsigned_range b_out. + eapply Properties.word.unsigned_inj; + repeat ( + cbv [word.wrap byte.wrap]; + rewrite ?byte.unsigned_of_Z, ?word.unsigned_of_Z, ?Properties.word.unsigned_and_nowrap, ?Z.land_ones, ?Z.mod_mod, ?Z.mod_small by blia; + change (Z.ones 8 mod 2 ^ 32) with (Z.ones 8)); + rewrite ?Z.mod_small; rewrite ?Z.mod_small; trivial; blia. } + left; split; eauto. + eexists nil, x0; repeat split; cbv [any choice lightbulb_spec.spi_timeout]; eauto. + rewrite app_nil_r; trivial. } + + { destruct H10; intuition eauto. + { eexists. split. + { subst a0. subst a. + rewrite List.app_assoc; trivial. } + eexists. split. + { eapply Forall2_app; eauto. } + eexists _, _; split. + { subst v; trivial. } + left; split; eauto. + eapply concat_app; cbv [any choice lightbulb_spec.spi_timeout]; eauto. } + eexists. + subst a0. + subst a. + split. + { rewrite List.app_assoc; trivial. } + eexists. + split. + { eapply Forall2_app; eauto. } + eexists _, _; split. + { subst v. eauto. } + right. split; eauto. + cbv [lightbulb_spec.spi_xchg]. + + assert (Trace__concat_app : forall T (P Q:list T->Prop) x y, P x -> Q y -> (P +++ Q) (y ++ x)). { + cbv [concat]; eauto. } + + eauto using Trace__concat_app. } + Qed. +End WithParameters. From 07eee801eed2984383baef0b595a0bd5153cf86f Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Fri, 26 Jul 2024 02:58:57 -0400 Subject: [PATCH 53/62] clean up bits of PR --- bedrock2/src/bedrock2/MetricLogging.v | 20 +++++------ bedrock2/src/bedrock2/Notations.v | 2 +- compiler/src/compiler/RegAlloc.v | 48 +++++++++++---------------- 3 files changed, 31 insertions(+), 39 deletions(-) diff --git a/bedrock2/src/bedrock2/MetricLogging.v b/bedrock2/src/bedrock2/MetricLogging.v index cc7147055..37e798737 100644 --- a/bedrock2/src/bedrock2/MetricLogging.v +++ b/bedrock2/src/bedrock2/MetricLogging.v @@ -29,18 +29,18 @@ Section Riscv. Definition subMetricLoads n log := withLoads (loads log - n) log. Definition subMetricJumps n log := withJumps (jumps log - n) log. - Definition metricAdd(metric: MetricLog -> Z) finalM initialM : Z := - Z.add (metric finalM) (metric initialM). - Definition metricSub(metric: MetricLog -> Z) finalM initialM : Z := - Z.sub (metric finalM) (metric initialM). + Definition metricAdd(metric: MetricLog -> Z) m1 m2 : Z := + Z.add (metric m1) (metric m2). + Definition metricSub(metric: MetricLog -> Z) m1 m2 : Z := + Z.sub (metric m1) (metric m2). Definition metricsOp op : MetricLog -> MetricLog -> MetricLog := - fun initialM finalM => + fun m1 m2 => mkMetricLog - (op instructions initialM finalM) - (op stores initialM finalM) - (op loads initialM finalM) - (op jumps initialM finalM). + (op instructions m1 m2) + (op stores m1 m2) + (op loads m1 m2) + (op jumps m1 m2). Definition metricsAdd := metricsOp metricAdd. Definition metricsSub := metricsOp metricSub. @@ -71,7 +71,7 @@ Infix "+" := metricsAdd : MetricH_scope. Infix "-" := metricsSub : MetricH_scope. Infix "*" := metricsMul : MetricH_scope. -#[global] Hint Unfold +#[export] Hint Unfold withInstructions withLoads withStores diff --git a/bedrock2/src/bedrock2/Notations.v b/bedrock2/src/bedrock2/Notations.v index 5d1a500e9..c8176b6e2 100644 --- a/bedrock2/src/bedrock2/Notations.v +++ b/bedrock2/src/bedrock2/Notations.v @@ -1,4 +1,4 @@ -Require Import coqutil.Macros.subst coqutil.Macros.unique. +Require Import coqutil.Macros.subst. Notation "' x <- a | y ; f" := (match a with diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index 18c8a9551..20ceb8882 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -9,6 +9,7 @@ Require Import coqutil.Datatypes.ListSet. Require Import coqutil.Tactics.fwd. Require Import coqutil.Tactics.autoforward. Require Import compiler.Registers. +Require Import bedrock2.MetricLogging. Require Import bedrock2.MetricCosts. Open Scope Z_scope. @@ -698,8 +699,7 @@ Fixpoint check(m: list (srcvar * impvar))(s: stmt)(s': stmt'){struct s}: result assert_in y y' m;; assert (Z.eqb ofs ofs') err;; assert (access_size_beq sz sz') err;; - a <- assignment m x x';; - Success a + assignment m x x' | SStore sz x y ofs, SStore sz' x' y' ofs' => assert_in x x' m;; assert_in y y' m;; @@ -712,26 +712,22 @@ Fixpoint check(m: list (srcvar * impvar))(s: stmt)(s': stmt'){struct s}: result assert (negb (Z.eqb x' y')) err;; assert (access_size_beq sz sz') err;; assert (List.list_eqb Byte.eqb bs bs') err;; - a <- (assignment m x x');; - Success a + assignment m x x' | SStackalloc x n body, SStackalloc x' n' body' => assert (Z.eqb n n') err;; a <- assignment m x x';; check a body body' | SLit x z, SLit x' z' => assert (Z.eqb z z') err;; - a <- assignment m x x';; - Success a + assignment m x x' | SOp x op y z, SOp x' op' y' z' => assert_in y y' m;; assert_in_op z z' m;; assert (bopname_beq op op') err;; - a <- assignment m x x';; - Success a + assignment m x x' | SSet x y, SSet x' y' => assert_in y y' m;; - a <- assignment m x x';; - Success a + assignment m x x' | SIf c s1 s2, SIf c' s1' s2' => check_bcond m c c';; m1 <- check m s1 s1';; @@ -1140,16 +1136,13 @@ Section CheckerCorrect. Lemma check_regs_cost_SIf: forall (cond : bcond) (bThen bElse : stmt) (corresp : list (string * Z)) (cond0 : bcond') (s'1 s'2 : stmt') (a0 a1 : list (string * Z)) - (mc mcL : MetricLogging.MetricLog), + (mc mcL : MetricLog), check_bcond corresp cond cond0 = Success tt -> check corresp bThen s'1 = Success a0 -> check corresp bElse s'2 = Success a1 -> - forall mc' mcH' : MetricLogging.MetricLog, - MetricLogging.metricsLeq - (MetricLogging.metricsSub mc' (exec.cost_SIf isRegZ cond0 mcL)) - (MetricLogging.metricsSub mcH' (exec.cost_SIf isRegStr cond mc)) -> - MetricLogging.metricsLeq (MetricLogging.metricsSub mc' mcL) - (MetricLogging.metricsSub mcH' mc). + forall mc' mcH' : MetricLog, + (mc' - exec.cost_SIf isRegZ cond0 mcL <= mcH' - exec.cost_SIf isRegStr cond mc)%metricsH -> + (mc' - mcL <= mcH' - mc)%metricsH. Proof. intros. repeat (unfold check_bcond, assert_in, assignment in *; fwd). @@ -1163,15 +1156,14 @@ Section CheckerCorrect. Lemma check_regs_cost_SLoop_false: forall (cond : bcond) (body1 body2 : stmt) (corresp' : list (string * Z)) (s'1 : stmt') - (cond0 : bcond') (s'2 : stmt') (mc mcL : MetricLogging.MetricLog) (a : list (string * Z)), + (cond0 : bcond') (s'2 : stmt') (mc mcL : MetricLog) (a : list (string * Z)), check a body1 s'1 = Success corresp' -> check_bcond corresp' cond cond0 = Success tt -> forall a2 : list (string * Z), check corresp' body2 s'2 = Success a2 -> - forall mc' mcH' : MetricLogging.MetricLog, - MetricLogging.metricsLeq (MetricLogging.metricsSub mc' mcL) (MetricLogging.metricsSub mcH' mc) -> - MetricLogging.metricsLeq (MetricLogging.metricsSub (exec.cost_SLoop_false isRegZ cond0 mc') mcL) - (MetricLogging.metricsSub (exec.cost_SLoop_false isRegStr cond mcH') mc). + forall mc' mcH' : MetricLog, + (mc' - mcL <= mcH' - mc)%metricsH -> + (exec.cost_SLoop_false isRegZ cond0 mc' - mcL <= exec.cost_SLoop_false isRegStr cond mcH' - mc)%metricsH. Proof. intros. repeat (unfold check_bcond, assert_in, assignment in *; fwd). @@ -1195,7 +1187,7 @@ Section CheckerCorrect. (* : checker_hints.*) Ltac a := - repeat match goal with | |- MetricLogging.metricsLeq _ _ => fail 1 | _ => econstructor end; + repeat match goal with | |- metricsLeq _ _ => fail 1 | _ => econstructor end; eauto 10 with checker_hints. Ltac b := unfold assert_in, assignment, check_regs in *; cost_hammer. @@ -1208,7 +1200,7 @@ Section CheckerCorrect. states_compat lH (precond corresp s s') lL -> exec PreSpill isRegZ e' s' t m lL mcL (fun t' m' lL' mcL' => exists lH' mcH', states_compat lH' corresp' lL' /\ - MetricLogging.metricsLeq (MetricLogging.metricsSub mcL' mcL) (MetricLogging.metricsSub mcH' mcH) /\ + (mcL' - mcL <= mcH' - mcH)%metricsH /\ post t' m' lH' mcH'). Proof. induction 2; intros; @@ -1308,12 +1300,12 @@ Section CheckerCorrect. unfold loop_inv in SC. rewrite E in SC. eapply exec.loop with - (mid2 := (fun (t'0 : Semantics.trace) (m'0 : mem) (lL' : impLocals) (mcL' : MetricLogging.MetricLog) => - exists (lH' : srcLocals) (mcH' : MetricLogging.MetricLog), + (mid2 := (fun (t'0 : Semantics.trace) (m'0 : mem) (lL' : impLocals) (mcL' : MetricLog) => + exists (lH' : srcLocals) (mcH' : MetricLog), states_compat lH' a1 lL' /\ (exists mcHmid mcLmid, - MetricLogging.metricsLeq (MetricLogging.metricsSub mcLmid mcL) (MetricLogging.metricsSub mcHmid mc) /\ - MetricLogging.metricsLeq (MetricLogging.metricsSub mcL' mcLmid) (MetricLogging.metricsSub mcH' mcHmid)) /\ + mcLmid - mcL <= mcHmid - mc /\ + mcL' - mcLmid <= mcH' - mcHmid)%metricsH /\ mid2 t'0 m'0 lH' mcH')). + eapply IH1. 1: eassumption. eapply states_compat_precond. exact SC. + cbv beta. intros. fwd. eauto using states_compat_eval_bcond_None. From e4358eb18e1491fc59797a2e216f18f38f97c78b Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Fri, 26 Jul 2024 03:46:31 -0400 Subject: [PATCH 54/62] correct call/spilling costs --- bedrock2/src/bedrock2/MetricCosts.v | 18 ++++++------------ bedrock2/src/bedrock2/MetricSemantics.v | 4 ++-- .../src/bedrock2/MetricWeakestPrecondition.v | 4 ++-- compiler/src/compiler/FlatImp.v | 8 ++++---- compiler/src/compiler/FlatToRiscvFunctions.v | 4 ++-- compiler/src/compiler/Pipeline.v | 14 +++++++------- compiler/src/compiler/Spilling.v | 8 ++++++-- 7 files changed, 29 insertions(+), 31 deletions(-) diff --git a/bedrock2/src/bedrock2/MetricCosts.v b/bedrock2/src/bedrock2/MetricCosts.v index d6ec0d5e4..630f80c08 100644 --- a/bedrock2/src/bedrock2/MetricCosts.v +++ b/bedrock2/src/bedrock2/MetricCosts.v @@ -21,16 +21,10 @@ Section FlatImpExec. | PostSpill => mkMetricLog 50 50 50 50 end + mc. - Definition cost_call_internal mc := + Definition cost_call mc := match phase with - | PreSpill => mkMetricLog 100 100 100 100 - | PostSpill => mkMetricLog 50 50 50 50 - end + mc. - - Definition cost_call_external mc := - match phase with - | PreSpill => mkMetricLog 100 100 100 100 - | PostSpill => mkMetricLog 50 50 50 50 + | PreSpill => mkMetricLog 200 200 200 200 + | PostSpill => mkMetricLog 100 100 100 100 end + mc. (* TODO think about a non-fixed bound on the cost of function preamble and postamble *) @@ -119,9 +113,9 @@ Definition isRegStr (var : String.string) : bool := String.prefix "reg_" var. Ltac cost_unfold := - unfold cost_interact, cost_call_internal, cost_call_external, cost_load, - cost_store, cost_inlinetable, cost_stackalloc, cost_lit, cost_op, cost_set, - cost_if, cost_loop_true, cost_loop_false, EmptyMetricLog in *. + unfold cost_interact, cost_call, cost_load, cost_store, cost_inlinetable, + cost_stackalloc, cost_lit, cost_op, cost_set, cost_if, cost_loop_true, + cost_loop_false, EmptyMetricLog in *. Ltac cost_destr := repeat match goal with diff --git a/bedrock2/src/bedrock2/MetricSemantics.v b/bedrock2/src/bedrock2/MetricSemantics.v index bde80ac64..059ae32ff 100644 --- a/bedrock2/src/bedrock2/MetricSemantics.v +++ b/bedrock2/src/bedrock2/MetricSemantics.v @@ -146,11 +146,11 @@ Module exec. Section WithParams. params rets fbody (_ : map.get e fname = Some (params, rets, fbody)) args mc' (_ : eval_call_args m l arges mc = Some (args, mc')) lf (_ : map.of_list_zip params args = Some lf) - mid (_ : exec fbody t m lf (cost_call_internal PreSpill mc') mid) + mid (_ : exec fbody t m lf mc' mid) (_ : forall t' m' st1 mc'', mid t' m' st1 mc'' -> exists retvs, map.getmany_of_list st1 rets = Some retvs /\ exists l', map.putmany_of_list_zip binds retvs l = Some l' /\ - post t' m' l' (cost_call_external PreSpill mc'')) + post t' m' l' (cost_call PreSpill mc'')) : exec (cmd.call binds fname arges) t m l mc post | interact binds action arges t m l mc post diff --git a/bedrock2/src/bedrock2/MetricWeakestPrecondition.v b/bedrock2/src/bedrock2/MetricWeakestPrecondition.v index 46a079819..d64c4ad6b 100644 --- a/bedrock2/src/bedrock2/MetricWeakestPrecondition.v +++ b/bedrock2/src/bedrock2/MetricWeakestPrecondition.v @@ -107,8 +107,8 @@ Section WeakestPrecondition. | cmd.while _ _ => MetricSemantics.exec e c t m l mc post | cmd.call binds fname arges => exists args mc', dexprs m l arges mc (args, mc') /\ - MetricSemantics.call e fname t m args (cost_call_internal PreSpill mc') (fun t m rets mc'' => - exists l', map.putmany_of_list_zip binds rets l = Some l' /\ post t m l' (cost_call_external PreSpill mc'')) + MetricSemantics.call e fname t m args mc' (fun t m rets mc'' => + exists l', map.putmany_of_list_zip binds rets l = Some l' /\ post t m l' (cost_call PreSpill mc'')) | cmd.interact binds action arges => exists args mc', dexprs m l arges mc (args, mc') /\ exists mKeep mGive, map.split m mKeep mGive /\ diff --git a/compiler/src/compiler/FlatImp.v b/compiler/src/compiler/FlatImp.v index 829eac6c0..403c4443a 100644 --- a/compiler/src/compiler/FlatImp.v +++ b/compiler/src/compiler/FlatImp.v @@ -342,13 +342,13 @@ Module exec. map.get e fname = Some (params, rets, fbody) -> map.getmany_of_list l args = Some argvs -> map.putmany_of_list_zip params argvs map.empty = Some st0 -> - exec fbody t m st0 (cost_call_internal phase mc) outcome -> + exec fbody t m st0 mc outcome -> (forall t' m' mc' st1, outcome t' m' st1 mc' -> exists retvs l', map.getmany_of_list st1 rets = Some retvs /\ map.putmany_of_list_zip binds retvs l = Some l' /\ - post t' m' l' (cost_call_external phase mc')) -> + post t' m' l' (cost_call phase mc')) -> exec (SCall binds fname args) t m l mc post | load: forall t m l mc sz x a o v addr post, map.get l a = Some addr -> @@ -451,12 +451,12 @@ Module exec. map.get e fname = Some (params, rets, fbody) -> map.getmany_of_list l args = Some argvs -> map.putmany_of_list_zip params argvs map.empty = Some st -> - exec fbody t m st (cost_call_internal phase mc) + exec fbody t m st mc (fun t' m' st' mc' => exists retvs l', map.getmany_of_list st' rets = Some retvs /\ map.putmany_of_list_zip binds retvs l = Some l' /\ - post t' m' l' (cost_call_external phase mc')) -> + post t' m' l' (cost_call phase mc')) -> exec (SCall binds fname args) t m l mc post. Proof. intros. eapply call; try eassumption. diff --git a/compiler/src/compiler/FlatToRiscvFunctions.v b/compiler/src/compiler/FlatToRiscvFunctions.v index 6defbb0ae..3aa0b88e7 100644 --- a/compiler/src/compiler/FlatToRiscvFunctions.v +++ b/compiler/src/compiler/FlatToRiscvFunctions.v @@ -1633,8 +1633,8 @@ Section Proofs. split; eauto 8 with map_hints. split; eauto 8 with map_hints. split; eauto 8 with map_hints. - (* cost_SCall constraint: cost_SCall_L + (1,1,1,0) <= cost_SCall_internal + cost_SCall_external *) - unfold cost_SCall_L, cost_call_internal, cost_call_external in *. + (* cost_SCall constraint: cost_SCall_L + (1,1,1,0) <= cost_call TODO check *) + unfold cost_SCall_L, cost_call in *. MetricsToRiscv.solve_MetricLog. - idtac "Case compile_stmt_correct/SLoad". diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index bc2ceb4ff..c58ac300c 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -164,7 +164,7 @@ Section WithWordAndMem. exists argnames retnames fbody l, map.get e f = Some (argnames, retnames, fbody) /\ map.of_list_zip argnames argvals = Some l /\ - Exec e fbody t m l (cost_call_internal PreSpill mc) (fun t' m' l' mc' => + Exec e fbody t m l (cost_spill_spec mc) (fun t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals mc'). @@ -293,7 +293,7 @@ Section WithWordAndMem. + eapply @freshNameGenState_disjoint_fbody. - simpl. intros. fwd. eexists. split. + eauto using map.getmany_of_list_extends. - + eexists. split; [|eassumption]. cost_unfold; solve_MetricLog. + + eexists. split; [|eassumption]. unfold cost_spill_spec in *; solve_MetricLog. Qed. Lemma useimmediate_functions_NoDup: forall funs funs', @@ -334,7 +334,7 @@ Section WithWordAndMem. - eapply useImmediate_correct_aux; eauto. - simpl. destruct 1 as (?&?&?&?&?). repeat (eexists; split; try eassumption). - cost_unfold; solve_MetricLog. + unfold cost_spill_spec in *; solve_MetricLog. Qed. @@ -376,7 +376,7 @@ Section WithWordAndMem. exists retvals. split. + erewrite MapEauto.agree_on_getmany; [ eauto | eapply MapEauto.agree_on_comm; [ eassumption ] ]. - + eexists; split; eauto. unfold cost_call_internal in *. solve_MetricLog. + + eexists; split; eauto. unfold cost_spill_spec in *; solve_MetricLog. Qed. Lemma regalloc_functions_NoDup: forall funs funs', @@ -432,7 +432,7 @@ Section WithWordAndMem. rewrite H1 in P'. inversion P'. exact Cp. - simpl. intros. fwd. eexists. split. + eauto using states_compat_getmany. - + eexists. split; [|eassumption]. cost_unfold; solve_MetricLog. + + eexists. split; [|eassumption]. unfold cost_spill_spec in *; solve_MetricLog. Qed. Ltac debool := @@ -591,7 +591,7 @@ Section WithWordAndMem. Some (argnames, retnames, fbody) /\ map.of_list_zip argnames argvals = Some l /\ MetricSemantics.exec (map.of_list functions) fbody t mH l - (cost_call_internal PreSpill mc) + (cost_spill_spec mc) (fun t' m' l' mc' => exists retvals: list word, map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals mc')) -> @@ -666,7 +666,7 @@ Section WithWordAndMem. NoDup (map fst fs) -> compile fs = Success (instrs, finfo, req_stack_size) -> MetricWeakestPrecondition.call (map.of_list fs) fname t mH argvals - (cost_call_internal PreSpill mc) post -> + (cost_spill_spec mc) post -> forall mcL, map.get (map.of_list finfo) fname = Some f_rel_pos -> req_stack_size <= word.unsigned (word.sub stack_hi stack_lo) / bytes_per_word -> diff --git a/compiler/src/compiler/Spilling.v b/compiler/src/compiler/Spilling.v index 01acb0903..14c656809 100644 --- a/compiler/src/compiler/Spilling.v +++ b/compiler/src/compiler/Spilling.v @@ -1287,11 +1287,15 @@ Section Spilling. post t1' m1' l1' mc1' /\ (mc2' - mc2 <= mc1' - mc1)%metricsH). + (* TODO tighter / non-fixed bound *) + Definition cost_spill_spec mc := + (mkMetricLog 100 100 100 100 + mc)%metricsH. + Definition call_spec(e: env) '(argnames, retnames, fbody) (t: Semantics.trace)(m: mem)(argvals: list word)(mc: MetricLog) (post: Semantics.trace -> mem -> list word -> MetricLog -> Prop): Prop := forall l, map.of_list_zip argnames argvals = Some l -> - execpre e fbody t m l (cost_call_internal PreSpill mc) (fun t' m' l' mc' => + execpre e fbody t m l (cost_spill_spec mc) (fun t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals mc'). @@ -1475,7 +1479,7 @@ Section Spilling. { eassumption. } { add_bounds. - unfold cost_stackalloc, cost_call_internal in *. + unfold cost_stackalloc, cost_spill_spec in *. (* TODO XXX *) destruct (isRegZ fp); solve_MetricLog. } Qed. From ae6c57b5de03ffc0f1205d05da8721b582f36267 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Fri, 26 Jul 2024 05:29:28 -0400 Subject: [PATCH 55/62] naming and tactic improvements --- compiler/src/compiler/FlatToRiscvFunctions.v | 86 +++++--------------- compiler/src/compiler/LowerPipeline.v | 12 +-- compiler/src/compiler/Pipeline.v | 4 +- 3 files changed, 25 insertions(+), 77 deletions(-) diff --git a/compiler/src/compiler/FlatToRiscvFunctions.v b/compiler/src/compiler/FlatToRiscvFunctions.v index 3aa0b88e7..84aacc06e 100644 --- a/compiler/src/compiler/FlatToRiscvFunctions.v +++ b/compiler/src/compiler/FlatToRiscvFunctions.v @@ -432,7 +432,7 @@ Section Proofs. Local Notation exec := (exec PostSpill isRegZ). - Definition cost_SCall_L mc := + Definition cost_compile_spec mc := Platform.MetricLogging.addMetricInstructions 95 (Platform.MetricLogging.addMetricJumps 95 (Platform.MetricLogging.addMetricLoads 95 @@ -501,7 +501,7 @@ Section Proofs. (of_list (list_union Z.eqb (List.firstn binds_count (reg_class.all reg_class.arg)) [])) (singleton_set RegisterNames.ra)) (getRegs finalL) /\ - (getMetrics finalL - cost_SCall_L (getMetrics mach) <= + (getMetrics finalL - cost_compile_spec (getMetrics mach) <= lowerMetrics (finalMetricsH - mc))%metricsL /\ goodMachine finalTrace finalMH finalRegsH g finalL). Proof. @@ -1112,7 +1112,7 @@ Section Proofs. end end. cbn in H2p6. - (* cost_SCall constraint: cost_SCall_L >= (...93...) i think? *) + (* cost_compile_spec constraint: cost_compile_spec >= (...93...) i think? *) blia. + rename l into lH, finalRegsH into lFH', finalRegsH' into lH', st0 into lFH, @@ -1318,8 +1318,7 @@ Section Proofs. eapply List.Forall_filter. intros *. intro E. destr (reg_class.get a); try discriminate E. unfold reg_class.get in E0. fwd. - unfold FlatToRiscvDef.valid_FlatImp_var. - destruct_one_match_hyp. + unfold FlatToRiscvDef.valid_FlatImp_var. destruct_one_match_hyp. -- fwd. blia. -- destruct_one_match_hyp. 1: discriminate. destruct_one_match_hyp; discriminate. @@ -1351,6 +1350,15 @@ Section Proofs. Qed. + Ltac finishcost := + scost_unfold; + repeat match goal with + | H : ForallVars_bcond _ ?cond |- _ => destruct cond eqn:?; unfold ForallVars_bcond in H + | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * + | H : _ /\ _ |- _ => destruct H + end; + MetricsToRiscv.solve_MetricLog. + Lemma compile_stmt_correct: (forall resvars extcall argvars, compiles_FlatToRiscv_correctly compile_ext_call @@ -1633,8 +1641,8 @@ Section Proofs. split; eauto 8 with map_hints. split; eauto 8 with map_hints. split; eauto 8 with map_hints. - (* cost_SCall constraint: cost_SCall_L + (1,1,1,0) <= cost_call TODO check *) - unfold cost_SCall_L, cost_call in *. + (* cost_compile_spec constraint: cost_compile_spec + (1,1,1,0) <= cost_call *) + unfold cost_compile_spec, cost_call in *. MetricsToRiscv.solve_MetricLog. - idtac "Case compile_stmt_correct/SLoad". @@ -1948,21 +1956,7 @@ Section Proofs. { eapply run_Jal0; try safe_sidecond. solve_divisibleBy4. } simpl_MetricRiscvMachine_get_set. - intros. destruct_RiscvMachine mid. fwd. run1done. - destruct cond eqn:E; scost_unfold; - unfold ForallVars_bcond in *; - repeat match goal with - | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H - end; - repeat match goal with - | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end; - try MetricsToRiscv.solve_MetricLog. - destruct H6p0. - repeat match goal with - | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end. - MetricsToRiscv.solve_MetricLog. + intros. destruct_RiscvMachine mid. fwd. run1done. finishcost. - idtac "Case compile_stmt_correct/SIf/Else". (* execute branch instruction, which will jump over then-branch *) @@ -1987,21 +1981,7 @@ Section Proofs. all: try safe_sidecond. * (* at end of else-branch, i.e. also at end of if-then-else, just prove that computed post satisfies required post *) - simpl. intros. destruct_RiscvMachine middle. fwd. subst. run1done. - destruct cond eqn:E; scost_unfold; - unfold ForallVars_bcond in *; - repeat match goal with - | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H - end; - repeat match goal with - | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end; - try MetricsToRiscv.solve_MetricLog. - destruct H6p0. - repeat match goal with - | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end. - MetricsToRiscv.solve_MetricLog. + simpl. intros. destruct_RiscvMachine middle. fwd. subst. run1done. finishcost. - idtac "Case compile_stmt_correct/SLoop". match goal with @@ -2073,21 +2053,7 @@ Section Proofs. all: try safe_sidecond. } (* at end of loop, just prove that computed post satisfies required post *) - simpl. intros. destruct_RiscvMachine middle. fwd. run1done. - destruct cond eqn:blah; scost_unfold; - unfold ForallVars_bcond in *; - repeat match goal with - | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H - end; - repeat match goal with - | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end; - try MetricsToRiscv.solve_MetricLog. - destruct H11p0. - repeat match goal with - | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end. - MetricsToRiscv.solve_MetricLog. + simpl. intros. destruct_RiscvMachine middle. fwd. run1done. finishcost. * (* false: done, jump over body2 *) eapply runsToStep. { @@ -2096,21 +2062,7 @@ Section Proofs. try safe_sidecond. } simpl_MetricRiscvMachine_get_set. - intros. destruct_RiscvMachine mid. fwd. run1done. - destruct cond eqn:blah; scost_unfold; - unfold ForallVars_bcond in *; - repeat match goal with - | H : ForallVars_bcond _ |- _ => unfold ForallVars_bcond in H; destruct H - end; - repeat match goal with - | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end; - try MetricsToRiscv.solve_MetricLog. - destruct H11p0. - repeat match goal with - | H : valid_FlatImp_var _ |- _ => apply valid_FlatImp_var_isRegZ in H; rewrite H in * - end. - MetricsToRiscv.solve_MetricLog. + intros. destruct_RiscvMachine mid. fwd. run1done. finishcost. - idtac "Case compile_stmt_correct/SSeq". on hyp[(FlatImpConstraints.uses_standard_arg_regs s1); runsTo] diff --git a/compiler/src/compiler/LowerPipeline.v b/compiler/src/compiler/LowerPipeline.v index 15be08c5f..37ed158ff 100644 --- a/compiler/src/compiler/LowerPipeline.v +++ b/compiler/src/compiler/LowerPipeline.v @@ -396,7 +396,7 @@ Section LowerPipeline. forall p_funcs stack_start stack_pastend ret_addr Rdata Rexec (initial: MetricRiscvMachine), map.get initial.(getRegs) RegisterNames.ra = Some ret_addr -> initial.(getLog) = t -> - raiseMetrics (cost_SCall_L initial.(getMetrics)) = mc -> + raiseMetrics (cost_compile_spec initial.(getMetrics)) = mc -> word.unsigned ret_addr mod 4 = 0 -> arg_regs_contain initial.(getRegs) argvals -> req_stack_size <= word.unsigned (word.sub stack_pastend stack_start) / bytes_per_word -> @@ -692,16 +692,12 @@ Section LowerPipeline. eapply map.getmany_of_list_length. exact GM. + eexists. split; [|eassumption]. + subst. apply raise_metrics_ineq in H10p3. - destruct (getMetrics final). - destruct (getMetrics initial). - unfold raiseMetrics, lowerMetrics in *. unfold_MetricLog. - simpl_MetricLog. - simpl in *. - inversion H2. + cbn in H10p3. + cbn. solve_MetricLog. - (* XXX ???? *) + eapply only_differ_subset. 1: eassumption. rewrite ListSet.of_list_list_union. rewrite ?singleton_set_eq_of_list. diff --git a/compiler/src/compiler/Pipeline.v b/compiler/src/compiler/Pipeline.v index c58ac300c..f91f428e8 100644 --- a/compiler/src/compiler/Pipeline.v +++ b/compiler/src/compiler/Pipeline.v @@ -609,7 +609,7 @@ Section WithWordAndMem. word.unsigned ret_addr mod 4 = 0 -> arg_regs_contain initial.(getRegs) argvals -> initial.(getLog) = t -> - raiseMetrics (cost_SCall_L initial.(getMetrics)) = mcL -> + raiseMetrics (cost_compile_spec initial.(getMetrics)) = mcL -> machine_ok p_funcs stack_lo stack_hi instrs mH Rdata Rexec initial -> runsTo initial (fun final : MetricRiscvMachine => exists mH' retvals, @@ -676,7 +676,7 @@ Section WithWordAndMem. word.unsigned ret_addr mod 4 = 0 -> arg_regs_contain initial.(getRegs) argvals -> initial.(getLog) = t -> - raiseMetrics (cost_SCall_L initial.(getMetrics)) = mcL -> + raiseMetrics (cost_compile_spec initial.(getMetrics)) = mcL -> machine_ok p_funcs stack_lo stack_hi instrs mH Rdata Rexec initial -> runsTo initial (fun final : MetricRiscvMachine => exists mH' retvals, From bdecbb94fcd6715653d7266818fcdc9ff3e27e8c Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Fri, 26 Jul 2024 18:03:55 -0400 Subject: [PATCH 56/62] fix qed slowness metriclit and most of s should go somewhere else eventually --- bedrock2/src/bedrock2Examples/metric_ipow.v | 33 +++++++++++++++++---- 1 file changed, 27 insertions(+), 6 deletions(-) diff --git a/bedrock2/src/bedrock2Examples/metric_ipow.v b/bedrock2/src/bedrock2Examples/metric_ipow.v index d8ebc57b7..2ba342bdf 100644 --- a/bedrock2/src/bedrock2Examples/metric_ipow.v +++ b/bedrock2/src/bedrock2Examples/metric_ipow.v @@ -100,15 +100,36 @@ Proof. blia. Qed. -Lemma metriclit : forall a b c d a' b' c' d', - metricsAdd (mkMetricLog a b c d) (mkMetricLog a' b' c' d') = - mkMetricLog (a+a') (b+b') (c+c') (d+d'). -Proof. reflexivity. Qed. +Lemma metriclit : forall a b c d a' b' c' d' mc, + metricsAdd (mkMetricLog a b c d) (metricsAdd (mkMetricLog a' b' c' d') mc) = + metricsAdd (mkMetricLog (a+a') (b+b') (c+c') (d+d')) mc. +Proof. intros. rewrite MetricArith.add_assoc. reflexivity. Qed. Ltac s := unfold initCost, iterCost, endCost in *; + repeat ( + let H := match goal with + | H : context[cost_interact] |- _ => H + | H : context[cost_call] |- _ => H + | H : context[cost_load] |- _ => H + | H : context[cost_store] |- _ => H + | H : context[cost_inlinetable] |- _ => H + | H : context[cost_stackalloc] |- _ => H + | H : context[cost_lit] |- _ => H + | H : context[cost_op] |- _ => H + | H : context[cost_set] |- _ => H + | H : context[cost_if] |- _ => H + | H : context[cost_loop_true] |- _ => H + | H : context[cost_loop_false] |- _ => H + end in + let t := type of H in + let t' := eval cbv [cost_interact cost_call cost_load cost_store + cost_inlinetable cost_stackalloc cost_lit cost_op cost_set + cost_if cost_loop_true cost_loop_false] in t in + replace t with t' in H by (symmetry; reflexivity) + ); cost_unfold; - repeat rewrite MetricArith.add_assoc in *; cbn in *; - repeat rewrite metriclit in *; cbn in *; + cbn in *; + repeat rewrite metriclit in *; solve_MetricLog. Lemma ipow_ok : program_logic_goal_for_function! ipow. From 3aad17e966de56a952db1cb66c308ab0acc04aaf Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Wed, 31 Jul 2024 16:25:06 -0400 Subject: [PATCH 57/62] clean up metrics examples to mergable state --- bedrock2/src/bedrock2/MetricCosts.v | 31 +- bedrock2/src/bedrock2/MetricLogging.v | 40 +- bedrock2/src/bedrock2Examples/metric_SPI.v | 466 -------------------- bedrock2/src/bedrock2Examples/metric_ipow.v | 27 -- 4 files changed, 41 insertions(+), 523 deletions(-) delete mode 100644 bedrock2/src/bedrock2Examples/metric_SPI.v diff --git a/bedrock2/src/bedrock2/MetricCosts.v b/bedrock2/src/bedrock2/MetricCosts.v index 630f80c08..6cde700d7 100644 --- a/bedrock2/src/bedrock2/MetricCosts.v +++ b/bedrock2/src/bedrock2/MetricCosts.v @@ -112,10 +112,35 @@ Definition isRegZ (var : Z) : bool := Definition isRegStr (var : String.string) : bool := String.prefix "reg_" var. +(* awkward tactic use to avoid Qed slowness *) +(* symmetry; reflexivity is black magic that shuffles orders of terms to make + heuristics choose the right things *) Ltac cost_unfold := - unfold cost_interact, cost_call, cost_load, cost_store, cost_inlinetable, - cost_stackalloc, cost_lit, cost_op, cost_set, cost_if, cost_loop_true, - cost_loop_false, EmptyMetricLog in *. + repeat ( + let H := match goal with + | H : context[cost_interact] |- _ => H + | H : context[cost_call] |- _ => H + | H : context[cost_load] |- _ => H + | H : context[cost_store] |- _ => H + | H : context[cost_inlinetable] |- _ => H + | H : context[cost_stackalloc] |- _ => H + | H : context[cost_lit] |- _ => H + | H : context[cost_op] |- _ => H + | H : context[cost_set] |- _ => H + | H : context[cost_if] |- _ => H + | H : context[cost_loop_true] |- _ => H + | H : context[cost_loop_false] |- _ => H + end in + let t := type of H in + let t' := eval cbv [cost_interact cost_call cost_load cost_store + cost_inlinetable cost_stackalloc cost_lit cost_op cost_set + cost_if cost_loop_true cost_loop_false] in t in + replace t with t' in H by (symmetry; reflexivity) + ); + cbv [cost_interact cost_call cost_load cost_store cost_inlinetable + cost_stackalloc cost_lit cost_op cost_set cost_if cost_loop_true + cost_loop_false]; + unfold EmptyMetricLog in *. Ltac cost_destr := repeat match goal with diff --git a/bedrock2/src/bedrock2/MetricLogging.v b/bedrock2/src/bedrock2/MetricLogging.v index 37e798737..b8083f747 100644 --- a/bedrock2/src/bedrock2/MetricLogging.v +++ b/bedrock2/src/bedrock2/MetricLogging.v @@ -113,12 +113,25 @@ Ltac fold_MetricLog := Ltac simpl_MetricLog := cbn [instructions loads stores jumps] in *. +(* need this to define solve_MetricLog, but need solve_MetricLog inside of MetricArith, oops *) +Lemma add_assoc' : forall n m p, (n + (m + p) = n + m + p)%metricsH. +Proof. intros. unfold_MetricLog. f_equal; apply Z.add_assoc. Qed. + +Lemma metriclit : forall a b c d a' b' c' d' mc, + metricsAdd (mkMetricLog a b c d) (metricsAdd (mkMetricLog a' b' c' d') mc) = + metricsAdd (mkMetricLog (a+a') (b+b') (c+c') (d+d')) mc. +Proof. intros. rewrite add_assoc'. reflexivity. Qed. + +Ltac flatten_MetricLog := repeat rewrite metriclit in *. + Ltac solve_MetricLog := + flatten_MetricLog; repeat unfold_MetricLog; repeat simpl_MetricLog; blia. Ltac solve_MetricLog_piecewise := + flatten_MetricLog; repeat unfold_MetricLog; repeat simpl_MetricLog; f_equal; blia. @@ -163,30 +176,3 @@ Create HintDb metric_arith. #[export] Hint Resolve MetricArith.le_trans MetricArith.le_refl MetricArith.add_0_r MetricArith.sub_0_r MetricArith.add_comm MetricArith.add_assoc : metric_arith. #[export] Hint Resolve <- MetricArith.le_sub_mono : metric_arith. #[export] Hint Resolve -> MetricArith.le_sub_mono : metric_arith. - -Lemma applyAddInstructions n a b c d : addMetricInstructions n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a+n; stores := b; loads := c; jumps := d |}. Proof. auto. Qed. -Lemma applyAddStores n a b c d : addMetricStores n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a; stores := b+n; loads := c; jumps := d |}. Proof. auto. Qed. -Lemma applyAddLoads n a b c d : addMetricLoads n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a; stores := b; loads := c+n; jumps := d |}. Proof. auto. Qed. -Lemma applyAddJumps n a b c d : addMetricJumps n {| instructions := a; stores := b; loads := c; jumps := d |} = {| instructions := a; stores := b; loads := c; jumps := d+n |}. Proof. auto. Qed. - -Ltac applyAddMetricsGoal := ( - repeat ( - try rewrite applyAddInstructions; - try rewrite applyAddStores; - try rewrite applyAddLoads; - try rewrite applyAddJumps - ); - repeat rewrite <- Z.add_assoc; - cbn [Z.add Pos.add Pos.succ] - ). - -Ltac applyAddMetrics H := ( - repeat ( - try rewrite applyAddInstructions in H; - try rewrite applyAddStores in H; - try rewrite applyAddLoads in H; - try rewrite applyAddJumps in H - ); - repeat rewrite <- Z.add_assoc in H; - cbn [Z.add Pos.add Pos.succ] in H -). diff --git a/bedrock2/src/bedrock2Examples/metric_SPI.v b/bedrock2/src/bedrock2Examples/metric_SPI.v deleted file mode 100644 index 9acfc1e2e..000000000 --- a/bedrock2/src/bedrock2Examples/metric_SPI.v +++ /dev/null @@ -1,466 +0,0 @@ -Require Import bedrock2.Syntax bedrock2.NotationsCustomEntry Coq.Strings.String. -Require Import coqutil.Z.div_mod_to_equations. -Require Import coqutil.Z.Lia. -Require Import coqutil.Word.Interface. -Require Import coqutil.Byte. - -Import BinInt String List.ListNotations ZArith. -Local Open Scope Z_scope. Local Open Scope string_scope. Local Open Scope list_scope. - -Require bedrock2Examples.lightbulb_spec. -Local Notation patience := lightbulb_spec.patience. - -Definition spi_write := func! (b) ~> busy { - busy = $-1; - i = $patience; while i { i = i - $1; - io! busy = MMIOREAD($0x10024048); - if !(busy >> $31) { i = i^i } - }; - if !(busy >> $31) { - output! MMIOWRITE($0x10024048, b); - busy = (busy ^ busy) - } - }. - -Definition spi_read := func! () ~> (b, busy) { - busy = $-1; - b = $0x5a; - i = $patience; while i { i = i - $1; - io! busy = MMIOREAD($0x1002404c); - if !(busy >> $31) { - b = busy & $0xff; - i = i^i; - busy = i - } - } - }. - -Definition spi_xchg := func! (b) ~> (b, busy) { - unpack! busy = spi_write(b); - require !busy; - unpack! b, busy = spi_read() - }. - -Require Import bedrock2.MetricProgramLogic. -Require Import bedrock2.MetricWeakestPreconditionProperties. -Require Import bedrock2.FE310CSemantics bedrock2.MetricSemantics. -Require Import Coq.Lists.List. Import ListNotations. -Require Import bedrock2.TracePredicate. Import TracePredicateNotations. -Require Import bedrock2.ZnWords. - -Import coqutil.Map.Interface. -Import ReversedListNotations. - -Section WithParameters. - Context {word: word.word 32} {mem: map.map word Byte.byte}. - Context {word_ok: word.ok word} {mem_ok: map.ok mem}. - - Definition mmio_event_abstraction_relation - (h : lightbulb_spec.OP word) - (l : mem * string * list word * (mem * list word)) := - Logic.or - (exists a v, h = ("st", a, v) /\ l = (map.empty, "MMIOWRITE", [a; v], (map.empty, []))) - (exists a v, h = ("ld", a, v) /\ l = (map.empty, "MMIOREAD", [a], (map.empty, [v]))). - Definition mmio_trace_abstraction_relation := - List.Forall2 mmio_event_abstraction_relation. - Definition only_mmio_satisfying P t := - exists mmios, mmio_trace_abstraction_relation mmios t /\ P mmios. - - Global Instance spec_of_spi_write : spec_of "spi_write" := fun functions => forall t m b mc, - word.unsigned b < 2 ^ 8 -> - MetricWeakestPrecondition.call functions "spi_write" t m [b] mc (fun T M RETS MC => - M = m /\ exists iol, T = t ;++ iol /\ exists ioh, mmio_trace_abstraction_relation ioh iol /\ exists err, RETS = [err] /\ Logic.or - (((word.unsigned err <> 0) /\ lightbulb_spec.spi_write_full _ ^* ioh /\ Z.of_nat (length ioh) = patience)) - (word.unsigned err = 0 /\ lightbulb_spec.spi_write word (byte.of_Z (word.unsigned b)) ioh)). - - Global Instance spec_of_spi_read : spec_of "spi_read" := fun functions => forall t m mc, - MetricWeakestPrecondition.call functions "spi_read" t m [] mc (fun T M RETS MC => - M = m /\ exists iol, T = t ;++ iol /\ exists ioh, mmio_trace_abstraction_relation ioh iol /\ exists (b: byte) (err : word), RETS = [word.of_Z (byte.unsigned b); err] /\ Logic.or - (word.unsigned err <> 0 /\ lightbulb_spec.spi_read_empty _ ^* ioh /\ Z.of_nat (length ioh) = patience) - (word.unsigned err = 0 /\ lightbulb_spec.spi_read word b ioh)). - - Lemma nonzero_because_high_bit_set (x : word) (H : word.unsigned (word.sru x (word.of_Z 31)) <> 0) - : word.unsigned x <> 0. - Proof. ZnWords. Qed. - - Add Ring wring : (Properties.word.ring_theory (word := word)) - (preprocess [autorewrite with rew_word_morphism], - morphism (Properties.word.ring_morph (word := word)), - constants [Properties.word_cst]). - - Import coqutil.Tactics.letexists. - Import MetricLoops. - Lemma spi_write_ok : program_logic_goal_for_function! spi_write. - Proof. - repeat straightline. - rename H into Hb. - - (* WHY do theese parentheses matter? *) - refine ((atleastonce ["b"; "busy"; "i"] (fun v T M B BUSY I MC => - b = B /\ v = word.unsigned I /\ word.unsigned I <> 0 /\ M = m /\ - exists tl, T = tl++t /\ - exists th, mmio_trace_abstraction_relation th tl /\ - lightbulb_spec.spi_write_full _ ^* th /\ - Z.of_nat (length th) + word.unsigned I = patience - )) _ _ _ _ _ _ _); - cbn [reconstruct map.putmany_of_list HList.tuple.to_list - HList.hlist.foralls HList.tuple.foralls - HList.hlist.existss HList.tuple.existss - HList.hlist.apply HList.tuple.apply - HList.hlist - List.repeat Datatypes.length - HList.polymorphic_list.repeat HList.polymorphic_list.length - PrimitivePair.pair._1 PrimitivePair.pair._2] in *. - { repeat straightline. } - { eapply (Z.lt_wf 0). } - { eexists; eexists; split; repeat straightline. - 2: { - (* TODO why do i need to do all this manually now *) - repeat split; eauto. - exists []; split; eauto. - exists []; repeat split; eauto. - - unfold mmio_trace_abstraction_relation; eauto. - - apply kleene_empty. - - rewrite word.unsigned_of_Z; reflexivity. - } - exfalso. ZnWords. } - - { admit. } - (*{ repeat (split; trivial; []).*) - (* subst i. rewrite word.unsigned_of_Z.*) - (* split.*) - (* { discriminate. }*) - (* split; trivial.*) - (* eexists; split.*) - (* { rewrite app_nil_l; trivial. }*) - (* eexists; split.*) - (* { constructor. }*) - (* split.*) - (* { constructor. }*) - (* exact eq_refl. }*) - - repeat straightline. - eapply MetricWeakestPreconditionProperties.interact_nomem; repeat straightline. - letexists; split; [exact eq_refl|]; split; [split; trivial|]. - { - cbv [isMMIOAddr addr]. - ZnWords. } - repeat straightline. - (* The hnf inside this letexists used to substitute - - c := cmd.cond (expr.op bopname.sru "busy" 31) cmd.skip - (cmd.set "i" (expr.op bopname.xor "i" "i")) : cmd.cmd - - and also unfolded WeakestPrecondition.cmd of c into - - WeakestPrecondition.dexpr m l0 cond letboundEvar /\ - ThenCorrectness /\ - ElseCorrectness - *) - letexists. letexists. split. - { repeat straightline. } - split; intros. - { (* CASE if-condition was true (word.unsigned v0 <> 0), i.e. NOP, loop exit depends on whether timeout *) - repeat straightline. (* <-- does split on a postcondition of the form - (word.unsigned br <> 0 -> loop invariant still holds) /\ - (word.unsigned br = 0 -> code after loop is fine) - which corresponds to case distinction over whether loop was exited *) - - (* fails to generate any cases now *) - admit. - - (*{ (* SUBCASE loop condition was true (do loop again) *)*) - (* eexists; split.*) - (* { repeat (split; trivial; []). subst t0.*) - (* eexists (_ ;++ cons _ nil); split; [exact eq_refl|].*) - (* eexists; split.*) - (* { refine (List.Forall2_app _ _); try eassumption.*) - (* econstructor; [|constructor].*) - (* right; eexists _, _; repeat split. }*) - (* split.*) - (* { eapply kleene_app; eauto.*) - (* refine (kleene_step _ _ nil _ (kleene_empty _)).*) - (* repeat econstructor.*) - (* ZnWords. }*) - (* { ZnWordsL. } }*) - (* { ZnWords. } }*) - (*{ (* SUBCASE loop condition was false (exit loop because of timeout *)*) - (* letexists; split; [solve[repeat straightline]|split]; repeat straightline; try contradiction.*) - (* subst t0.*) - (* eexists (_ ;++ cons _ nil); split.*) - (* { rewrite <-app_assoc; cbn [app]; f_equal. }*) - (* eexists. split.*) - (* { eapply Forall2_app; eauto.*) - (* constructor; [|constructor].*) - (* right; eauto. }*) - (* eexists. split; trivial.*) - (* { left; repeat split; eauto using nonzero_because_high_bit_set.*) - (* { (* copied from above -- trace element for "fifo full" *)*) - (* eapply kleene_app; eauto.*) - (* refine (kleene_step _ _ nil _ (kleene_empty _)).*) - (* repeat econstructor.*) - (* ZnWords. }*) - (* { ZnWordsL. } } }*) - } - - (* CASE if-condition was false (word.unsigned v0 = 0), i.e. we'll set i=i^i and exit loop *) - repeat straightline. - (* fails to generate any cases now *) - admit. - - (*{ subst i.*) - (* rewrite Properties.word.unsigned_xor_nowrap in *; rewrite Z.lxor_nilpotent in *; contradiction. }*) - (*(* evaluate condition then split if *) letexists; split; [solve[repeat straightline]|split].*) - (*1:contradiction.*) - (*repeat straightline.*) - (*eapply WeakestPreconditionProperties.interact_nomem; repeat straightline.*) - (*letexists; letexists; split; [exact eq_refl|]; split; [split; trivial|].*) - (*{ cbv [isMMIOAddr]. ZnWords. }*) - (*repeat straightline.*) - (*subst t0.*) - (*eexists (_ ;++ cons _ (cons _ nil)). split.*) - (*{ rewrite <-app_assoc. cbn [app]. f_equal. }*) - (*eexists. split.*) - (*{ eapply List.Forall2_app; eauto.*) - (* { constructor.*) - (* { left. eexists _, _; repeat split. }*) - (* { right; [|constructor].*) - (* right; eexists _, _; repeat split. } } }*) - (*eexists; split; trivial.*) - (*right.*) - (*subst busy.*) - (*split.*) - (*{ f_equal. rewrite Properties.word.unsigned_xor_nowrap; rewrite Z.lxor_nilpotent; reflexivity. }*) - (*cbv [lightbulb_spec.spi_write].*) - (*eexists _, _; split; eauto; []; split; eauto.*) - (*eexists (cons _ nil), (cons _ nil); split; cbn [app]; eauto.*) - (*split; repeat econstructor.*) - (*{ ZnWords. }*) - (*{ cbv [lightbulb_spec.spi_write_enqueue one].*) - (* repeat f_equal.*) - (* eapply Properties.word.unsigned_inj.*) - (* rewrite byte.unsigned_of_Z; cbv [byte.wrap]; rewrite Z.mod_small; ZnWords. }*) - Admitted. - - Local Ltac split_if := - lazymatch goal with - |- WeakestPrecondition.cmd _ ?c _ _ _ ?post => - let c := eval hnf in c in - lazymatch c with - | cmd.cond _ _ _ => letexists; split; [solve[repeat straightline]|split] - end - end. - - Lemma spi_read_ok : program_logic_goal_for_function! spi_read. - repeat straightline. - refine ((atleastonce ["b"; "busy"; "i"] (fun v T M B BUSY I => - v = word.unsigned I /\ word.unsigned I <> 0 /\ M = m /\ - B = word.of_Z (byte.unsigned (byte.of_Z (word.unsigned B))) /\ - exists tl, T = tl++t /\ - exists th, mmio_trace_abstraction_relation th tl /\ - lightbulb_spec.spi_read_empty _ ^* th /\ - Z.of_nat (length th) + word.unsigned I = patience - )) - _ _ _ _ _ _ _); - cbn [reconstruct map.putmany_of_list HList.tuple.to_list - HList.hlist.foralls HList.tuple.foralls - HList.hlist.existss HList.tuple.existss - HList.hlist.apply HList.tuple.apply - HList.hlist - List.repeat Datatypes.length - HList.polymorphic_list.repeat HList.polymorphic_list.length - PrimitivePair.pair._1 PrimitivePair.pair._2] in *; repeat straightline. - { exact (Z.lt_wf 0). } - { exfalso. ZnWords. } - { subst i. rewrite word.unsigned_of_Z. - split; [inversion 1|]. - split; trivial. - subst b; rewrite byte.unsigned_of_Z; cbv [byte.wrap]; - rewrite Z.mod_small; rewrite word.unsigned_of_Z. - 2: { cbv. split; congruence. } - split; trivial. - eexists nil; split; trivial. - eexists nil; split; try split; solve [constructor]. } - { eapply WeakestPreconditionProperties.interact_nomem; repeat straightline. - letexists; split; [exact eq_refl|]; split; [split; trivial|]. - { cbv [isMMIOAddr]. ZnWords. } - repeat ((split; trivial; []) || straightline || split_if). - { - letexists. split; split. - { subst v'; exact eq_refl. } - { split; trivial. - split; trivial. - split; trivial. - eexists (x2 ;++ cons _ nil); split; cbn [app]; eauto. - eexists. split. - { econstructor; try eassumption; right; eauto. } - split. - { - refine (kleene_app _ (cons _ nil) _ x3 _); eauto. - refine (kleene_step _ (cons _ nil) nil _ (kleene_empty _)). - eexists; split. - { exact eq_refl. } - { ZnWords. } } - { ZnWordsL. } } - { ZnWords. } - { ZnWords. } } - { eexists (x2 ;++ cons _ nil); split; cbn [app]; eauto. - eexists. split. - { econstructor; try eassumption; right; eauto. } - eexists (byte.of_Z (word.unsigned x)), _; split. - { f_equal. eassumption. } - left; repeat split; eauto using nonzero_because_high_bit_set. - { refine (kleene_app _ (cons _ nil) _ x3 _); eauto. - refine (kleene_step _ (cons _ nil) nil _ (kleene_empty _)). - eexists; split. - { exact eq_refl. } - { ZnWords. } } - { ZnWordsL. } } - { repeat straightline. - repeat letexists; split. - 1: split. - { repeat straightline. } - 2: { - subst v'. - subst v. - subst i. - rewrite Properties.word.unsigned_xor_nowrap, Z.lxor_nilpotent. - ZnWords. } - repeat straightline. - repeat (split; trivial; []). - split. - { subst b. - (* automatable: multi-word bitwise *) - change (255) with (Z.ones 8). - pose proof Properties.word.unsigned_range v0. - eapply Properties.word.unsigned_inj. - repeat ( - cbv [byte.wrap word.wrap]; - rewrite ?byte.unsigned_of_Z, ?word.unsigned_of_Z, ?Properties.word.unsigned_and_nowrap, - ?Z.land_ones, ?Z.mod_mod, ?Z.mod_small - by blia; - change (Z.ones 8 mod 2 ^ 32) with (Z.ones 8)). - symmetry; eapply Z.mod_small. - pose proof Z.mod_pos_bound (word.unsigned v0) (2^8) eq_refl. - clear. Z.div_mod_to_equations. blia. } - { (* copy-paste from above, trace manipulation *) - eexists (x2 ;++ cons _ nil); split; cbn [app]; eauto. - eexists. split. - { econstructor; try eassumption; right; eauto. } - subst i. - rewrite Properties.word.unsigned_xor_nowrap, Z.lxor_nilpotent in H1; contradiction. } } - { (* copy-paste from above, trace manipulation *) - eexists (x2 ;++ cons _ nil); split; cbn [app]; eauto. - eexists. split. - { econstructor; try eassumption; right; eauto. } - eexists (byte.of_Z (word.unsigned b)), _; split. - { subst b; f_equal. - (* tag:bitwise *) - (* automatable: multi-word bitwise *) - change (255) with (Z.ones 8). - pose proof Properties.word.unsigned_range v0. - eapply Properties.word.unsigned_inj. - repeat ( - cbv [byte.wrap word.wrap]; - rewrite ?byte.unsigned_of_Z, ?word.unsigned_of_Z, ?Properties.word.unsigned_and_nowrap, - ?Z.land_ones, ?Z.mod_mod, ?Z.mod_small - by blia; - change (Z.ones 8 mod 2 ^ 32) with (Z.ones 8)). - symmetry; eapply Z.mod_small. - pose proof Z.mod_pos_bound (word.unsigned v0) (2^8) eq_refl. - clear. Z.div_mod_to_equations. blia. } - (* tag:symex *) - { right; split. - { subst_words. rewrite Properties.word.unsigned_xor_nowrap, Z.lxor_nilpotent; exact eq_refl. } - eexists x3, (cons _ nil); split; cbn [app]; eauto. - split; eauto. - eexists; split; cbv [one]; trivial. - split. - (* tag:bitwise *) - { ZnWords. } - subst b. - (* automatable: multi-word bitwise *) - change (255) with (Z.ones 8). - pose proof Properties.word.unsigned_range v0. - eapply byte.unsigned_inj. - repeat ( - cbv [byte.wrap word.wrap]; - rewrite ?byte.unsigned_of_Z, ?word.unsigned_of_Z, ?Properties.word.unsigned_and_nowrap, - ?Z.land_ones, ?Z.mod_mod, ?Z.mod_small - by blia; - change (Z.ones 8 mod 2 ^ 32) with (Z.ones 8)). - trivial. } } } - Qed. - - Global Instance spec_of_spi_xchg : spec_of "spi_xchg" := fun functions => forall t m b_out, - word.unsigned b_out < 2 ^ 8 -> - WeakestPrecondition.call functions "spi_xchg" t m [b_out] (fun T M RETS => - M = m /\ exists iol, T = t ;++ iol /\ exists ioh, mmio_trace_abstraction_relation ioh iol /\ exists (b_in:byte) (err : word), RETS = [word.of_Z (byte.unsigned b_in); err] /\ Logic.or - (word.unsigned err <> 0 /\ (any +++ lightbulb_spec.spi_timeout _) ioh) - (word.unsigned err = 0 /\ lightbulb_spec.spi_xchg word (byte.of_Z (word.unsigned b_out)) b_in ioh)). - - Lemma spi_xchg_ok : program_logic_goal_for_function! spi_xchg. - Proof. - repeat ( - match goal with - | |- ?F ?a ?b ?c => - match F with WeakestPrecondition.get => idtac end; - let f := (eval cbv beta delta [WeakestPrecondition.get] in F) in - change (f a b c); cbv beta - | H : _ /\ _ \/ ?Y /\ _, G : not ?X |- _ => - constr_eq X Y; let Z := fresh in destruct H as [|[Z ?]]; [|case (G Z)] - | H : not ?Y /\ _ \/ _ /\ _, G : ?X |- _ => - constr_eq X Y; let Z := fresh in destruct H as [[Z ?]|]; [case (Z G)|] - end || - - straightline || straightline_call || split_if || refine (conj _ _) || eauto). - - { eexists. split. - { exact eq_refl. } - eexists. split. - { eauto. } - eexists. eexists. split. - { repeat f_equal. - instantiate (1 := byte.of_Z (word.unsigned b_out)). - (* automatable: multi-word bitwise *) - change (255) with (Z.ones 8). - pose proof Properties.word.unsigned_range b_out. - eapply Properties.word.unsigned_inj; - repeat ( - cbv [word.wrap byte.wrap]; - rewrite ?byte.unsigned_of_Z, ?word.unsigned_of_Z, ?Properties.word.unsigned_and_nowrap, ?Z.land_ones, ?Z.mod_mod, ?Z.mod_small by blia; - change (Z.ones 8 mod 2 ^ 32) with (Z.ones 8)); - rewrite ?Z.mod_small; rewrite ?Z.mod_small; trivial; blia. } - left; split; eauto. - eexists nil, x0; repeat split; cbv [any choice lightbulb_spec.spi_timeout]; eauto. - rewrite app_nil_r; trivial. } - - { destruct H10; intuition eauto. - { eexists. split. - { subst a0. subst a. - rewrite List.app_assoc; trivial. } - eexists. split. - { eapply Forall2_app; eauto. } - eexists _, _; split. - { subst v; trivial. } - left; split; eauto. - eapply concat_app; cbv [any choice lightbulb_spec.spi_timeout]; eauto. } - eexists. - subst a0. - subst a. - split. - { rewrite List.app_assoc; trivial. } - eexists. - split. - { eapply Forall2_app; eauto. } - eexists _, _; split. - { subst v. eauto. } - right. split; eauto. - cbv [lightbulb_spec.spi_xchg]. - - assert (Trace__concat_app : forall T (P Q:list T->Prop) x y, P x -> Q y -> (P +++ Q) (y ++ x)). { - cbv [concat]; eauto. } - - eauto using Trace__concat_app. } - Qed. -End WithParameters. diff --git a/bedrock2/src/bedrock2Examples/metric_ipow.v b/bedrock2/src/bedrock2Examples/metric_ipow.v index 2ba342bdf..2c931529b 100644 --- a/bedrock2/src/bedrock2Examples/metric_ipow.v +++ b/bedrock2/src/bedrock2Examples/metric_ipow.v @@ -100,36 +100,9 @@ Proof. blia. Qed. -Lemma metriclit : forall a b c d a' b' c' d' mc, - metricsAdd (mkMetricLog a b c d) (metricsAdd (mkMetricLog a' b' c' d') mc) = - metricsAdd (mkMetricLog (a+a') (b+b') (c+c') (d+d')) mc. -Proof. intros. rewrite MetricArith.add_assoc. reflexivity. Qed. - Ltac s := unfold initCost, iterCost, endCost in *; - repeat ( - let H := match goal with - | H : context[cost_interact] |- _ => H - | H : context[cost_call] |- _ => H - | H : context[cost_load] |- _ => H - | H : context[cost_store] |- _ => H - | H : context[cost_inlinetable] |- _ => H - | H : context[cost_stackalloc] |- _ => H - | H : context[cost_lit] |- _ => H - | H : context[cost_op] |- _ => H - | H : context[cost_set] |- _ => H - | H : context[cost_if] |- _ => H - | H : context[cost_loop_true] |- _ => H - | H : context[cost_loop_false] |- _ => H - end in - let t := type of H in - let t' := eval cbv [cost_interact cost_call cost_load cost_store - cost_inlinetable cost_stackalloc cost_lit cost_op cost_set - cost_if cost_loop_true cost_loop_false] in t in - replace t with t' in H by (symmetry; reflexivity) - ); cost_unfold; cbn in *; - repeat rewrite metriclit in *; solve_MetricLog. Lemma ipow_ok : program_logic_goal_for_function! ipow. From 9b9f221d8242c8c5fc68929b3a1fc4e172851ccc Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Fri, 9 Aug 2024 04:52:57 -0400 Subject: [PATCH 58/62] remove unfinished loop lemmas to be restored later --- bedrock2/src/bedrock2/MetricLoops.v | 294 ---------------------------- 1 file changed, 294 deletions(-) diff --git a/bedrock2/src/bedrock2/MetricLoops.v b/bedrock2/src/bedrock2/MetricLoops.v index 52210674e..50855548f 100644 --- a/bedrock2/src/bedrock2/MetricLoops.v +++ b/bedrock2/src/bedrock2/MetricLoops.v @@ -358,300 +358,6 @@ Section Loops. constructor. intros [y|] pf; eauto. constructor. intros [] []. Qed. - - Lemma atleastonce_localsmap - {e c t} {m : mem} {l mc} {post : _->_->_->_-> Prop} - {measure : Type} (invariant:_->_->_->_->_->Prop) lt - (Hwf : well_founded lt) - (Henter : exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) - /\ (word.unsigned br = 0%Z -> post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc)) - /\ (word.unsigned br <> 0 -> exists v', invariant v' t m l (cost_loop_true isRegStr UNK (Some UNK) brmc)) - ) - (v0 : measure) (Hpre : invariant v0 t m l mc) - (Hbody : forall v t m l mc, invariant v t m l mc -> - cmd call c t m l mc (fun t m l mc => - exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) - /\ (word.unsigned br <> 0 -> exists v', invariant v' t m l (cost_loop_true isRegStr UNK (Some UNK) brmc) /\ lt v' v) - /\ (word.unsigned br = 0 -> post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc)))) - : cmd call (cmd.while e c) t m l mc post. - Proof. - eapply wp_while. - eexists (option measure), (with_bottom lt), (fun ov t m l mc => - exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) - /\ ((word.unsigned br <> 0 -> exists v, ov = Some v /\ invariant v t m l (cost_loop_true isRegStr UNK (Some UNK) brmc)) - /\ (word.unsigned br = 0 -> ov = None /\ post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc)))). - split; auto using well_founded_with_bottom; []. split. - { destruct Henter as [br [brmc [He [Henter0 Henterm]]]]. - destruct (BinInt.Z.eq_dec (word.unsigned br) 0). - { exists None, br, brmc; split; trivial. - split; intros; try contradiction; split; eauto. } - { assert (Hnonzero := n). - apply Henterm in Hnonzero. - destruct Hnonzero as [v Hinv]. - exists (Some v), br, brmc. - split; trivial; []; split; try contradiction. - exists v; split; trivial. } } - intros vi ti mi li mci (br&brmc&Ebr&Hcontinue&Hexit). - eexists; eexists; split; [eassumption|]; split. - { admit. } - (* - { intros Hc; destruct (Hcontinue Hc) as (v&?&Hinv); subst. - eapply Proper_cmd. - 2: { - eapply Hbody. - } - eapply Proper_cmd; [ |eapply Hbody; eassumption]. - intros t' m' l' mc' (br'&brmc'&Ebr'&Hinv'&Hpost'). - destruct (BinInt.Z.eq_dec (word.unsigned br') 0). - { exists None; split; try constructor. - exists br'; split; trivial; []. - split; intros; try contradiction. - split; eauto. } - { destruct (Hinv' ltac:(trivial)) as (v'&inv'<v'v). - exists (Some v'); split; trivial. (* NOTE: this [trivial] simpl-reduces [with_bottom] *) - exists br'; split; trivial. - split; intros; try contradiction. - eexists; split; eauto. } } - *) - - eapply Hexit. - (*Qed.*) - Admitted. - - Lemma atleastonce - {e c t l mc} {m : mem} - (variables : list String.string) - {localstuple : tuple word (length variables)} - {Pl : enforce variables localstuple l} - {measure : Type} (invariant:_->_->_-> ufunc word (length variables) (MetricLog -> Prop)) - lt (Hwf : well_founded lt) - {post : _->_->_->_->Prop} - (Henter : exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ (word.unsigned br = 0%Z -> post t m l (cost_loop_false isRegStr UNK (Some UNK) mc')) /\ - (word.unsigned br <> 0 -> exists v', tuple.apply (invariant v' t m) localstuple (cost_loop_true isRegStr UNK (Some UNK) mc')) - - ) - (v0 : measure) (Hpre : tuple.apply (invariant v0 t m) localstuple mc) - (Hbody : forall v t m mc, tuple.foralls (fun localstuple => - tuple.apply (invariant v t m) localstuple mc -> - cmd call c t m (reconstruct variables localstuple) mc (fun t m l mc => - exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ - (word.unsigned br <> 0 -> Markers.unique (Markers.left (tuple.existss (fun localstuple => enforce variables localstuple l /\ Markers.right (Markers.unique (exists v', tuple.apply (invariant v' t m) localstuple (cost_loop_true isRegStr UNK (Some UNK) mc') /\ lt v' v)))))) /\ - (word.unsigned br = 0 -> post t m l (cost_loop_false isRegStr UNK (Some UNK) mc'))))) - : cmd call (cmd.while e c) t m l mc post. - Proof. - eapply (atleastonce_localsmap (fun v t m l mc => exists localstuple, Logic.and (enforce variables localstuple l) (tuple.apply (invariant v t m) localstuple mc))); eauto. - { - destruct Henter as [br0 [mc0 [Henter [Henter0 Henternonzero]]]]. - exists br0. exists mc0. repeat (split; eauto). - intro Hnonzero. apply Henternonzero in Hnonzero. destruct Hnonzero as [v Hnonzero]. - exists v. exists localstuple. repeat (split; eauto). - } - { - intros vi ti mi li mci (?&X&Y). - specialize (Hbody vi ti mi). - eapply hlist.foralls_forall in Hbody. - specialize (Hbody Y). - rewrite <-(reconstruct_enforce _ _ _ X) in Hbody. - eapply Proper_cmd; [|eapply Hbody]. - intros t' m' l' mc' (?&?&?&HH). - eexists; eexists; split; eauto. destruct HH as [HH1 HH2]. - split; intros; eauto. - specialize (HH1 ltac:(eauto)). - eapply hlist.existss_exists in HH1; destruct HH1 as (?&?&?&?&?). - eexists; split; eauto. - } - Qed. - - (* - Lemma tailrec_earlyout_localsmap - {e c t} {m : mem} {l mc} {post : _-> _->_->_-> Prop} - {measure : Type} (spec:_->_->_->_->_->(Prop*(_->_->_->_-> Prop))) lt - (Hwf : well_founded lt) - (v0 : measure) (P0 := (spec v0 t m l mc).(1)) (Hpre : P0) - (Q0 := (spec v0 t m l mc).(2)) - (Hbody : forall v t m l mc, - let S := spec v t m l mc in let (P, Q) := S in - P -> - exists br mc', expr m l e mc (eq (Datatypes.pair br mc')) /\ - (word.unsigned br <> 0%Z -> cmd call c t m l mc' - (fun t' m' l' mc'' => - (exists br mc', expr m' l' e mc'' (eq (Datatypes.pair br mc')) /\ word.unsigned br = 0 /\ Q t' m' l' mc') \/ - exists v', let S' := spec v' t' m' l' in let '(P', Q') := S' in - P' /\ - lt v' v /\ - forall T M L MC, Q' T M L -> Q T M L)) /\ - (word.unsigned br = 0%Z -> Q t m l)) - (Hpost : forall t m l, Q0 t m l -> post t m l) - : cmd call (cmd.while e c) t m l post. - Proof. - eapply wp_while. - eexists (option measure), (with_bottom lt), (fun v t m l => - match v with - | None => exists br, expr m l e (eq br) /\ word.unsigned br = 0 /\ Q0 t m l - | Some v => - let S := spec v t m l in let '(P, Q) := S in - P /\ forall T M L, Q T M L -> Q0 T M L - end). - split; auto using well_founded_with_bottom; []; cbv [Markers.split] in *. - split. - { exists (Some v0); eauto. } - intros [vi|] ti mi li inv_i; [destruct inv_i as (?&Qi)|destruct inv_i as (br&Hebr&Hbr0&HQ)]. - { destruct (Hbody _ _ _ _ ltac:(eassumption)) as (br&?&X); exists br; split; [assumption|]. - destruct X as (Htrue&Hfalse). split; intros Hbr; - [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; eauto. - eapply Proper_cmd; [ |eapply Hpc]. - intros tj mj lj [(br'&Hbr'&Hz&HQ)|(vj&dP&?&dQ)]; - [exists None | exists (Some vj)]; cbn [with_bottom]; eauto 9. } - repeat esplit || eauto || intros; contradiction. - Qed. - - Lemma tailrec_earlyout - {e c t localsmap} {m : mem} - (ghosttypes : polymorphic_list.list Type) - (variables : list String.string) - {l0 : tuple word (length variables)} - {Pl : enforce variables l0 localsmap} - {post : _->_->_-> Prop} - {measure : Type} (spec:_->HList.arrows ghosttypes (_->_->ufunc word (length variables) (Prop*(_->_->ufunc word (length variables) Prop)))) lt - (Hwf : well_founded lt) - (v0 : measure) - : hlist.foralls (fun (g0 : hlist ghosttypes) => forall - (Hpre : (tuple.apply (hlist.apply (spec v0) g0 t m) l0).(1)) - (Hbody : forall v, hlist.foralls (fun g => forall t m, tuple.foralls (fun l => - @dlet _ (fun _ => Prop) (reconstruct variables l) (fun localsmap : locals => - match tuple.apply (hlist.apply (spec v) g t m) l with S_ => - S_.(1) -> - Markers.unique (Markers.left (exists br, expr m localsmap e (eq br) /\ Markers.right ( - (word.unsigned br <> 0%Z -> cmd call c t m localsmap - (fun t' m' localsmap' => - Markers.unique (Markers.left (hlist.existss (fun l' => enforce variables l' localsmap' /\ Markers.right ( - Markers.unique (Markers.left (exists br, expr m' localsmap' e (eq br) /\ Markers.right ( word.unsigned br = 0 /\ tuple.apply (S_.(2) t' m') l') ) ) \/ - Markers.unique (Markers.left (hlist.existss (fun g' => exists v', - match tuple.apply (hlist.apply (spec v') g' t' m') l' with S' => - S'.(1) /\ Markers.right ( - lt v' v /\ - forall T M, hlist.foralls (fun L => tuple.apply (S'.(2) T M) L -> tuple.apply (S_.(2) T M) L)) end))))))))) /\ - (word.unsigned br = 0%Z -> tuple.apply (S_.(2) t m) l))))end)))) - (Hpost : match (tuple.apply (hlist.apply (spec v0) g0 t m) l0).(2) with Q0 => forall t m, hlist.foralls (fun l => tuple.apply (Q0 t m) l -> post t m (reconstruct variables l))end) - , cmd call (cmd.while e c) t m localsmap post ). - Proof. - eapply hlist_forall_foralls; intros g0 **. - eapply wp_while. - eexists (option measure), (with_bottom lt), (fun vi ti mi localsmapi => - exists li, localsmapi = reconstruct variables li /\ - match vi with None => exists br, expr mi localsmapi e (eq br) /\ word.unsigned br = 0 /\ tuple.apply ((tuple.apply (hlist.apply (spec v0) g0 t m) l0).(2) ti mi) li | Some vi => - exists gi, match tuple.apply (hlist.apply (spec vi) gi ti mi) li with S_ => - S_.(1) /\ forall T M L, tuple.apply (S_.(2) T M) L -> - tuple.apply ((tuple.apply (hlist.apply (spec v0) g0 t m) l0).(2) T M) L end end). - cbv [Markers.unique Markers.split Markers.left Markers.right] in *. - split; eauto using well_founded_with_bottom. - split. { exists (Some v0), l0. split. 1: eapply reconstruct_enforce; eassumption. exists g0; split; eauto. } - intros [vi|] ti mi lmapi. - 2: { intros (ld&Hd&br&Hbr&Hz&Hdone). - eexists; split; eauto. - split; intros; try contradiction. - subst; eapply (hlist.foralls_forall (Hpost ti mi) _ Hdone). } - intros (?&?&gi&?&Qi); subst. - destruct (hlist.foralls_forall (hlist.foralls_forall (Hbody vi) gi ti mi) _ ltac:(eassumption)) as (br&?&X). - exists br; split; [assumption|]. destruct X as (Htrue&Hfalse). split; intros Hbr; - [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. - { eapply Proper_cmd; [ |eapply Hpc]. - intros tj mj lmapj Hlj; eapply hlist.existss_exists in Hlj. - destruct Hlj as (lj&Elj&HE); eapply reconstruct_enforce in Elj; subst lmapj. - destruct HE as [(br'&Hevalr'&Hz'&Hdone)|HE]. - { exists None; cbn. eauto 9. } - { eapply hlist.existss_exists in HE. destruct HE as (l&?&?&?&HR). - pose proof fun T M => hlist.foralls_forall (HR T M); clear HR. - eexists (Some _); eauto 9. } } - { pose proof fun t m => hlist.foralls_forall (Hpost t m); clear Hpost; eauto. } - Qed. - -<<<<<<< HEAD - - Lemma atleastonce - {e c t l} {m : mem} - (variables : list String.string) - {localstuple : tuple word (length variables)} - {Pl : enforce variables localstuple l} - {measure : Type} (invariant:_->_->_->ufunc word (length variables) Prop) - lt (Hwf : well_founded lt) - {post : _->_->_-> Prop} - (Henter : exists br, expr m l e (eq br) /\ (word.unsigned br = 0%Z -> post t m l)) - (v0 : measure) (Hpre : tuple.apply (invariant v0 t m) localstuple) - (Hbody : forall v t m, tuple.foralls (fun localstuple => - tuple.apply (invariant v t m) localstuple -> - cmd call c t m (reconstruct variables localstuple) (fun t m l => - exists br, expr m l e (eq br) /\ - (word.unsigned br <> 0 -> Markers.unique (Markers.left (tuple.existss (fun localstuple => enforce variables localstuple l /\ Markers.right (Markers.unique (exists v', tuple.apply (invariant v' t m) localstuple /\ lt v' v)))))) /\ - (word.unsigned br = 0 -> post t m l)))) - : cmd call (cmd.while e c) t m l post. - Proof. - eapply (atleastonce_localsmap (fun v t m l => exists localstuple, Logic.and (enforce variables localstuple l) (tuple.apply (invariant v t m) localstuple))); eauto. - intros vi ti mi li (?&X&Y). - specialize (Hbody vi ti mi). - eapply hlist.foralls_forall in Hbody. - specialize (Hbody Y). - rewrite <-(reconstruct_enforce _ _ _ X) in Hbody. - eapply Proper_cmd; [ |eapply Hbody]. - intros t' m' l' (?&?&HH&?). - eexists; split; eauto. - split; intros; eauto. - specialize (HH ltac:(eauto)). - eapply hlist.existss_exists in HH; destruct HH as (?&?&?&?&?). - eexists; split; eauto. - Qed. - -======= ->>>>>>> 94c30cc1 (made rather a mess trying to get andy's metric stuff to work; commiting so I can ask andres for help) - Lemma while_zero_iterations {e c t l} {m : mem} {post : _->_->_-> Prop} - (HCond: expr m l e (eq (word.of_Z 0))) - (HPost: post t m l) - : cmd call (cmd.while e c) t m l post. - Proof. - eapply (while_localsmap (fun n t' m' l' => t' = t /\ m' = m /\ l' = l) (PeanoNat.Nat.lt_wf 0) 0%nat). - 1: unfold split; auto. intros *. intros (? & ? & ?). subst. - eexists. split. 1: exact HCond. - rewrite Properties.word.unsigned_of_Z_0. - split; intros; congruence. - Qed. - - - (* Bedrock-style loop rule *) - Local Open Scope sep_scope. - Local Infix "*" := Separation.sep : type_scope. - Local Infix "==>" := Lift1Prop.impl1. - - Lemma tailrec_sep - e c t (m : mem) l (post : _->_->_-> Prop) - {measure : Type} P Q lt (Hwf : well_founded lt) (v0 : measure) R0 - (Hpre : (P v0 t l * R0) m) - (Hbody : forall v t m l R, (P v t l * R) m -> - exists br, expr m l e (eq br) /\ - (word.unsigned br <> 0%Z -> cmd call c t m l - (fun t' m' l' => exists v' dR, (P v' t' l' * (R * dR)) m' /\ - lt v' v /\ - forall T L, Q v' T L * dR ==> Q v T L)) /\ - (word.unsigned br = 0%Z -> (Q v t l * R) m)) - (Hpost : forall t m l, (Q v0 t l * R0) m -> post t m l) - : cmd call (cmd.while e c) t m l post. - Proof. - eapply wp_while. - eexists measure, lt, (fun v t m l => exists R, (P v t l * R) m /\ - forall T L, Q v T L * R ==> Q v0 T L * R0). - split; [assumption|]. - split. { exists v0, R0. split; [assumption|]. intros. reflexivity. } - intros vi ti mi li (Ri&?&Qi). - destruct (Hbody _ _ _ _ _ ltac:(eassumption)) as (br&?&X); exists br; split; [assumption|]. - destruct X as (Htrue&Hfalse). split; intros Hbr; - [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. - { eapply Proper_cmd; [ |eapply Hpc]. - intros tj mj lj (vj&dR&dP&?&dQ). - exists vj; split; [|assumption]. - exists (Ri * dR); split; [assumption|]. - intros. rewrite (sep_comm _ dR), <-(sep_assoc _ dR), dQ; trivial. } - { eapply Hpost, Qi, Hpc. } - Qed. -*) End Loops. Ltac loop_simpl := From 1d848252798f12014efa80de06cb2e39bb1c184d Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Fri, 9 Aug 2024 04:53:17 -0400 Subject: [PATCH 59/62] slightly nicer cost_unfold --- bedrock2/src/bedrock2/MetricCosts.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/bedrock2/src/bedrock2/MetricCosts.v b/bedrock2/src/bedrock2/MetricCosts.v index 6cde700d7..fcafde638 100644 --- a/bedrock2/src/bedrock2/MetricCosts.v +++ b/bedrock2/src/bedrock2/MetricCosts.v @@ -135,7 +135,7 @@ Ltac cost_unfold := let t' := eval cbv [cost_interact cost_call cost_load cost_store cost_inlinetable cost_stackalloc cost_lit cost_op cost_set cost_if cost_loop_true cost_loop_false] in t in - replace t with t' in H by (symmetry; reflexivity) + replace t with t' in H by (exact (eq_refl t')) ); cbv [cost_interact cost_call cost_load cost_store cost_inlinetable cost_stackalloc cost_lit cost_op cost_set cost_if cost_loop_true From 46e8ccc441e251a9fcc0fae29ec5c2f908e63412 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Fri, 9 Aug 2024 13:33:43 -0400 Subject: [PATCH 60/62] oops forgot to delete this --- compiler/src/compiler/FlatToRiscvMetric.v | 383 ---------------------- 1 file changed, 383 deletions(-) delete mode 100644 compiler/src/compiler/FlatToRiscvMetric.v diff --git a/compiler/src/compiler/FlatToRiscvMetric.v b/compiler/src/compiler/FlatToRiscvMetric.v deleted file mode 100644 index b0724fe6f..000000000 --- a/compiler/src/compiler/FlatToRiscvMetric.v +++ /dev/null @@ -1,383 +0,0 @@ -(* -TODO why did pratap delete this file? - -Require Import riscv.Spec.Primitives. -Require Import riscv.Platform.RiscvMachine. -Require Import riscv.Platform.MetricRiscvMachine. -Require Import riscv.Platform.MetricLogging. -Require Import riscv.Utility.Utility. -Require Import riscv.Utility.runsToNonDet. -Require Import riscv.Utility.InstructionCoercions. -Require Import compiler.util.Common. -Require Import compiler.eqexact. -Require Import coqutil.Tactics.Simp. -Require Import compiler.on_hyp_containing. -Require Import compiler.SeparationLogic. -Require Export coqutil.Word.SimplWordExpr. -Require Import compiler.GoFlatToRiscv. -Require Import compiler.DivisibleBy4. -Require Import compiler.MetricsToRiscv. -Require Import compiler.FlatImp. -Require Import compiler.RiscvWordProperties. -Require Import compiler.FlatToRiscvDef. -Require Import compiler.FlatToRiscvCommon. -Require Import compiler.FlatToRiscvLiterals. -Require Import coqutil.Tactics.fwd. - -Open Scope ilist_scope. - -Local Arguments Z.mul: simpl never. -Local Arguments Z.add: simpl never. -Local Arguments Z.of_nat: simpl never. -Local Arguments Z.modulo : simpl never. -Local Arguments Z.pow: simpl never. -Local Arguments Z.sub: simpl never. - -Section Proofs. - Context {iset: Decode.InstructionSet}. - Context {pos_map: map.map String.string Z}. - Context (compile_ext_call: pos_map -> Z -> Z -> stmt Z -> list Decode.Instruction). - Context {width: Z} {BW: Bitwidth width} {word: word.word width} {word_ok: word.ok word}. - Context {word_riscv_ok: RiscvWordProperties.word.riscv_ok word}. - Context {locals: map.map Z word} {locals_ok: map.ok locals}. - Context {mem: map.map word byte} {mem_ok: map.ok mem}. - Context {env: map.map String.string (list Z * list Z * stmt Z)} {env_ok: map.ok env}. - Context {M: Type -> Type}. - Context {MM: Monads.Monad M}. - Context {RVM: Machine.RiscvProgram M word}. - Context {PRParams: PrimitivesParams M MetricRiscvMachine}. - Context {ext_spec: Semantics.ExtSpec} {ext_spec_ok: Semantics.ext_spec.ok ext_spec}. - Context {PR: MetricPrimitives.MetricPrimitives PRParams}. - Context {BWM: bitwidth_iset width iset}. - - Add Ring wring : (word.ring_theory (word := word)) - (preprocess [autorewrite with rew_word_morphism], - morphism (word.ring_morph (word := word)), - constants [word_cst]). - - Local Notation RiscvMachineL := (MetricRiscvMachine (width := width)). - - Ltac run1done := - apply runsToDone; - simpl_MetricRiscvMachine_get_set; - rewrite ?word.sru_ignores_hibits, - ?word.slu_ignores_hibits, - ?word.srs_ignores_hibits, - ?word.mulhuu_simpl, - ?word.divu0_simpl, - ?word.modu0_simpl; - simpl in *; - eexists; (* finalMH *) - eexists; (* finalMetricsH *) - repeat split; - simpl_word_exprs word_ok; - first - [ solve [eauto] - | solve_MetricLog - | solve_word_eq word_ok - | solve [wcancel_assumption] - | eapply rearrange_footpr_subset; [ eassumption | wwcancel ] - | solve [solve_valid_machine word_ok] - | idtac ]. - - Ltac IH_sidecondition := - simpl_word_exprs word_ok; - try solve - [ reflexivity - | auto - | solve_word_eq word_ok - | simpl; solve_divisibleBy4 - | solve_valid_machine word_ok - | eapply rearrange_footpr_subset; [ eassumption | wwcancel ] - | wcancel_assumption ]. - - Hypothesis no_ext_calls: forall t mGive action argvals outcome, - ext_spec t mGive action argvals outcome -> False. - - Hypothesis stackalloc_always_0: forall x n body t m (l: locals) mc post, - FlatImp.exec map.empty (SStackalloc x n body) t m l mc post -> n = 0. - - Hypothesis sp_always_set: forall l: locals, - map.get l RegisterNames.sp = Some (word.of_Z 42). - - (* not needed any more *) - Definition stmt_not_too_big(s: stmt Z): Prop := True. - - Lemma compile_stmt_correct: - forall (s: stmt Z) t initialMH initialRegsH postH initialMetricsH, - FlatImp.exec map.empty s t initialMH initialRegsH initialMetricsH postH -> - forall R Rexec (initialL: RiscvMachineL) insts pos, - compile_stmt iset compile_ext_call map.empty pos 12345678 s = insts -> - stmt_not_too_big s -> - valid_FlatImp_vars s -> - divisibleBy4 initialL.(getPc) -> - initialL.(getRegs) = initialRegsH -> - subset (footpr (program iset initialL.(getPc) insts * Rexec)%sep) (of_list initialL.(getXAddrs)) -> - (program iset initialL.(getPc) insts * Rexec * eq initialMH * R)%sep initialL.(getMem) -> - initialL.(getLog) = t -> - initialL.(getNextPc) = add initialL.(getPc) (word.of_Z 4) -> - valid_machine initialL -> - runsTo initialL (fun finalL => exists finalMH finalMetricsH, - postH finalL.(getLog) finalMH finalL.(getRegs) finalMetricsH /\ - subset (footpr (program iset initialL.(getPc) insts * Rexec)%sep) - (of_list finalL.(getXAddrs)) /\ - (program iset initialL.(getPc) insts * Rexec * eq finalMH * R)%sep finalL.(getMem) /\ - finalL.(getPc) = add initialL.(getPc) - (word.mul (word.of_Z 4) (word.of_Z (Z.of_nat (length insts)))) /\ - finalL.(getNextPc) = word.add finalL.(getPc) (word.of_Z 4) /\ - (finalL.(getMetrics) - initialL.(getMetrics) <= - lowerMetrics (finalMetricsH - initialMetricsH))%metricsL /\ - valid_machine finalL). - Proof. - induction 1; intros; - repeat match goal with - | m: _ |- _ => destruct_RiscvMachine m; simpl_MetricRiscvMachine_get_set - end; - simpl in *; - subst; - simp. - - - (* SInteract *) - exfalso. eapply no_ext_calls. eassumption. - - - (* SCall *) - lazymatch goal with - | A: map.get map.empty _ = Some _ |- _ => - exfalso; simpl in *; - rewrite map.get_empty in A - end. - discriminate. - - - (* SLoad *) - unfold Memory.load, Memory.load_Z in *. simp. subst_load_bytes_for_eq. - run1det. run1done. - - - (* SStore *) - simpl_MetricRiscvMachine_get_set. - assert ((eq m * (program iset initialL_pc [[compile_store iset sz a v o]] * Rexec * R))%sep - initialL_mem) as A by ecancel_assumption. - match goal with - | H: _ |- _ => pose proof (store_bytes_frame H A) as P; move H at bottom; - unfold Memory.store, Memory.store_Z, Memory.store_bytes in H - end. - destruct P as (finalML & P1 & P2). - simp. - destruct (eq_sym (LittleEndianList.length_le_split (Memory.bytes_per(width:=width) sz) (word.unsigned val))) in t0, E. - subst_load_bytes_for_eq. - run1det. run1done. - eapply preserve_subset_of_xAddrs. 1: assumption. - ecancel_assumption. - - - (* SInlinetable *) - run1det. - assert (map.get (map.put l x (word.add initialL_pc (word.of_Z 4))) i = Some index). { - rewrite map.get_put_diff by congruence. assumption. - } - run1det. - assert (Memory.load sz initialL_mem - (word.add (word.add (word.add initialL_pc (word.of_Z 4)) index) (word.of_Z 0)) - = Some v). { - rewrite word.add_0_r. - eapply load_from_compile_byte_list. 1: eassumption. - wcancel_assumption. - } - run1det. - rewrite !map.put_put_same in *. - run1done. - - - (* SStackalloc *) - assert (valid_register RegisterNames.sp) by (cbv; auto). - specialize (stackalloc_always_0 x n body t mSmall l mc post). move stackalloc_always_0 at bottom. - assert (n = 0). { - eapply stackalloc_always_0. econstructor; eauto. - } - subst n. - run1det. - eapply runsTo_weaken. { - eapply H1 with (mStack := map.empty) (mCombined := mSmall). - { unfold Memory.anybytes. exists nil. reflexivity. } - { rewrite map.split_empty_r. reflexivity. } - all: IH_sidecondition. - } - simpl. - intros. - unfold Memory.anybytes, Memory.ftprint, map.of_disjoint_list_zip in *. simpl in *. - simp. - rewrite map.split_empty_r in H6p0p1. subst mSmall'. - repeat match goal with - | m: _ |- _ => destruct_RiscvMachine m; simpl_MetricRiscvMachine_get_set - end. - eexists. eexists. - split; [eassumption|]. - split; [solve [sidecondition]|]. - split; [solve [sidecondition]|]. - split. { - subst. solve_word_eq word_ok. - } - split; [solve [sidecondition]|]. - split. { - solve_MetricLog. - } - assumption. - - - (* SLit *) - RunInstruction.get_runsTo_valid_for_free. - eapply compile_lit_correct_full. - + sidecondition. - + use_sep_assumption. cbn. - (* ecancel. (* The term "Tree.Leaf (subset (footpr (program iset initialL_pc (compile_lit x v) * Rexec)%sep))" has type "Tree.Tree (set word -> Prop)" while it is expected to have type "Tree.Tree (?map -> Prop)". *) *) - eapply RelationClasses.reflexivity. - + unfold compile_stmt. simpl. ecancel_assumption. - + sidecondition. - + assumption. - + simpl. run1done; - remember (updateMetricsForLiteral v initialL_metrics) as finalMetrics; - symmetry in HeqfinalMetrics; - pose proof update_metrics_for_literal_bounded (width := width) as Hlit; - specialize Hlit with (1 := HeqfinalMetrics); - solve_MetricLog. - - - (* SOp *) - match goal with - | o: Syntax.bopname.bopname |- _ => destruct o - end; - simpl in *. - all: match goal with - | y: operand, H: context[Syntax.bopname.eq] |- _ => - destr y; simpl in *; - [ run1det; simpl_MetricRiscvMachine_get_set; run1det; run1done - | ]; try fwd - | y: operand |- _ => - destr y; simpl in *; - [ run1det; run1done - | ]; try fwd - end; simpl in *; fwd. - - all: try match goal with - | H: context[Decode.InvalidInstruction] |- _ => - assert (Encode.verify (Decode.InvalidInstruction (-1)) iset \/ - valid_InvalidInstruction (Decode.InvalidInstruction (-1))) by - ( eapply invert_ptsto_instr; ecancel_assumption) - | H: _ |- _ => run1det; run1done - end. - - all: - try match goal with - | H: Encode.verify (Decode.InvalidInstruction (-1)) iset \/ - valid_InvalidInstruction (Decode.InvalidInstruction (-1)) |- _ => - exfalso; destruct H; - [ unfold Encode.verify in H; simpl in H; - destruct H; assumption | unfold valid_InvalidInstruction in H]; fwd - end. - all: - try match goal with - | H: 0 <= -1 < 2^32 |- False - => destruct H; - match goal with - | H: 0 <= -1 |- False => destruct H; simpl; reflexivity - end - end. - all: simpl in *; fwd. - all: try match goal with - | H: ?post _ _ _ |- ?post _ _ _ => eqexact H - end. - all : rewrite ?word.srs_ignores_hibits, - ?word.sru_ignores_hibits, - ?word.slu_ignores_hibits, - ?word.mulhuu_simpl, - ?word.divu0_simpl, - ?word.modu0_simpl; trivial. - all: try solve_MetricLog. - simpl. rewrite reduce_eq_to_sub_and_lt. symmetry. apply map.put_put_same. - - - (* SSet *) - run1det. run1done. - - - (* SIf/Then *) - (* execute branch instruction, which will not jump *) - eapply runsTo_det_step_with_valid_machine; simpl in *; subst. - + assumption. - + simulate'. simpl_MetricRiscvMachine_get_set. - destruct cond; [destruct op | ]; - simpl in *; simp; repeat (simulate'; simpl_bools; simpl); try reflexivity. - + intro V. eapply runsTo_trans; simpl_MetricRiscvMachine_get_set. - * (* use IH for then-branch *) - eapply IHexec; IH_sidecondition. - * (* jump over else-branch *) - simpl. intros. destruct_RiscvMachine middle. simp. subst. - run1det. run1done. - - - (* SIf/Else *) - (* execute branch instruction, which will jump over then-branch *) - eapply runsTo_det_step_with_valid_machine; simpl in *; subst. - + assumption. - + simulate'. - destruct cond; [destruct op | ]; - simpl in *; simp; repeat (simulate'; simpl_bools; simpl); try reflexivity. - + intro V. eapply runsTo_trans; simpl_MetricRiscvMachine_get_set. - * (* use IH for else-branch *) - eapply IHexec; IH_sidecondition. - * (* at end of else-branch, i.e. also at end of if-then-else, just prove that - computed post satisfies required post *) - simpl. intros. destruct_RiscvMachine middle. simp. subst. run1done. - - - (* SLoop/again *) - on hyp[(stmt_not_too_big body1); runsTo] do (fun H => rename H into IH1). - on hyp[(stmt_not_too_big body2); runsTo] do (fun H => rename H into IH2). - on hyp[(stmt_not_too_big (SLoop body1 cond body2)); runsTo] do (fun H => rename H into IH12). - eapply runsTo_trans; simpl_MetricRiscvMachine_get_set. - + (* 1st application of IH: part 1 of loop body *) - eapply IH1; IH_sidecondition. - + simpl in *. simpl. intros. destruct_RiscvMachine middle. simp. subst. - destruct (eval_bcond middle_regs cond) as [condB|] eqn: E. - 2: exfalso; - match goal with - | H: context [_ <> None] |- _ => solve [eapply H; eauto] - end. - destruct condB. - * (* true: iterate again *) - eapply runsTo_det_step_with_valid_machine; simpl in *; subst. - { assumption. } - { simulate'. - destruct cond; [destruct op | ]; - simpl in *; simp; repeat (simulate'; simpl_bools; simpl); try reflexivity. } - { intro V. eapply runsTo_trans; simpl_MetricRiscvMachine_get_set. - - (* 2nd application of IH: part 2 of loop body *) - eapply IH2; IH_sidecondition; simpl_MetricRiscvMachine_get_set; - try eassumption; IH_sidecondition. - - simpl in *. simpl. intros. destruct_RiscvMachine middle. simp. subst. - (* jump back to beginning of loop: *) - run1det. - eapply runsTo_trans; simpl_MetricRiscvMachine_get_set. - + (* 3rd application of IH: run the whole loop again *) - eapply IH12 with (pos := pos); IH_sidecondition; simpl_MetricRiscvMachine_get_set; - try eassumption; IH_sidecondition. - + (* at end of loop, just prove that computed post satisfies required post *) - simpl. intros. destruct_RiscvMachine middle. simp. subst. - run1done. } - * (* false: done, jump over body2 *) - eapply runsTo_det_step_with_valid_machine; simpl in *; subst. - { assumption. } - { simulate'. - destruct cond; [destruct op | ]; - simpl in *; simp; repeat (simulate'; simpl_bools; simpl); try reflexivity. } - { intro V. simpl in *. run1done. } - - - (* SSeq *) - on hyp[(stmt_not_too_big s1); runsTo] do (fun H => rename H into IH1). - on hyp[(stmt_not_too_big s2); runsTo] do (fun H => rename H into IH2). - eapply runsTo_trans. - + eapply IH1; IH_sidecondition. - + simpl. intros. destruct_RiscvMachine middle. simp. subst. - eapply runsTo_trans. - * eapply IH2; IH_sidecondition; simpl_MetricRiscvMachine_get_set; - try eassumption; IH_sidecondition. - * simpl. intros. destruct_RiscvMachine middle. simp. subst. run1done. - - - (* SSkip *) - run1done. - Qed. - -End Proofs. - -*) From b8b232e9f6c6ac804af764dad56bac24014fa20f Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Fri, 9 Aug 2024 13:36:50 -0400 Subject: [PATCH 61/62] update comments to reflect current code --- bedrock2/src/bedrock2/MetricCosts.v | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/bedrock2/src/bedrock2/MetricCosts.v b/bedrock2/src/bedrock2/MetricCosts.v index fcafde638..83d0e9e36 100644 --- a/bedrock2/src/bedrock2/MetricCosts.v +++ b/bedrock2/src/bedrock2/MetricCosts.v @@ -113,8 +113,7 @@ Definition isRegStr (var : String.string) : bool := String.prefix "reg_" var. (* awkward tactic use to avoid Qed slowness *) -(* symmetry; reflexivity is black magic that shuffles orders of terms to make - heuristics choose the right things *) +(* this is slow with (eq_refl t) and fast with (eq_refl t') due to black box heuristics *) Ltac cost_unfold := repeat ( let H := match goal with From 070a202d33bebda804033a500896b954cc0d3bc2 Mon Sep 17 00:00:00 2001 From: Andy Tockman Date: Fri, 16 Aug 2024 07:28:15 -0400 Subject: [PATCH 62/62] restore loop lemmas with metrics adapted proofs --- bedrock2/src/bedrock2/MetricLoops.v | 244 +++++++++++++++++++++++++++- 1 file changed, 242 insertions(+), 2 deletions(-) diff --git a/bedrock2/src/bedrock2/MetricLoops.v b/bedrock2/src/bedrock2/MetricLoops.v index 50855548f..8517cd9eb 100644 --- a/bedrock2/src/bedrock2/MetricLoops.v +++ b/bedrock2/src/bedrock2/MetricLoops.v @@ -295,8 +295,6 @@ Section Loops. split. { exists v0, g0, l0. split. 1: eapply reconstruct_enforce; eassumption. split; eauto. } intros vi ti mi lmapi mci (gi&?&?&?&Qi); subst. destruct (hlist.foralls_forall (hlist.foralls_forall (Hbody vi) gi ti mi mci) _ ltac:(eassumption)) as (brv&brmc&?&X). - - exists brv; exists brmc; split; [assumption|]. destruct X as (Htrue&Hfalse). split; intros Hbr; [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. { eapply Proper_cmd; [ |eapply Hpc]. @@ -358,6 +356,248 @@ Section Loops. constructor. intros [y|] pf; eauto. constructor. intros [] []. Qed. + + Lemma atleastonce_localsmap + {e c t} {m : mem} {l mc} {post : _->_->_->_-> Prop} + {measure : Type} (invariant:measure->_->_->_->_->Prop) lt + (Hwf : well_founded lt) + (Henter : exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) + /\ (word.unsigned br <> 0 -> exists v', invariant v' t m l brmc) + /\ (word.unsigned br = 0%Z -> post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc))) + (Hbody : forall v t m l mc, invariant v t m l mc -> + cmd call c t m l mc (fun t m l mc => + exists br brmc, expr m l e (cost_loop_true isRegStr UNK (Some UNK) mc) (eq (Datatypes.pair br brmc)) + /\ (word.unsigned br <> 0 -> exists v', invariant v' t m l brmc /\ lt v' v) + /\ (word.unsigned br = 0 -> post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc)))) + : cmd call (cmd.while e c) t m l mc post. + Proof. + eapply wp_while. + eexists (option measure), (with_bottom lt), (fun ov t m l mc => + exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) + /\ ((word.unsigned br <> 0 -> exists v, ov = Some v /\ invariant v t m l brmc) + /\ (word.unsigned br = 0 -> ov = None /\ post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc)))). + split; auto using well_founded_with_bottom; []. split. + { destruct Henter as [br [brmc [He [Henterm Henter0]]]]. + destruct (BinInt.Z.eq_dec (word.unsigned br) 0). + { exists None, br, brmc; split; trivial. + split; intros; try contradiction; split; eauto. } + { destruct (Henterm n) as [v Hinv]. + exists (Some v), br, brmc. + split; trivial; []; split; try contradiction. + exists v; split; trivial. } } + intros vi ti mi li mci (br&brmc&Ebr&Hcontinue&Hexit). + eexists; eexists; split; [eassumption|]; split. + { intros Hc; destruct (Hcontinue Hc) as (v&?&Hinv); subst. + eapply Proper_cmd; [ |eapply Hbody; eassumption]. + intros t' m' l' mc' (br'&brmc'&Ebr'&Hinv'&Hpost'). + destruct (BinInt.Z.eq_dec (word.unsigned br') 0). + { exists None; split; try constructor. + exists br', brmc'; split; trivial; []. + split; intros; try contradiction. + split; eauto. } + { destruct (Hinv' ltac:(trivial)) as (v'&inv'<v'v). + exists (Some v'); split; trivial. (* NOTE: this [trivial] simpl-reduces [with_bottom] *) + exists br', brmc'; split; trivial. + split; intros; try contradiction. + eexists; split; eauto. } } + eapply Hexit. + Qed. + + Lemma atleastonce + {e c t l mc} {m : mem} + (variables : list String.string) + {localstuple : tuple word (length variables)} + {Pl : enforce variables localstuple l} + {measure : Type} (invariant:measure->_->_->_->ufunc word (length variables) Prop) + lt (Hwf : well_founded lt) + {post : _->_->_->_-> Prop} + (Henter : exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) + /\ (word.unsigned br <> 0 -> exists v', tuple.apply (invariant v' t m brmc) localstuple) + /\ (word.unsigned br = 0%Z -> post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc))) + (Hbody : forall v t m mc, tuple.foralls (fun localstuple => + tuple.apply (invariant v t m mc) localstuple -> + cmd call c t m (reconstruct variables localstuple) mc (fun t m l mc => + exists br brmc, expr m l e (cost_loop_true isRegStr UNK (Some UNK) mc) (eq (Datatypes.pair br brmc)) + /\ (word.unsigned br <> 0 -> Markers.unique (Markers.left (tuple.existss (fun localstuple => enforce variables localstuple l /\ Markers.right (Markers.unique (exists v', tuple.apply (invariant v' t m brmc) localstuple /\ lt v' v)))))) + /\ (word.unsigned br = 0 -> post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc))))) + : cmd call (cmd.while e c) t m l mc post. + Proof. + eapply (atleastonce_localsmap (fun v t m l mc => exists localstuple, Logic.and (enforce variables localstuple l) (tuple.apply (invariant v t m mc) localstuple))); eauto. + 1: { + destruct Henter as [br [brmc [He [Henterm Henter0]]]]. + do 3 eexists; eauto; split; eauto. + intro Hbr; destruct (Henterm Hbr); eauto. + } + intros vi ti mi li mci (?&X&Y). + specialize (Hbody vi ti mi mci). + eapply hlist.foralls_forall in Hbody. + specialize (Hbody Y). + rewrite <-(reconstruct_enforce _ _ _ X) in Hbody. + eapply Proper_cmd; [ |eapply Hbody]. + intros t' m' l' mc' (?&?&?&HH&?). + eexists; eexists; split; eauto. + split; intros; eauto. + specialize (HH ltac:(eauto)). + eapply hlist.existss_exists in HH; destruct HH as (?&?&?&?&?). + eexists; split; eauto. + Qed. + + Lemma tailrec_earlyout_localsmap + {e c t} {m : mem} {l mc} {post : _-> _->_->_-> Prop} + {measure : Type} (spec:_->_->_->_->_->(Prop*(_->_->_->_-> Prop))) lt + (Hwf : well_founded lt) + (v0 : measure) (P0 := (spec v0 t m l mc).(1)) (Hpre : P0) + (Q0 := (spec v0 t m l mc).(2)) + (Hbody : forall v t m l mc, + let S := spec v t m l mc in let (P, Q) := S in + P -> + exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) /\ + (word.unsigned br <> 0%Z -> cmd call c t m l brmc + (fun t' m' l' mc' => + (exists br brmc, + expr m' l' e (cost_loop_true isRegStr UNK (Some UNK) mc') (eq (Datatypes.pair br brmc)) + /\ word.unsigned br = 0 /\ Q t' m' l' (cost_loop_false isRegStr UNK (Some UNK) brmc)) \/ + exists v', let S' := spec v' t' m' l' (cost_loop_true isRegStr UNK (Some UNK) mc') in let '(P', Q') := S' in + P' /\ + lt v' v /\ + forall T M L MC, Q' T M L MC -> Q T M L MC)) /\ + (word.unsigned br = 0%Z -> Q t m l (cost_loop_false isRegStr UNK (Some UNK) brmc))) + (Hpost : forall t m l mc, Q0 t m l mc -> post t m l mc) + : cmd call (cmd.while e c) t m l mc post. + Proof. + eapply wp_while. + eexists (option measure), (with_bottom lt), (fun v t m l mc => + match v with + | None => exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) /\ word.unsigned br = 0 /\ Q0 t m l (cost_loop_false isRegStr UNK (Some UNK) brmc) + | Some v => + let S := spec v t m l mc in let '(P, Q) := S in + P /\ forall T M L MC, Q T M L MC -> Q0 T M L MC + end). + split; auto using well_founded_with_bottom; []; cbv [Markers.split] in *. + split. + { exists (Some v0); eauto. } + intros [vi|] ti mi li mci inv_i; [destruct inv_i as (?&Qi)|destruct inv_i as (br&Hebr&Hbr0&HQ)]. + { destruct (Hbody _ _ _ _ _ ltac:(eassumption)) as (br&brmc&?&X); exists br, brmc; split; [assumption|]. + destruct X as (Htrue&Hfalse). split; intros Hbr; + [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; eauto. + eapply Proper_cmd; [ |eapply Hpc]. + intros tj mj lj mcj [(br'&brmc'&Hbr'&Hz&HQ)|(vj&dP&?&dQ)]; + [exists None | exists (Some vj)]; cbn [with_bottom]; eauto 9. } + repeat esplit; destruct HQ; eauto; contradiction. + Qed. + + Lemma tailrec_earlyout + {e c t localsmap} {m : mem} + (ghosttypes : polymorphic_list.list Type) + (variables : list String.string) + {l0 : tuple word (length variables)} {mc} + {Pl : enforce variables l0 localsmap} + {post : _->_->_->_-> Prop} + {measure : Type} (spec:_->HList.arrows ghosttypes (_->_->ufunc word (length variables) (MetricLog->Prop*(_->_->ufunc word (length variables) (MetricLog->Prop))))) lt + (Hwf : well_founded lt) + (v0 : measure) + : hlist.foralls (fun (g0 : hlist ghosttypes) => forall + (Hpre : (tuple.apply (hlist.apply (spec v0) g0 t m) l0 mc).(1)) + (Hbody : forall v, hlist.foralls (fun g => forall t m mc, tuple.foralls (fun l => + @dlet _ (fun _ => Prop) (reconstruct variables l) (fun localsmap : locals => + match tuple.apply (hlist.apply (spec v) g t m) l mc with S_ => + S_.(1) -> + Markers.unique (Markers.left (exists br brmc, expr m localsmap e mc (eq (Datatypes.pair br brmc)) /\ Markers.right ( + (word.unsigned br <> 0%Z -> cmd call c t m localsmap brmc + (fun t' m' localsmap' mc' => + Markers.unique (Markers.left (hlist.existss (fun l' => enforce variables l' localsmap' /\ Markers.right ( + Markers.unique (Markers.left (exists br brmc, expr m' localsmap' e (cost_loop_true isRegStr UNK (Some UNK) mc') (eq (Datatypes.pair br brmc)) /\ Markers.right ( word.unsigned br = 0 /\ tuple.apply (S_.(2) t' m') l' (cost_loop_false isRegStr UNK (Some UNK) brmc)) ) ) \/ + Markers.unique (Markers.left (hlist.existss (fun g' => exists v', + match tuple.apply (hlist.apply (spec v') g' t' m') l' (cost_loop_true isRegStr UNK (Some UNK) mc') with S' => + S'.(1) /\ Markers.right ( + lt v' v /\ + forall T M, hlist.foralls (fun L => forall MC, tuple.apply (S'.(2) T M) L MC -> tuple.apply (S_.(2) T M) L MC)) end))))))))) /\ + (word.unsigned br = 0%Z -> tuple.apply (S_.(2) t m) l (cost_loop_false isRegStr UNK (Some UNK) brmc)))))end)))) + (Hpost : match (tuple.apply (hlist.apply (spec v0) g0 t m) l0 mc).(2) with Q0 => forall t m, hlist.foralls (fun l => forall mc, tuple.apply (Q0 t m) l mc -> post t m (reconstruct variables l) mc)end) + , cmd call (cmd.while e c) t m localsmap mc post). + Proof. + eapply hlist_forall_foralls; intros g0 **. + eapply wp_while. + eexists (option measure), (with_bottom lt), (fun vi ti mi localsmapi mci => + exists li, localsmapi = reconstruct variables li /\ + match vi with + | None => exists br brmc, expr mi localsmapi e mci (eq (Datatypes.pair br brmc)) + /\ word.unsigned br = 0 /\ tuple.apply ((tuple.apply (hlist.apply (spec v0) g0 t m) l0 mc).(2) ti mi) li (cost_loop_false isRegStr UNK (Some UNK) brmc) + | Some vi => exists gi, match tuple.apply (hlist.apply (spec vi) gi ti mi) li mci with S_ => + S_.(1) /\ forall T M L MC, tuple.apply (S_.(2) T M) L MC -> + tuple.apply ((tuple.apply (hlist.apply (spec v0) g0 t m) l0 mc).(2) T M) L MC end end). + cbv [Markers.unique Markers.split Markers.left Markers.right] in *. + split; eauto using well_founded_with_bottom. + split. { exists (Some v0), l0. split. 1: eapply reconstruct_enforce; eassumption. exists g0; split; eauto. } + intros [vi|] ti mi lmapi mci. + 2: { intros (ld&Hd&br&brmc&Hbr&Hz&Hdone). + eexists; eexists; split; eauto. + split; intros; try contradiction. + subst; eapply (hlist.foralls_forall (Hpost ti mi) _ _ Hdone). } + intros (?&?&gi&?&Qi); subst. + destruct (hlist.foralls_forall (hlist.foralls_forall (Hbody vi) gi ti mi mci) _ ltac:(eassumption)) as (br&brmc&?&X). + exists br, brmc; split; [assumption|]. destruct X as (Htrue&Hfalse). split; intros Hbr; + [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. + { eapply Proper_cmd; [ |eapply Hpc]. + intros tj mj lmapj mcj Hlj; eapply hlist.existss_exists in Hlj. + destruct Hlj as (lj&Elj&HE); eapply reconstruct_enforce in Elj; subst lmapj. + destruct HE as [(br'&brmc'&Hevalr'&Hz'&Hdone)|HE]. + { exists None; cbn. eauto 9. } + { eapply hlist.existss_exists in HE. destruct HE as (l&?&?&?&HR). + pose proof fun T M => hlist.foralls_forall (HR T M); clear HR. + eexists (Some _); eauto 9. } } + { pose proof fun t m => hlist.foralls_forall (Hpost t m); clear Hpost; eauto. } + Qed. + + Lemma while_zero_iterations {e c t l mc brmc} {m : mem} {post : _->_->_->_-> Prop} + (HCond: expr m l e mc (eq (Datatypes.pair (word.of_Z 0) brmc))) + (HPost: post t m l (cost_loop_false isRegStr UNK (Some UNK) brmc)) + : cmd call (cmd.while e c) t m l mc post. + Proof. + eapply (while_localsmap (fun n t' m' l' mc' => t' = t /\ m' = m /\ l' = l /\ mc' = mc) (PeanoNat.Nat.lt_wf 0) 0%nat). + 1: unfold split; auto. intros *. intros (? & ? & ? & ?). subst. + eexists. eexists. split. 1: exact HCond. + rewrite Properties.word.unsigned_of_Z_0. + split; intros; congruence. + Qed. + + + (* Bedrock-style loop rule *) + Local Open Scope sep_scope. + Local Infix "*" := Separation.sep : type_scope. + Local Infix "==>" := Lift1Prop.impl1. + + Lemma tailrec_sep + e c t (m : mem) l mc (post : _->_->_->_-> Prop) + {measure : Type} P Q lt (Hwf : well_founded lt) (v0 : measure) R0 + (Hpre : (P v0 t l mc * R0) m) + (Hbody : forall v t m l mc R, (P v t l mc * R) m -> + exists br brmc, expr m l e mc (eq (Datatypes.pair br brmc)) /\ + (word.unsigned br <> 0%Z -> cmd call c t m l brmc + (fun t' m' l' mc' => exists v' dR, (P v' t' l' (cost_loop_true isRegStr UNK (Some UNK) mc') * (R * dR)) m' /\ + lt v' v /\ + forall T L MC, Q v' T L MC * dR ==> Q v T L MC)) /\ + (word.unsigned br = 0%Z -> (Q v t l (cost_loop_false isRegStr UNK (Some UNK) brmc) * R) m)) + (Hpost : forall t m l mc, (Q v0 t l mc * R0) m -> post t m l mc) + : cmd call (cmd.while e c) t m l mc post. + Proof. + eapply wp_while. + eexists measure, lt, (fun v t m l mc => exists R, (P v t l mc * R) m /\ + forall T L MC, Q v T L MC * R ==> Q v0 T L MC * R0). + split; [assumption|]. + split. { exists v0, R0. split; [assumption|]. intros. reflexivity. } + intros vi ti mi li mci (Ri&?&Qi). + destruct (Hbody _ _ _ _ _ _ ltac:(eassumption)) as (br&brmc&?&X); exists br, brmc; split; [assumption|]. + destruct X as (Htrue&Hfalse). split; intros Hbr; + [pose proof(Htrue Hbr)as Hpc|pose proof(Hfalse Hbr)as Hpc]; clear Hbr Htrue Hfalse. + { eapply Proper_cmd; [ |eapply Hpc]. + intros tj mj lj mcj (vj&dR&dP&?&dQ). + exists vj; split; [|assumption]. + exists (Ri * dR); split; [assumption|]. + intros. rewrite (sep_comm _ dR), <-(sep_assoc _ dR), dQ; trivial. } + { eapply Hpost, Qi, Hpc. } + Qed. + End Loops. Ltac loop_simpl :=