diff --git a/bedrock2/src/bedrock2/FE310CSemantics.v b/bedrock2/src/bedrock2/FE310CSemantics.v index de4b7f845..37d965de5 100644 --- a/bedrock2/src/bedrock2/FE310CSemantics.v +++ b/bedrock2/src/bedrock2/FE310CSemantics.v @@ -1,5 +1,5 @@ Require Import Coq.ZArith.ZArith. -Require Import bedrock2.Syntax bedrock2.Semantics. +Require Import bedrock2.Syntax bedrock2.Semantics bedrock2.LeakageSemantics bedrock2.MetricLeakageSemantics. Require coqutil.Datatypes.String coqutil.Map.SortedList coqutil.Map.SortedListString. Require Import coqutil.Word.Interface. Require Export coqutil.Word.Bitwidth32. @@ -27,24 +27,25 @@ Section WithParameters. Definition isMMIOAligned (n : nat) (addr : word) := n = 4%nat /\ word.unsigned addr mod 4 = 0. - Global Instance ext_spec: ExtSpec := - fun (t : trace) (mGive : mem) a (args: list word) (post: mem -> list word -> Prop) => +(* FE310 is a simple enough processor that our leakage assumptions are likely to hold. There is no official documentation of whether multiply always takes the maximum time or not, but both https://eprint.iacr.org/2019/794.pdf and https://pure.tue.nl/ws/portalfiles/portal/169647601/Berg_S._ES_CSE.pdf quote a fixed number of cycles for FE310 multiplication in the context of cryptography. *) + Global Instance leakage_ext_spec: LeakageSemantics.ExtSpec := + fun (t : trace) (mGive : mem) a (args: list word) (post: mem -> list word -> list word -> Prop) => if String.eqb "MMIOWRITE" a then exists addr val, args = [addr; val] /\ (mGive = Interface.map.empty /\ isMMIOAddr addr /\ word.unsigned addr mod 4 = 0) /\ - post Interface.map.empty nil + post Interface.map.empty nil [addr] else if String.eqb "MMIOREAD" a then exists addr, args = [addr] /\ (mGive = Interface.map.empty /\ isMMIOAddr addr /\ word.unsigned addr mod 4 = 0) /\ - forall val, post Interface.map.empty [val] + forall val, post Interface.map.empty [val] [addr] else False. - Global Instance ext_spec_ok : ext_spec.ok ext_spec. + Global Instance leakage_ext_spec_ok : ext_spec.ok leakage_ext_spec. Proof. split; - cbv [ext_spec Morphisms.Proper Morphisms.respectful Morphisms.pointwise_relation Basics.impl]; + cbv [leakage_ext_spec Morphisms.Proper Morphisms.respectful Morphisms.pointwise_relation Basics.impl]; intros. all : repeat match goal with @@ -53,9 +54,16 @@ Section WithParameters. | H: exists _, _ |- _ => destruct H | H: _ /\ _ |- _ => destruct H | H: False |- _ => destruct H - end; subst; eauto 8 using Properties.map.same_domain_refl. + end; subst; + repeat match goal with + | H: _ :: _ = _ :: _ |- _ => injection H; intros; subst; clear H + end; + eauto 8 using Properties.map.same_domain_refl. Qed. + Global Instance ext_spec : Semantics.ExtSpec := deleakaged_ext_spec. + Global Instance ext_spec_ok : Semantics.ext_spec.ok ext_spec := deleakaged_ext_spec_ok. + Global Instance locals: Interface.map.map String.string word := SortedListString.map _. Global Instance env: Interface.map.map String.string (list String.string * list String.string * cmd) := SortedListString.map _. diff --git a/bedrock2/src/bedrock2/LeakageLoops.v b/bedrock2/src/bedrock2/LeakageLoops.v new file mode 100644 index 000000000..b981a0be6 --- /dev/null +++ b/bedrock2/src/bedrock2/LeakageLoops.v @@ -0,0 +1,597 @@ +Require Import List. +Import ListNotations. +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 Semantics LeakageSemantics Markers. +From bedrock2 Require Import LeakageWeakestPrecondition LeakageWeakestPreconditionProperties. + +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: LeakageSemantics.ExtSpec} {pick_sp: PickSp}. + Context {word_ok : word.ok word} {mem_ok : map.ok mem}. + Context {locals_ok : map.ok locals}. + Context {ext_spec_ok : LeakageSemantics.ext_spec.ok ext_spec}. + + Context {fs : env}. + Let call := fs. + + Lemma wp_while: forall e c k t m l (post: _ -> _ -> _ -> _ -> Prop), + (exists measure (lt:measure->measure->Prop) (inv:measure->leakage->trace->mem->locals->Prop), + Coq.Init.Wf.well_founded lt /\ + (exists v, inv v k t m l) /\ + (forall v k t m l, inv v k t m l -> + exists b k', dexpr m l k e b k' /\ + (word.unsigned b <> 0%Z -> cmd call c (leak_bool true :: k') t m l (fun k'' t' m l => + exists v', inv v' k'' t' m l /\ lt v' v)) /\ + (word.unsigned b = 0%Z -> post (leak_bool false :: k') t m l))) -> + cmd call (cmd.while e c) k t m l post. + Proof. + intros. destruct H as (measure & lt & inv & Hwf & HInit & Hbody). + destruct HInit as (v0 & HInit). + revert k t m l HInit. pattern v0. revert v0. + eapply (well_founded_ind Hwf). intros. + specialize Hbody with (1 := HInit). destruct Hbody as (b & k' & Hexpr & Ht & Hf). + eapply expr_sound in Hexpr. destruct Hexpr as (b' & k'' & Hb & Hk' & ?). subst b' k''. + 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 k t} {m: mem} {l} {post : leakage -> trace -> mem -> locals -> Prop} + {measure: Type} {Ghost: Type} + (P Q: measure -> Ghost -> leakage -> trace -> mem -> locals -> Prop) + (lt: measure -> measure -> Prop) + (Hwf: well_founded lt) + (v0: measure) (g0: Ghost) + (Hpre: P v0 g0 k t m l) + (Hbody: forall v g k t m l, + P v g k t m l -> + exists br k', dexpr m l k e br k' /\ + (word.unsigned br <> 0%Z -> cmd call c (leak_bool true :: k') t m l + (fun k'' t' m' l' => exists v' g', + P v' g' k'' t' m' l' /\ + lt v' v /\ + (forall k''' t'' m'' l'', Q v' g' k''' t'' m'' l'' -> Q v g k''' t'' m'' l''))) /\ + (word.unsigned br = 0%Z -> Q v g (leak_bool false :: k') t m l)) + (Hpost: forall k t m l, Q v0 g0 k t m l -> post k t m l) + : cmd call (cmd.while e c) k t m l post. + Proof. + eapply wp_while. + eexists measure, lt, (fun v k t m l => + exists g, P v g k t m l /\ forall k'' t'' m'' l'', Q v g k'' t'' m'' l'' -> Q v0 g0 k'' t'' m'' l''). + split; [assumption|]. + split; [solve[eauto]|]. + intros vi ki ti mi li (gi & HPi & HQimpl). + specialize (Hbody vi gi ki ti mi li HPi). + destruct Hbody as (br & k' & ? & Hbody). exists br, k'; 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 kj tj mj lj (vj& gj & HPj & Hlt & Qji); eauto 9. } + { eauto. } + Qed. + + Lemma tailrec_localsmap_1ghost_parameterized_finalpost + {e c rest k t} {m: mem} {l} + {measure: Type} {Ghost: Type} + (P Q: measure -> Ghost -> leakage -> trace -> mem -> locals -> Prop) + (lt: measure -> measure -> Prop) + (Hwf: well_founded lt) + (v0: measure) (g0: Ghost) + (Hpre: P v0 g0 k t m l) + (Hbody: forall v g k t m l, + P v g k t m l -> + exists br k', dexpr m l k e br k' /\ + (word.unsigned br <> 0%Z -> cmd call c (leak_bool true :: k') t m l + (fun k'' t' m' l' => exists v' g', + P v' g' k'' t' m' l' /\ + lt v' v /\ + (forall k''' t'' m'' l'', Q v' g' k''' t'' m'' l'' -> Q v g k''' t'' m'' l''))) /\ + (word.unsigned br = 0%Z -> cmd call rest (leak_bool false :: k') t m l (Q v g))) + : cmd call (cmd.seq (cmd.while e c) rest) k t m l (Q v0 g0). + Proof. + cbn. eapply tailrec_localsmap_1ghost with + (Q := fun v g k t m l => cmd call rest k t m l (Q v g)). + 1: eassumption. + 1: exact Hpre. + 2: intros *; exact id. + intros vi gi ki ti mi li HPi. + specialize (Hbody vi gi ki ti mi li HPi). + destruct Hbody as (br & k' & ? & Hbody). exists br, k'; 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 kj tj mj lj (vj& gj & HPj & Hlt & Qji). do 2 eexists. + split. 1: eassumption. split. 1: assumption. + intros. + eapply Proper_cmd. + 2: eassumption. + intros kk 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 k t l} {m : mem} + {measure : Type} (invariant:_->_->_->_->_->Prop) + {lt} (Hwf : well_founded lt) (v0 : measure) + {post : _->_->_->_-> Prop} + (Hpre : invariant v0 k t m l) + (Hbody : forall v k t m l, + invariant v k t m l -> + exists br k', dexpr m l k e br k' /\ + (word.unsigned br <> 0 -> + cmd fs c (leak_bool true :: k') t m l (fun k t m l => exists v', invariant v' k t m l /\ lt v' v)) /\ + (word.unsigned br = 0 -> post (leak_bool false :: k') t m l)) + : cmd fs (cmd.while e c) k t m l post. + Proof. + eapply wp_while. + eexists measure, lt, invariant. + split. 1: exact Hwf. + split. 1: eauto. + exact Hbody. + Qed. + + Lemma while + {e c k t l} {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 k t m) localstuple) + (Hbody : forall v k t m, tuple.foralls (fun localstuple => + tuple.apply (invariant v k t m) localstuple -> + let l := reconstruct variables localstuple in + exists br k', dexpr m l k e br k' /\ + (word.unsigned br <> 0 -> + cmd call c (leak_bool true :: k') t m l (fun k t m l => + Markers.unique (Markers.left (tuple.existss (fun localstuple => + enforce variables localstuple l /\ + Markers.right (Markers.unique (exists v', + tuple.apply (invariant v' k t m) localstuple /\ lt v' v))))))) /\ + (word.unsigned br = 0 -> post (leak_bool false :: k') t m l))) + : cmd call (cmd.while e c) k t m l post. + Proof. + eapply (while_localsmap (fun v k t m l => + exists localstuple, enforce variables localstuple l /\ + tuple.apply (invariant v k t m) localstuple)); + unfold Markers.split; eauto. + intros vi ki ti mi li (?&X&Y). + specialize (Hbody vi ki ti mi). + eapply hlist.foralls_forall in Hbody. + specialize (Hbody Y). + rewrite <-(reconstruct_enforce _ _ _ X) in Hbody. + destruct Hbody as (br & k' & Cond & Again & Done). + exists br, k'. 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 k'' t' m' l' 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 k 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 k t m) l0).(1)) + (Hbody : forall v, hlist.foralls (fun g => forall k t m, tuple.foralls (fun l => + @dlet _ (fun _ => Prop) (reconstruct variables l) (fun localsmap : locals => + match tuple.apply (hlist.apply (spec v) g k t m) l with S_ => + S_.(1) -> + Markers.unique (Markers.left (exists br k', dexpr m localsmap k e br k' /\ Markers.right ( + (word.unsigned br <> 0%Z -> cmd call c (leak_bool true :: k') t m localsmap + (fun k'' t' m' localsmap' => + 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' k'' t' m') l' with S' => + S'.(1) /\ Markers.right ( + lt v' v /\ + forall K T M, hlist.foralls (fun L => tuple.apply (S'.(2) K T M) L -> tuple.apply (S_.(2) K T M) L)) end))))))))) /\ + (word.unsigned br = 0%Z -> tuple.apply (S_.(2) (leak_bool false :: k') t m) l))))end)))) + (Hpost : match (tuple.apply (hlist.apply (spec v0) g0 k t m) l0).(2) with Q0 => forall k t m, hlist.foralls (fun l => tuple.apply (Q0 k t m) l -> post k t m (reconstruct variables l))end) + , cmd call (cmd.while e c) k t m localsmap post ). + Proof. + eapply hlist_forall_foralls; intros g0 **. + eapply wp_while. + eexists measure, lt, (fun vi ki ti mi localsmapi => + exists gi li, localsmapi = reconstruct variables li /\ + match tuple.apply (hlist.apply (spec vi) gi ki ti mi) li with S_ => + S_.(1) /\ forall K T M L, tuple.apply (S_.(2) K T M) L -> + tuple.apply ((tuple.apply (hlist.apply (spec v0) g0 k t m) l0).(2) K T M) L 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 ki ti mi lmapi (gi&?&?&?&Qi); subst. + destruct (hlist.foralls_forall (hlist.foralls_forall (Hbody vi) gi ki ti mi) _ ltac:(eassumption)) as (br&k'&?&X). + exists br, k'; 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 kj tj mj lmapj Hlj; eapply hlist.existss_exists in Hlj. + 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 K T M => hlist.foralls_forall (HR K T M); clear HR. + eauto 9. } + { pose proof fun k t m => hlist.foralls_forall (Hpost k t m); clear Hpost; eauto. } + Qed. + + Lemma tailrec_localsmap + {e c k t} {m : mem} {l} {post : _->_->_->_-> Prop} + {measure : Type} (spec:_->_->_->_->_->(Prop*(_->_->_->_-> Prop))) lt + (Hwf : well_founded lt) + (v0 : measure) (P0 := (spec v0 k t m l).(1)) (Hpre : P0) + (Q0 := (spec v0 k t m l).(2)) + (Hbody : forall v k t m l, + let S := spec v k t m l in let (P, Q) := S in + P -> + exists br k', dexpr m l k e br k' /\ + (word.unsigned br <> 0%Z -> cmd call c (leak_bool true :: k') t m l + (fun k' t' m' l' => exists v', + let S' := spec v' k' t' m' l' in let '(P', Q') := S' in + P' /\ + lt v' v /\ + forall K T M L, Q' K T M L -> Q K T M L)) /\ + (word.unsigned br = 0%Z -> Q (leak_bool false :: k') t m l)) + (Hpost : forall k t m l, Q0 k t m l -> post k t m l) + : cmd call (cmd.while e c) k t m l post. + Proof. + eapply wp_while. + eexists measure, lt, (fun v k t m l => + let S := spec v k t m l in let '(P, Q) := S in + P /\ forall K T M L, Q K T M L -> Q0 K T M L). + split; [assumption|]. + cbv [Markers.split] in *. + split; [solve[eauto]|]. + intros vi ki ti mi li (?&Qi). + destruct (Hbody _ _ _ _ _ ltac:(eassumption)) as (br&k'&?&X); exists br, k'; 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 kj tj mj lj (vj&dP&?&dQ); eauto 9. } + { 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 k t} {m : mem} {l} {post : _->_->_->_-> Prop} + {measure : Type} (invariant:_->_->_->_->_->Prop) lt + (Hwf : well_founded lt) + (v0 : measure) + (Henter : exists br k', dexpr m l k e br k' /\ + (word.unsigned br = 0%Z -> post (leak_bool false :: k') t m l) /\ + (word.unsigned br <> 0%Z -> invariant v0 (leak_bool true :: k') t m l)) + (Hbody : forall v k t m l, invariant v k t m l -> + cmd call c k t m l (fun k t m l => + exists br k', dexpr m l k e br k' /\ + (word.unsigned br <> 0 -> exists v', invariant v' (leak_bool true :: k') t m l /\ lt v' v) /\ + (word.unsigned br = 0 -> post (leak_bool false :: k') t m l))) + : cmd call (cmd.while e c) k t m l post. + Proof. + eapply wp_while. + eexists (option measure), (with_bottom lt), (fun ov k t m l => + exists br k', dexpr m l k e br k' /\ + ((word.unsigned br <> 0 -> exists v, ov = Some v /\ invariant v (leak_bool true :: k') t m l) /\ + (word.unsigned br = 0 -> ov = None /\ post (leak_bool false :: k') t m l))). + split; auto using well_founded_with_bottom; []. split. + { destruct Henter as (br & k' & He & Henterfalse & Hentertrue). + destruct (BinInt.Z.eq_dec (word.unsigned br) 0). + { exists None, br, k'; split; trivial. + split; intros; try contradiction; split; eauto. } + { exists (Some v0), br, k'. + split; trivial; []; split; try contradiction. + exists v0; split; auto. } } + intros vi ki ti mi li (br&k'&Ebr&Hcontinue&Hexit). + eexists; eexists; split; [eassumption|]; split. + { intros Hc; destruct (Hcontinue Hc) as (v&?&Hinv); subst. + eapply Proper_cmd; [ |eapply Hbody; eassumption]. + intros k'' t' m' l' (br'&k'2&Ebr'&Hinv'&Hpost'). + destruct (BinInt.Z.eq_dec (word.unsigned br') 0). + { exists None; split; try constructor. + exists br', k'2; 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', k'2; split; trivial. + split; intros; try contradiction. + eexists; split; eauto. } } + eapply Hexit. + Qed. + + Lemma tailrec_earlyout_localsmap + {e c k t} {m : mem} {l} {post : _->_->_->_-> Prop} + {measure : Type} (spec:_->_->_->_->_->(Prop*(_->_->_->_-> Prop))) lt + (Hwf : well_founded lt) + (v0 : measure) (P0 := (spec v0 k t m l).(1)) (Hpre : P0) + (Q0 := (spec v0 k t m l).(2)) + (Hbody : forall v k t m l, + let S := spec v k t m l in let (P, Q) := S in + P -> + exists br k', dexpr m l k e br k' /\ + (word.unsigned br <> 0%Z -> cmd call c (leak_bool true :: k') t m l + (fun k'' t' m' l' => + (exists br k''', dexpr m' l' k'' e br k''' /\ word.unsigned br = 0 /\ Q (leak_bool false :: k''') t' m' l') \/ + exists v', let S' := spec v' k'' t' m' l' in let '(P', Q') := S' in + P' /\ + lt v' v /\ + forall K T M L, Q' K T M L -> Q K T M L)) /\ + (word.unsigned br = 0%Z -> Q (leak_bool false :: k') t m l)) + (Hpost : forall k t m l, Q0 k t m l -> post k t m l) + : cmd call (cmd.while e c) k t m l post. + Proof. + eapply wp_while. + eexists (option measure), (with_bottom lt), (fun v k t m l => + match v with + | None => exists br k', dexpr m l k e br k' /\ word.unsigned br = 0 /\ Q0 (leak_bool false :: k') t m l + | Some v => + let S := spec v k t m l in let '(P, Q) := S in + P /\ forall K T M L, Q K T M L -> Q0 K T M L + end). + split; auto using well_founded_with_bottom; []; cbv [Markers.split] in *. + split. + { exists (Some v0); eauto. } + intros [vi|] ki ti mi li inv_i; [destruct inv_i as (?&Qi)|destruct inv_i as (br&k'&Hebr&Hbr0&HQ)]. + { destruct (Hbody _ _ _ _ _ ltac:(eassumption)) as (br&k'&?&X); exists br, k'; 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 kj tj mj lj [(br'&k''&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 k 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 k t m) l0).(1)) + (Hbody : forall v, hlist.foralls (fun g => forall k t m, tuple.foralls (fun l => + @dlet _ (fun _ => Prop) (reconstruct variables l) (fun localsmap : locals => + match tuple.apply (hlist.apply (spec v) g k t m) l with S_ => + S_.(1) -> + Markers.unique (Markers.left (exists br k', dexpr m localsmap k e br k' /\ Markers.right ( + (word.unsigned br <> 0%Z -> cmd call c (leak_bool true :: k') t m localsmap + (fun k'' t' m' localsmap' => + Markers.unique (Markers.left (hlist.existss (fun l' => enforce variables l' localsmap' /\ Markers.right ( + Markers.unique (Markers.left (exists br k''', dexpr m' localsmap' k'' e br k''' /\ Markers.right ( word.unsigned br = 0 /\ tuple.apply (S_.(2) (leak_bool false :: k''') t' m') l') ) ) \/ + Markers.unique (Markers.left (hlist.existss (fun g' => exists v', + match tuple.apply (hlist.apply (spec v') g' k'' t' m') l' with S' => + S'.(1) /\ Markers.right ( + lt v' v /\ + forall K T M, hlist.foralls (fun L => tuple.apply (S'.(2) K T M) L -> tuple.apply (S_.(2) K T M) L)) end))))))))) /\ + (word.unsigned br = 0%Z -> tuple.apply (S_.(2) (leak_bool false :: k') t m) l))))end)))) + (Hpost : match (tuple.apply (hlist.apply (spec v0) g0 k t m) l0).(2) with Q0 => forall k t m, hlist.foralls (fun l => tuple.apply (Q0 k t m) l -> post k t m (reconstruct variables l))end) + , cmd call (cmd.while e c) k t m localsmap post ). + Proof. + eapply hlist_forall_foralls; intros g0 **. + eapply wp_while. + eexists (option measure), (with_bottom lt), (fun vi ki ti mi localsmapi => + exists li, localsmapi = reconstruct variables li /\ + match vi with + | None => exists br ki', dexpr mi localsmapi ki e br ki' /\ + word.unsigned br = 0 /\ tuple.apply ((tuple.apply (hlist.apply (spec v0) g0 k t m) l0).(2) (leak_bool false :: ki') ti mi) li + | Some vi => exists gi, + match tuple.apply (hlist.apply (spec vi) gi ki ti mi) li with + | S_ => + S_.(1) /\ forall K T M L, tuple.apply (S_.(2) K T M) L -> + tuple.apply ((tuple.apply (hlist.apply (spec v0) g0 k t m) l0).(2) K 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|] ki ti mi lmapi. + 2: { intros (ld&Hd&br&ki'&Hbr&Hz&Hdone). + eexists; eexists; split; eauto. + split; intros; try contradiction. + subst; eapply (hlist.foralls_forall (Hpost (leak_bool false :: ki') ti mi) _ Hdone). } + intros (?&?&gi&?&Qi); subst. + destruct (hlist.foralls_forall (hlist.foralls_forall (Hbody vi) gi ki ti mi) _ ltac:(eassumption)) as (br&k'&?&X). + exists br, k'; 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 kj 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'&k''&Hevalr'&Hz'&Hdone)|HE]. + { exists None; cbn. eauto 9. } + { eapply hlist.existss_exists in HE. destruct HE as (l&?&?&?&HR). + pose proof fun K T M => hlist.foralls_forall (HR K T M); clear HR. + eexists (Some _); eauto 9. } } + { pose proof fun k t m => hlist.foralls_forall (Hpost k t m); clear Hpost; eauto. } + Qed. + + + Lemma atleastonce + {e c k 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} + (v0 : measure) + (Henter : exists br k', dexpr m l k e br k' /\ + (word.unsigned br = 0%Z -> post (leak_bool false :: k') t m l) /\ + (word.unsigned br <> 0%Z -> tuple.apply (invariant v0 (leak_bool true :: k') t m) localstuple)) + (Hbody : forall v k t m, tuple.foralls (fun localstuple => + tuple.apply (invariant v k t m) localstuple -> + cmd call c k t m (reconstruct variables localstuple) (fun k t m l => + exists br k', dexpr m l k e br k' /\ + (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' (leak_bool true :: k') t m) localstuple /\ lt v' v)))))) /\ + (word.unsigned br = 0 -> post (leak_bool false :: k') t m l)))) + : cmd call (cmd.while e c) k t m l post. + Proof. + destruct Henter as (br & k' & Hbr & Henterfalse & Hentertrue). + eapply (atleastonce_localsmap (fun v k t m l => exists localstuple, Logic.and (enforce variables localstuple l) (tuple.apply (invariant v k t m) localstuple))); eauto. + { eexists. eexists. split; eauto. split; eauto. } + intros vi ki ti mi li (?&X&Y). + specialize (Hbody vi ki ti mi). + eapply hlist.foralls_forall in Hbody. + specialize (Hbody Y). + rewrite <-(reconstruct_enforce _ _ _ X) in Hbody. + eapply Proper_cmd; [ |eapply Hbody]. + intros k'' t' m' l' (?&?&?&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 while_zero_iterations {e c k t l} {m : mem} {post : _->_->_->_-> Prop} + (HCondPost: exists k', dexpr m l k e (word.of_Z 0) k' /\ post (leak_bool false :: k') t m l) + : cmd call (cmd.while e c) k t m l post. + Proof. + destruct HCondPost as (k' & HCond & Hpost). + eapply (while_localsmap (fun n k' t' m' l' => k' = k /\ t' = t /\ m' = m /\ l' = l) (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 k t (m : mem) l (post : _->_->_->_-> Prop) + {measure : Type} P Q lt (Hwf : well_founded lt) (v0 : measure) R0 + (Hpre : (P v0 k t l * R0) m) + (Hbody : forall v k t m l R, (P v k t l * R) m -> + exists br k', dexpr m l k e br k' /\ + (word.unsigned br <> 0%Z -> cmd call c (leak_bool true :: k') t m l + (fun k'' t' m' l' => exists v' dR, (P v' k'' t' l' * (R * dR)) m' /\ + lt v' v /\ + forall K T L, Q v' K T L * dR ==> Q v K T L)) /\ + (word.unsigned br = 0%Z -> (Q v (leak_bool false :: k') t l * R) m)) + (Hpost : forall k t m l, (Q v0 k t l * R0) m -> post k t m l) + : cmd call (cmd.while e c) k t m l post. + Proof. + eapply wp_while. + eexists measure, lt, (fun v k t m l => exists R, (P v k t l * R) m /\ + forall K T L, Q v K T L * R ==> Q v0 K T L * R0). + split; [assumption|]. + split. { exists v0, R0. split; [assumption|]. intros. reflexivity. } + intros vi ki ti mi li (Ri&?&Qi). + destruct (Hbody _ _ _ _ _ _ ltac:(eassumption)) as (br&ki'&?&X); exists br, ki'; 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 kj 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/LeakageProgramLogic.v b/bedrock2/src/bedrock2/LeakageProgramLogic.v new file mode 100644 index 000000000..cfe1c9c1f --- /dev/null +++ b/bedrock2/src/bedrock2/LeakageProgramLogic.v @@ -0,0 +1,446 @@ +From coqutil.Tactics Require Import Tactics letexists eabstract rdelta reference_to_string ident_of_string. +Require Import coqutil.Map.Interface. +Require coqutil.Datatypes.ListSet. +Require Import bedrock2.Syntax bedrock2.LeakageSemantics. +Require Import bedrock2.LeakageWeakestPrecondition. +Require Import bedrock2.LeakageWeakestPreconditionProperties. +Require Import bedrock2.LeakageLoops. +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 replace callees by callees' to avoid duplicates *) + +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. + +Fixpoint callees' (c : Syntax.cmd) : list String.string := + match c with + | cmd.cond _ c1 c2 | cmd.seq c1 c2 => ListSet.list_union String.eqb (callees c1) (callees c2) + | cmd.while _ c | cmd.stackalloc _ _ c => callees' c + | cmd.call _ f _ => cons f nil + | _ => nil + end. + +(* returns a list of (caller, list_of_direct_callees) tuples *) +Definition callgraph: list (String.string * Syntax.func) -> + list (String.string * list String.string) := + List.map (fun '(fname, (_, _, fbody)) => (fname, callees' fbody)). + +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]. + +Ltac special_intro := + repeat match goal with + | |- ?P -> ?Q => + let P' := eval hnf in P in + match P' with + | exists _, _ => intros ?H'; destruct H' as [?f ?H](*important to do this now so that f is in the context of evars*); revert H + end + end; + match goal with + | |- forall (x : ?P), ?Q => intro + | |- _ => idtac (*avoid unfolding things to do intro*) + end. + +(* 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; + repeat (special_intro || match goal with | |- let _ := _ in _ => intro end); + intros; + 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 LeakageWeakestPrecondition. +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 + | |- Basics.impl _ _ => cbv [Basics.impl] (*why does swap break without this?*) + | |- program_logic_goal_for ?f _ => + enter f; + repeat match goal with + | |- exists _, _ => eexists + end; intros; + match goal with + | H: map.get ?functions ?fname = Some _ |- _ => + eapply start_func; [exact H | clear H] + end; + cbv match beta delta [LeakageWeakestPrecondition.func] + | |- LeakageWeakestPrecondition.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; eexists (*for reasons I do not comprehend, putting a letexists here causes the evar's context to be too small*); 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' _ _ _ (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] + | |- @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 ?P => + let get' := eval cbv [get] in @get in + change (get' w W L l x P); cbv beta + | |- load _ _ _ _ => cbv beta delta [load] + | |- @LeakageLoops.enforce ?width ?word ?locals ?names ?values ?map => + let values := eval cbv in values in + change (@LeakageLoops.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 => (*unsure whether still need this, or just case below*) + let x := fresh x in refine (let x := _ in ex_intro (fun x => P /\ Q) x _); + split; [solve [repeat straightline]|] + | |- exists x y, ?P /\ ?Q => + let x := fresh x in + refine (let x := _ in ex_intro (fun x => exists y, P /\ Q) x _); + eexists (*it is important, for some reason, that this is not a letexists*); split; [solve [repeat straightline] | ] + | |- exists x, Markers.split (?P /\ ?Q) => (*unsure whether still need this, or just case below*) + let x := fresh x in refine (let x := _ in ex_intro (fun x => P /\ Q) x _); + split; [solve [repeat straightline]|] + | |- exists x y, Markers.split (?P /\ ?Q) => + let x := fresh x in + refine (let x := _ in ex_intro (fun x => exists y, P /\ Q) x _); + eexists (*it is important, for some reason, that this is not a letexists*); split; [solve [repeat straightline] | ] + | |- Markers.unique (exists x, Markers.split (?P /\ ?Q)) => (*unsure whether we still need this, or just need case below*) + let x := fresh x in refine (let x := _ in ex_intro (fun x => P /\ Q) x _); + split; [solve [repeat straightline]|] + | |- Markers.unique (exists x y, Markers.split (?P /\ ?Q)) => + let x := fresh x in + refine (let x := _ in ex_intro (fun x => exists y, P /\ Q) x _); + eexists (*it is important, for some reason, that this is not a letexists*); 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 + | |- 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. + +(*for use when the spec of callee has been written using regular 'fnspec' notation*) +Ltac straightline_call_noex := + lazymatch goal with + | |- LeakageWeakestPrecondition.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 + repeat destruct Hcall as [?f Hcall]; + eapply LeakageWeakestPreconditionProperties.Proper_call; cycle -1; + [ eapply Hcall | try eabstract (solve [Morphisms.solve_proper]) .. ]; + [ .. | intros ? ? ? ?] + end. + +(*for use when the spec of callee has been written using 'fnspec exists' notation*) +Ltac straightline_call_ex := + lazymatch goal with + | |- call ?functions ?callee _ _ _ _ _ => + let Hcall := multimatch goal with + | H: context [ call functions callee _ _ _ _ _ ] |- _ => H + end in + eapply LeakageWeakestPreconditionProperties.Proper_call; cycle -1; + [ eapply Hcall | try eabstract solve [ Morphisms.solve_proper ].. ]; + [ .. | intros ? ? ? ? ] + end. + +(* TODO: once we can automatically prove some calls, include the success-only version of this in [straightline] *) +Ltac straightline_call := straightline_call_noex || straightline_call_ex. + +Ltac current_trace_mem_locals := + lazymatch goal with + | |- LeakageWeakestPrecondition.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 ?pick_sp ?E ?c ?F ?G ?H ?I => + let c' := eval cbv in c in + change (@cmd width BW word mem locals ext_spec pick_sp 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/LeakageSemantics.v b/bedrock2/src/bedrock2/LeakageSemantics.v new file mode 100644 index 000000000..b875c469b --- /dev/null +++ b/bedrock2/src/bedrock2/LeakageSemantics.v @@ -0,0 +1,409 @@ +Require Import coqutil.sanity coqutil.Byte. +Require Import coqutil.Tactics.fwd. +Require Import coqutil.Map.Properties. +Require coqutil.Map.SortedListString. +Require Import bedrock2.Syntax coqutil.Map.Interface coqutil.Map.OfListWord. +Require Import BinIntDef coqutil.Word.Interface coqutil.Word.Bitwidth. +Require Export bedrock2.Memory. +Require Import bedrock2.Semantics. +Require Import Coq.Lists.List. + +(*list lemmas and tactics that are useful for proving equality of leakage traces*) +Lemma align_trace_cons {T} x xs cont t (H : xs = List.app cont t) : @List.cons T x xs = List.app (cons x cont) t. +Proof. intros. cbn. congruence. Qed. + +Lemma align_trace_app {T} x xs cont t (H : xs = List.app cont t) : @List.app T x xs = List.app (List.app x cont) t. +Proof. intros. cbn. subst. rewrite List.app_assoc; trivial. Qed. + +Lemma align_trace_nil {A : Type} (l : list A) : l = nil ++ l. Proof. reflexivity. Qed. + +Ltac align_trace_bad := + repeat match goal with + | t:=_ :: _:_ |- _ => subst t + end; + repeat eapply align_trace_app || eapply align_trace_cons || eapply align_trace_nil || reflexivity. + +Goal exists (x y z : list nat), x ++ y = z. do 3 eexists. Fail Timeout 1 align_trace_bad. Abort. + +Ltac align_trace := + repeat match goal with + | t:= _ |- _ => subst t + end; + repeat exact eq_refl || eapply align_trace_app || eapply align_trace_cons || eapply align_trace_nil. + +Goal exists (x y z : list nat), x ++ y = z. do 3 eexists. align_trace. Abort. + +Lemma associate_one_left {A : Type} (x : A) l1 l2 : + l1 ++ x :: l2 = (l1 ++ (x :: nil)) ++ l2. +Proof. rewrite <- app_assoc. reflexivity. Qed. + +Ltac simpl_rev := repeat (match goal with + | |- context[rev (?a ++ ?b)] => rewrite (rev_app_distr a b) + end + || match goal with + | H: context[rev (?a ++ ?b)] |- _ => rewrite (rev_app_distr a b) in H + end + || rewrite rev_involutive in * || cbn [List.app List.rev] in * ). + +Inductive leakage_event {width: Z}{BW: Bitwidth width}{word: word.word width} : Type := +| leak_unit +| leak_bool (b : bool) +| leak_word (w : word) +| leak_list (l : list word). +(* ^sometimes it's convenient that one io call leaks only one event + See Interact case of spilling transform_trace function for an example. *) +Definition leakage {width: Z}{BW: Bitwidth width}{word: word.word width} : Type := + list leakage_event. + +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, function call results, and leakage trace, *) + (mem -> list word -> 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 := + { + (* Given a trace of previous interactions, the action name and arguments + uniquely determine what chunk of memory the ext_spec will chop off from + the whole memory m *) + mGive_unique: forall t m mGive1 mKeep1 mGive2 mKeep2 a args post1 post2, + map.split m mKeep1 mGive1 -> + map.split m mKeep2 mGive2 -> + ext_spec t mGive1 a args post1 -> + ext_spec t mGive2 a args post2 -> + mGive1 = mGive2; + + #[global] weaken :: forall t mGive act args, + Morphisms.Proper + (Morphisms.respectful + (Morphisms.pointwise_relation Interface.map.rep + (Morphisms.pointwise_relation (list word) + (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 -> 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 klist => + post1 mReceive resvals klist /\ post2 mReceive resvals klist); + }. +End ext_spec. +Arguments ext_spec.ok {_ _ _ _} _. + +Definition PickSp {width: Z}{BW: Bitwidth width}{word: word.word width} : Type := + leakage -> word. +Existing Class PickSp. + +Section binops. + Context {width : Z} {BW : Bitwidth width} {word : Word.Interface.word width}. + Definition leak_binop (bop : bopname) (x1 : word) (x2 : word) : leakage := + match bop with + | bopname.divu | bopname.remu => cons (leak_word x2) (cons (leak_word x1) nil) + | bopname.sru | bopname.slu | bopname.srs => cons (leak_word x2) nil + | _ => nil + end. +End binops. + +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}. + + Section WithMemAndLocals. + Context (m : mem) (l : locals). + + Local Notation "' x <- a ; f" := (match a with Some x => f | None => None end) + (right associativity, at level 70, x pattern). + + Fixpoint eval_expr (e : expr) (k : leakage) : option (word * leakage) := + match e with + | expr.literal v => Some (word.of_Z v, k) + | expr.var x => 'v <- map.get l x; Some (v, k) + | expr.inlinetable aSize t index => + '(index', k') <- eval_expr index k; + 'v <- load aSize (map.of_list_word t) index'; + Some (v, leak_word index' :: k') + | expr.load aSize a => + '(a', k') <- eval_expr a k; + 'v <- load aSize m a'; + Some (v, leak_word a' :: k') + | expr.op op e1 e2 => + '(v1, k') <- eval_expr e1 k; + '(v2, k'') <- eval_expr e2 k'; + Some (interp_binop op v1 v2, leak_binop op v1 v2 ++ k'') + | expr.ite c e1 e2 => + '(vc, k') <- eval_expr c k; + let b := word.eqb vc (word.of_Z 0) in + eval_expr (if b then e2 else e1) (leak_bool (negb b) :: k') + end. + + Fixpoint eval_call_args (arges : list expr) (k : leakage) := + match arges with + | e :: tl => + '(v, k') <- eval_expr e k; + '(args, k'') <- eval_call_args tl k'; + Some (v :: args, k'') + | _ => Some (nil, k) + 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} {pick_sp: PickSp}. + Section WithEnv. + Context (e: env). + + Implicit Types post : leakage -> trace -> mem -> locals -> Prop. (* COQBUG: unification finds Type instead of Prop and fails to downgrade *) + Inductive exec: cmd -> leakage -> trace -> mem -> locals -> + (leakage -> trace -> mem -> locals -> Prop) -> Prop := + | skip: forall k t m l post, + post k t m l -> + exec cmd.skip k t m l post + | set: forall x e k t m l post v k', + eval_expr m l e k = Some (v, k') -> + post k' t m (map.put l x v) -> + exec (cmd.set x e) k t m l post + | unset: forall x k t m l post, + post k t m (map.remove l x) -> + exec (cmd.unset x) k t m l post + | store: forall sz ea ev k t m l post a k' v k'' m', + eval_expr m l ea k = Some (a, k') -> + eval_expr m l ev k' = Some (v, k'') -> + store sz m a v = Some m' -> + post (leak_word a :: k'') t m' l -> + exec (cmd.store sz ea ev) k t m l post + | stackalloc: forall x n body k t mSmall l post, + Z.modulo n (bytes_per_word width) = 0 -> + (forall mStack mCombined, + let a := pick_sp k in + anybytes a n mStack -> + map.split mCombined mSmall mStack -> + exec body (leak_unit :: k) t mCombined (map.put l x a) + (fun k' t' mCombined' l' => + exists mSmall' mStack', + anybytes a n mStack' /\ + map.split mCombined' mSmall' mStack' /\ + post k' t' mSmall' l')) -> + exec (cmd.stackalloc x n body) k t mSmall l post + | if_true: forall k t m l e c1 c2 post v k', + eval_expr m l e k = Some (v, k') -> + word.unsigned v <> 0 -> + exec c1 (leak_bool true :: k') t m l post -> + exec (cmd.cond e c1 c2) k t m l post + | if_false: forall e c1 c2 k t m l post v k', + eval_expr m l e k = Some (v, k') -> + word.unsigned v = 0 -> + exec c2 (leak_bool false :: k') t m l post -> + exec (cmd.cond e c1 c2) k t m l post + | seq: forall c1 c2 k t m l post mid, + exec c1 k t m l mid -> + (forall k' t' m' l', mid k' t' m' l' -> exec c2 k' t' m' l' post) -> + exec (cmd.seq c1 c2) k t m l post + | while_false: forall e c k t m l post v k', + eval_expr m l e k = Some (v, k') -> + word.unsigned v = 0 -> + post (leak_bool false :: k') t m l -> + exec (cmd.while e c) k t m l post + | while_true: forall e c k t m l post v k' mid, + eval_expr m l e k = Some (v, k') -> + word.unsigned v <> 0 -> + exec c (leak_bool true :: k') t m l mid -> + (forall k'' t' m' l', mid k'' t' m' l' -> exec (cmd.while e c) k'' t' m' l' post) -> + exec (cmd.while e c) k t m l post + | call: forall binds fname arges k t m l post params rets fbody args k' lf mid, + map.get e fname = Some (params, rets, fbody) -> + eval_call_args m l arges k = Some (args, k') -> + map.of_list_zip params args = Some lf -> + exec fbody (leak_unit :: k') t m lf mid -> + (forall k'' t' m' st1, mid k'' t' m' st1 -> + exists retvs, map.getmany_of_list st1 rets = Some retvs /\ + exists l', map.putmany_of_list_zip binds retvs l = Some l' /\ + post k'' t' m' l') -> + exec (cmd.call binds fname arges) k t m l post + | interact: forall binds action arges args k' k t m l post mKeep mGive mid, + map.split m mKeep mGive -> + eval_call_args m l arges k = Some (args, k') -> + ext_spec t mGive action args mid -> + (forall mReceive resvals klist, mid mReceive resvals klist -> + exists l', map.putmany_of_list_zip binds resvals l = Some l' /\ + forall m', map.split m' mKeep mReceive -> + post (leak_list klist :: k') (cons ((mGive, action, args), (mReceive, resvals)) t) m' l') -> + exec (cmd.interact binds action arges) k t m l post. + + Context {word_ok: word.ok word} {mem_ok: map.ok mem} {ext_spec_ok: ext_spec.ok ext_spec}. + + Lemma interact_cps: forall binds action arges args k' k t m l post mKeep mGive, + map.split m mKeep mGive -> + eval_call_args m l arges k = Some (args, k') -> + ext_spec t mGive action args (fun mReceive resvals klist => + exists l', map.putmany_of_list_zip binds resvals l = Some l' /\ + forall m', map.split m' mKeep mReceive -> + post (leak_list klist :: k') (cons ((mGive, action, args), (mReceive, resvals)) t) m' l') -> + exec (cmd.interact binds action arges) k t m l post. + Proof. intros. eauto using interact. Qed. + + Lemma seq_cps: forall c1 c2 k t m l post, + exec c1 k t m l (fun k' t' m' l' => exec c2 k' t' m' l' post) -> + exec (cmd.seq c1 c2) k t m l post. + Proof. intros. eauto using seq. Qed. + + Lemma weaken: forall k t l m s post1, + exec s k t m l post1 -> + forall post2, + (forall k' t' m' l', post1 k' t' m' l' -> post2 k' t' m' l') -> + exec s k t m l 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 k t l m s post1, + exec s k t m l post1 -> + forall post2, + exec s k t m l post2 -> + exec s k t m l (fun k' t' m' l' => post1 k' t' m' l' /\ post2 k' t' m' l'). + 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, ?k1), H2: ?e = Some (?v2, ?k2) |- _ => + replace v2 with v1 in * by congruence; + replace k2 with k1 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.mGive_unique as P. + specialize P with (1 := H) (2 := H7) (3 := H1) (4 := H14). + subst mGive0. + destruct (map.split_diff (map.same_domain_refl mGive) H H7) as (? & _). + subst mKeep0. + 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 k t m l post, + exec e1 c k t m l post -> + exec e2 c k t m l 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} {pick_sp : PickSp}. + + Implicit Types (l: locals) (m: mem) (post: leakage -> trace -> mem -> list word -> Prop). + + Definition call e fname k t m args 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 k t m l (fun k' t' m' l' => exists rets, + map.getmany_of_list l' retnames = Some rets /\ post k' t' m' rets). + + Lemma weaken_call: forall e fname k t m args post1, + call e fname k t m args post1 -> + forall post2, (forall k' t' m' rets, post1 k' t' m' rets -> post2 k' t' m' rets) -> + call e fname k t m args 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 k t m rets post, + call e1 f k t m rets post -> + call e2 f k t m rets 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/LeakageWeakestPrecondition.v b/bedrock2/src/bedrock2/LeakageWeakestPrecondition.v new file mode 100644 index 000000000..4525bca32 --- /dev/null +++ b/bedrock2/src/bedrock2/LeakageWeakestPrecondition.v @@ -0,0 +1,486 @@ +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 bedrock2.LeakageSemantics. +Require Import Coq.Lists.List. + +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} {pick_sp: PickSp}. + Implicit Types (k : leakage) (t : trace) (m : mem) (l : locals). + + Definition literal v (post : word -> Prop) : Prop := + dlet! v := word.of_Z v in post v. + Definition get (l : locals) (x : String.string) (post : word -> Prop) : Prop := + exists v, map.get l x = Some v /\ post v. + Definition load s m a (post : _ -> Prop) : Prop := + exists v, load s m a = Some v /\ post v. + 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 (k : leakage) (e : Syntax.expr) (post : leakage -> word -> Prop) : Prop := + match e with + | expr.literal v => + literal v (post k) + | expr.var x => + get l x (post k) + | expr.op op e1 e2 => + rec k e1 (fun k' v1 => + rec k' e2 (fun k'' v2 => + post (leak_binop op v1 v2 ++ k'') (interp_binop op v1 v2))) + | expr.load s e => + rec k e (fun k' a => + load s m a (post (leak_word a :: k'))) + | expr.inlinetable s t e => + rec k e (fun k' a => + load s (map.of_list_word t) a (post (leak_word a :: k'))) + | expr.ite c e1 e2 => + rec k c (fun k' b => + let b := word.eqb b (word.of_Z 0) in + rec (leak_bool (negb b) :: k') (if b then e2 else e1) post) + end. + Fixpoint expr k e := expr_body expr k e. + End WithMemAndLocals. + + Section WithF. + Context {A B} (f: A -> (B -> Prop) -> Prop). + Definition list_map_body rec (xs : list A) (post : list B -> Prop) : Prop := + match xs with + | nil => post nil + | cons x xs' => + f x (fun y => + rec xs' (fun ys' => + post (cons y ys'))) + end. + Fixpoint list_map xs := list_map_body list_map xs. + End WithF. + + Section WithF'. + Context {A B T} (f: T -> A -> (T -> B -> Prop) -> Prop). + Definition list_map'_body rec (acc : T) (xs : list A) (post : T -> list B -> Prop) : Prop := + match xs with + | nil => post acc nil + | cons x xs' => + f acc x (fun acc' y => + rec acc' xs' (fun acc'' ys' => + post acc'' (cons y ys'))) + end. + Fixpoint list_map' acc xs := list_map'_body list_map' acc xs. + End WithF'. + + Section WithFunctions. + Context (e: env). + Definition dexpr m l k e v k' := expr m l k e (fun k'2 v2 => eq k' k'2 /\ eq v v2). + Definition dexprs m l k es vs k' := list_map' (expr m l) k es (fun k'2 vs2 => eq k' k'2 /\ eq vs vs2). + (* 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) (k : leakage) (t : trace) (m : mem) (l : locals) + (post : leakage -> trace -> mem -> locals -> Prop) : Prop := + (* give value of each pure expression when stating its subproof *) + match c with + | cmd.skip => post k t m l + | cmd.set x ev => + exists v k', dexpr m l k ev v k' /\ + dlet! l := map.put l x v in + post k' t m l + | cmd.unset x => + dlet! l := map.remove l x in + post k t m l + | cmd.store sz ea ev => + exists a k', dexpr m l k ea a k' /\ + exists v k'', dexpr m l k' ev v k'' /\ + store sz m a v (fun m => + post (leak_word a :: k'') t m l) + | cmd.stackalloc x n c => + Z.modulo n (bytes_per_word width) = 0 /\ + forall mStack mCombined, + let a := pick_sp k in + anybytes a n mStack -> map.split mCombined m mStack -> + dlet! l := map.put l x a in + rec c (leak_unit :: k) t mCombined l (fun k' t' mCombined' l' => + exists m' mStack', + anybytes a n mStack' /\ map.split mCombined' m' mStack' /\ + post k' t' m' l') + | cmd.cond br ct cf => + exists v k', dexpr m l k br v k' /\ + (word.unsigned v <> 0%Z -> rec ct (leak_bool true :: k') t m l post) /\ + (word.unsigned v = 0%Z -> rec cf (leak_bool false :: k') t m l post) + | cmd.seq c1 c2 => + rec c1 k t m l (fun k t m l => rec c2 k t m l post) + | cmd.while _ _ => LeakageSemantics.exec e c k t m l post + | cmd.call binds fname arges => + exists args k', dexprs m l k arges args k' /\ + LeakageSemantics.call e fname (leak_unit :: k') t m args (fun k'' t m rets => + exists l', map.putmany_of_list_zip binds rets l = Some l' /\ + post k'' t m l') + | cmd.interact binds action arges => + exists args k', dexprs m l k arges args k' /\ + exists mKeep mGive, map.split m mKeep mGive /\ + ext_spec t mGive action args (fun mReceive rets klist => + exists l', map.putmany_of_list_zip binds rets l = Some l' /\ + forall m', map.split m' mKeep mReceive -> + post (leak_list klist :: k') (cons ((mGive, action, args), (mReceive, rets)) t) m' l') + end. + Fixpoint cmd c := cmd_body cmd c. + End WithFunctions. + + Definition func call '(innames, outnames, c) (k : leakage) (t : trace) (m : mem) (args : list word) (post : leakage -> trace -> mem -> list word -> Prop) := + exists l, map.of_list_zip innames args = Some l /\ + cmd call c k t m l (fun k t m l => + list_map (get l) outnames (fun rets => + post k t m rets)). + + Definition program := cmd. +End WeakestPrecondition. +Notation call := LeakageSemantics.call (only parsing). + +Ltac unfold1_cmd e := + lazymatch e with + @cmd ?width ?BW ?word ?mem ?locals ?ext_spec ?pick_sp ?CA ?c ?k ?t ?m ?l ?post => + let c := eval hnf in c in + constr:(@cmd_body width BW word mem locals ext_spec pick_sp CA + (@cmd width BW word mem locals ext_spec pick_sp CA) c k t m l 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 ?BW ?word ?mem ?locals ?m ?l ?k ?arg ?post => + let arg := eval hnf in arg in + constr:(@expr_body width BW word mem locals m l (@expr width BW word mem locals m l) k arg 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 ?post => + let arg := eval hnf in arg in + constr:(@list_map_body A B P (@list_map A B P) arg 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_list_map' e := + lazymatch e with + @list_map' ?A ?B ?T ?P ?t ?arg ?post => + let arg := eval hnf in arg in + constr:(@list_map'_body A B T P (@list_map' A B T P) t arg 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. + +Import Coq.ZArith.ZArith. + +Notation "'fnspec!' name a0 .. an '/' g0 .. gn '~>' r0 .. rn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (forall a0, + .. (forall an, + (forall g0, + .. (forall gn, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem (cons a0 .. (cons an nil) ..) + (fun k' tr' mem' rets => + (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, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name a0 .. an '/' g0 .. gn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (forall a0, + .. (forall an, + (forall g0, + .. (forall gn, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem (cons a0 .. (cons an nil) ..) + (fun k' tr' mem' rets => + rets = nil /\ post))) ..)) ..)) + (at level 200, + name at level 0, + a0 binder, an binder, + g0 binder, gn binder, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name a0 .. an '~>' r0 .. rn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (forall a0, + .. (forall an, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem (cons a0 .. (cons an nil) ..) + (fun k' tr' mem' rets => + (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, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name '/' g0 .. gn '~>' r0 .. rn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (forall an, + (forall g0, + .. (forall gn, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem nil + (fun k' tr' mem' rets => + (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, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name a0 .. an ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (forall a0, + .. (forall an, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem (cons a0 .. (cons an nil) ..) + (fun k' tr' mem' rets => + rets = nil /\ post))) ..)) + (at level 200, + name at level 0, + a0 binder, an binder, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name '/' g0 .. gn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (forall g0, + .. (forall gn, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem nil + (fun k' tr' mem' rets => + rets = nil /\ post))) ..)) + (at level 200, + name at level 0, + g0 binder, gn binder, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' name '~>' r0 .. rn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem nil + (fun k' tr' mem' rets => + (exists r0, + .. (exists rn, + rets = (cons r0 .. (cons rn nil) ..) /\ + post) ..)))) + (at level 200, + name at level 0, + r0 closed binder, rn closed binder, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +(*with existent function(s)*) + +Notation "'fnspec!' 'exists' f0 .. fn ',' name a0 .. an '/' g0 .. gn '~>' r0 .. rn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (exists f0, + .. (exists fn, + (forall a0, + .. (forall an, + (forall g0, + .. (forall gn, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem (cons a0 .. (cons an nil) ..) + (fun k' tr' mem' rets => + (exists r0, + .. (exists rn, + rets = (cons r0 .. (cons rn nil) ..) /\ + post) ..)))) ..)) ..)) ..)) + (at level 200, + name at level 0, + f0 binder, fn binder, + a0 binder, an binder, + g0 binder, gn binder, + r0 closed binder, rn closed binder, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' 'exists' f0 .. fn ',' name a0 .. an '/' g0 .. gn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (exists f0, + .. (exists fn, + (forall a0, + .. (forall an, + (forall g0, + .. (forall gn, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem (cons a0 .. (cons an nil) ..) + (fun k' tr' mem' rets => + rets = nil /\ post))) ..)) ..)) ..)) + (at level 200, + name at level 0, + f0 binder, fn binder, + a0 binder, an binder, + g0 binder, gn binder, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' 'exists' f0 .. fn ',' name a0 .. an '~>' r0 .. rn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (exists f0, + .. (exists fn, + (forall a0, + .. (forall an, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem (cons a0 .. (cons an nil) ..) + (fun k' tr' mem' rets => + (exists r0, + .. (exists rn, + rets = (cons r0 .. (cons rn nil) ..) /\ + post) ..)))) ..)) ..)) + (at level 200, + name at level 0, + f0 binder, fn binder, + a0 binder, an binder, + r0 closed binder, rn closed binder, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' 'exists' f0 .. fn ',' name '/' g0 .. gn '~>' r0 .. rn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (exists f0, + .. (exists fn, + (forall an, + (forall g0, + .. (forall gn, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem nil + (fun k' tr' mem' rets => + (exists r0, + .. (exists rn, + rets = (cons r0 .. (cons rn nil) ..) /\ + post) ..)))) ..))) ..)) + (at level 200, + name at level 0, + f0 binder, fn binder, + g0 binder, gn binder, + r0 closed binder, rn closed binder, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' 'exists' f0 .. fn ',' name a0 .. an ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (exists f0, + .. (exists fn, + (forall a0, + .. (forall an, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem (cons a0 .. (cons an nil) ..) + (fun k' tr' mem' rets => + rets = nil /\ post))) ..)) ..)) + (at level 200, + name at level 0, + f0 binder, fn binder, + a0 binder, an binder, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' 'exists' f0 .. fn ',' name '/' g0 .. gn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (exists f0, + .. (exists fn, + (forall g0, + .. (forall gn, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem nil + (fun k' tr' mem' rets => + rets = nil /\ post))) ..)) ..)) + (at level 200, + name at level 0, + f0 binder, fn binder, + g0 binder, gn binder, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). + +Notation "'fnspec!' 'exists' f0 .. fn ',' name '~>' r0 .. rn ',' '{' 'requires' k tr mem := pre ';' 'ensures' k' tr' mem' ':=' post '}'" := + (fun functions => + (exists f0, + .. (exists fn, + (forall k tr mem, + pre -> + LeakageWeakestPrecondition.call + functions name k tr mem nil + (fun k' tr' mem' rets => + (exists r0, + .. (exists rn, + rets = (cons r0 .. (cons rn nil) ..) /\ + post) ..)))) ..)) + (at level 200, + name at level 0, + f0 binder, fn binder, + r0 closed binder, rn closed binder, + k name, k' name, tr name, tr' name, mem name, mem' name, + pre at level 200, + post at level 200). diff --git a/bedrock2/src/bedrock2/LeakageWeakestPreconditionProperties.v b/bedrock2/src/bedrock2/LeakageWeakestPreconditionProperties.v new file mode 100644 index 000000000..50630d975 --- /dev/null +++ b/bedrock2/src/bedrock2/LeakageWeakestPreconditionProperties.v @@ -0,0 +1,443 @@ +Require Import coqutil.Macros.subst coqutil.Macros.unique coqutil.Map.Interface coqutil.Word.Properties. +Require Import coqutil.Word.Bitwidth. +Require bedrock2.LeakageWeakestPrecondition. + +Require Import Coq.Classes.Morphisms. + +Section WeakestPrecondition. + 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: LeakageSemantics.ExtSpec} {pick_sp: LeakageSemantics.PickSp}. + + 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 pick_sp + | 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 _ Basics.impl) ==> Basics.impl)) LeakageWeakestPrecondition.literal. + Proof using. clear. cbv [LeakageWeakestPrecondition.literal]; cbv [Proper respectful pointwise_relation Basics.impl dlet.dlet]. eauto. Qed. + + Global Instance Proper_get : Proper (pointwise_relation _ (pointwise_relation _ ((pointwise_relation _ Basics.impl) ==> Basics.impl))) LeakageWeakestPrecondition.get. + Proof using. clear. cbv [LeakageWeakestPrecondition.get]; cbv [Proper respectful pointwise_relation Basics.impl]; intros * ? (?&?&?); eauto. Qed. + + Global Instance Proper_load : Proper (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ ((pointwise_relation _ Basics.impl) ==> Basics.impl)))) LeakageWeakestPrecondition.load. + Proof using. clear. cbv [LeakageWeakestPrecondition.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))))) LeakageWeakestPrecondition.store. + Proof using. clear. cbv [LeakageWeakestPrecondition.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 _ (pointwise_relation _ Basics.impl) ==> Basics.impl)))))) LeakageWeakestPrecondition.expr. + Proof using. + clear. + cbv [Proper respectful pointwise_relation Basics.impl LeakageSemantics.leak_binop]; ind_on Syntax.expr.expr; + cbn in *; intuition (try typeclasses eauto with core). + { eapply Proper_literal; eauto. cbv [pointwise_relation Basics.impl]. auto. } + { eapply Proper_get; eauto. cbv [pointwise_relation Basics.impl]. auto. } + { eapply IHa2; eauto; intuition idtac. eapply Proper_load; eauto using Proper_load. + cbv [pointwise_relation Basics.impl]. auto. } + { eapply IHa2; eauto; intuition idtac. eapply Proper_load; eauto using Proper_load. + cbv [pointwise_relation Basics.impl]. auto. } + { eapply IHa2_1; eauto; intuition idtac. + eapply IHa2_2; eauto; intuition idtac. eauto. } + { eapply IHa2_1; eauto; intuition idtac. + Tactics.destruct_one_match; eauto using Proper_load. } + Qed. + + Global Instance Proper_list_map {A B} : + Proper ((pointwise_relation _ (pointwise_relation _ Basics.impl ==> Basics.impl)) ==> pointwise_relation _ (pointwise_relation _ Basics.impl ==> Basics.impl)) (LeakageWeakestPrecondition.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). + Qed. + + Global Instance Proper_list_map' {A B T} : + Proper ((pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ Basics.impl) ==> Basics.impl)) ==> pointwise_relation _ (pointwise_relation _ (pointwise_relation _ (pointwise_relation _ Basics.impl) ==> Basics.impl)))) (LeakageWeakestPrecondition.list_map' (A:=A) (B:=B) (T:=T)). + Proof using. + clear. + cbv [Proper respectful pointwise_relation Basics.impl]; ind_on (list A); + cbn in *; intuition (try typeclasses eauto with core). + Qed. + + Context {word_ok : word.ok word} {mem_ok : map.ok mem}. + Context {locals_ok : map.ok locals}. + Context {ext_spec_ok : LeakageSemantics.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)))))))) LeakageWeakestPrecondition.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 eauto. + { destruct H1 as (?&?&?&?). eexists. eexists. split. + 1: eapply Proper_expr. + 1: cbv [pointwise_relation Basics.impl]; intuition eauto 2. + all: eauto. all: destruct H3; auto. } + { destruct H1 as (?&?&?&?). eexists. eexists. split. + { eapply Proper_expr; eauto; cbv [pointwise_relation Basics.impl]; eauto. } + { destruct H2 as (?&?&?&?). eexists. eexists. split. + { eapply Proper_expr; eauto; cbv [pointwise_relation Basics.impl]; 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; eauto; cbv [pointwise_relation Basics.impl]; eauto. } + { intuition eauto 6. } } + { eapply H4; eauto. simpl. intros. eauto. } + { eapply LeakageSemantics.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 LeakageSemantics.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 LeakageSemantics.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))))))))) LeakageWeakestPrecondition.call. + Proof using word_ok mem_ok locals_ok ext_spec_ok. + cbv [Proper respectful pointwise_relation Basics.impl]. + intros. eapply LeakageSemantics.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))))))) LeakageWeakestPrecondition.program. + Proof using word_ok mem_ok locals_ok ext_spec_ok. + cbv [Proper respectful pointwise_relation Basics.impl LeakageWeakestPrecondition.program]; intros. + eapply Proper_cmd; + cbv [Proper respectful pointwise_relation Basics.flip Basics.impl LeakageWeakestPrecondition.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 LeakageWeakestPrecondition.dexpr LeakageWeakestPrecondition.dexprs LeakageWeakestPrecondition.store] in * + end; eauto. + + Lemma expr_sound: forall m l e k post (H : LeakageWeakestPrecondition.expr m l k e post), + exists v k', LeakageSemantics.eval_expr m l e k = Some (v, k') /\ post k' v. + Proof using word_ok. + induction e; t. + { destruct H. destruct H. eexists. eexists. rewrite H. eauto. } + { eapply IHe in H; t. cbv [LeakageWeakestPrecondition.load] in H0; t. rewrite H. rewrite H0. eauto. } + { eapply IHe in H; t. cbv [LeakageWeakestPrecondition.load] in H0; t. rewrite H. rewrite H0. 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. + { eapply IHe3 in H0; t. } + { eapply IHe2 in H0; t. } } + Qed. + + Import ZArith coqutil.Tactics.Tactics. + + Lemma expr_complete: forall m l e k v k', + LeakageSemantics.eval_expr m l e k = Some (v, k') -> + LeakageWeakestPrecondition.dexpr m l k e v k'. + Proof using word_ok. + induction e; cbn; intros. + - inversion_clear H. split; reflexivity. + - destruct_one_match_hyp. 2: discriminate. inversion H. subst. + eexists. eauto. + - repeat (destruct_one_match_hyp; try discriminate; []). + inversion H. subst. eapply Proper_expr. + 2: { eapply IHe. eassumption. } + intros addr ? (?&?). subst. unfold LeakageWeakestPrecondition.load. eauto. + - repeat (destruct_one_match_hyp; try discriminate; []). + inversion H. subst. eapply Proper_expr. + 2: { eapply IHe. eassumption. } + intros addr ? (?&?). subst. unfold LeakageWeakestPrecondition.load. eauto. + - repeat (destruct_one_match_hyp; try discriminate; []). + inversion H. subst. eapply Proper_expr. + 2: { eapply IHe1. eassumption. } + intros v1 ? (?&?). subst. + eapply Proper_expr. + 2: { eapply IHe2. eassumption. } + intros v2 ? (?&?). subst. + auto. + - repeat (destruct_one_match_hyp; try discriminate; []). + inversion H. subst. eapply Proper_expr. + 2: { eapply IHe1. eassumption. } + intros vc ? (?&?). subst. + destr (word.eqb a (word.of_Z 0)). + + eapply IHe3. eassumption. + + eapply IHe2. eassumption. + Qed. + + Lemma sound_args : forall m l args k P, + LeakageWeakestPrecondition.list_map' (LeakageWeakestPrecondition.expr m l) k args P -> + exists x k', LeakageSemantics.eval_call_args m l args k = Some (x, k') /\ P k' x. + Proof using 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 P : + LeakageWeakestPrecondition.list_map (LeakageWeakestPrecondition.get l) a P + -> exists vs, map.getmany_of_list l a = Some vs /\ P vs. + Proof using. + cbv [map.getmany_of_list] in *. + revert P l; induction a; cbn; repeat (subst; t). + cbv [LeakageWeakestPrecondition.get] in H; t. + epose proof (IHa _ l _); clear IHa; t. + rewrite H. erewrite H1. eexists; split; eauto. exact H2. + Unshelve. + eapply Proper_list_map; try exact H0. + all : cbv [respectful pointwise_relation Basics.impl LeakageWeakestPrecondition.get]; intros; cbv beta; t. + Qed. + + Local Hint Constructors LeakageSemantics.exec : core. + Lemma sound_cmd e c k t m l post (H: LeakageWeakestPrecondition.cmd e c k t m l post) + : LeakageSemantics.exec e c k t m l post. + Proof. + ind_on Syntax.cmd; repeat (t; try match reverse goal with H : LeakageWeakestPrecondition.expr _ _ _ _ _ |- _ => eapply expr_sound in H end). + { destruct (BinInt.Z.eq_dec (Interface.word.unsigned x) (BinNums.Z0)) as [Hb|Hb]; cycle 1. + { econstructor; t. } + { eapply LeakageSemantics.exec.if_false; t. } } + { inversion H0. t. eapply sound_args in H; t. } + { eapply sound_args in H; t. } + Qed. + + Lemma weaken_cmd: forall e c k t m l (post1 post2: _->_->_->_->Prop), + LeakageWeakestPrecondition.cmd e c k t m l post1 -> + (forall k t m l, post1 k t m l -> post2 k t m l) -> + LeakageWeakestPrecondition.cmd e c k t m l post2. + 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 k vs k', + LeakageSemantics.eval_call_args m l args k = Some (vs, k') -> + LeakageWeakestPrecondition.dexprs m l k args vs k'. + Proof using word_ok. + induction args; cbn; repeat (subst; t). + 1: inversion H; auto. + destruct_one_match_hyp. 2: discriminate. + do 2 destruct_one_match_hyp. 2: discriminate. + destruct_one_match_hyp. + inversion H. subst. clear H. + eapply Proper_expr. 2: eapply expr_complete. 2: eassumption. + intros x ? (?&?). subst. + eapply Proper_list_map'. 3: eapply IHargs; eassumption. + { eapply Proper_expr. } + { intros ? ? (?&?). subst. split; reflexivity. } + Qed. + + Lemma complete_cmd: forall e c k t m l post, + LeakageSemantics.exec e c k t m l post -> + LeakageWeakestPrecondition.cmd e c k t m l post. + Proof. + induction 1. + { eassumption. } + { eapply expr_complete in H. eexists. eexists. split. 1: exact H. + eassumption. } + { eauto. } + { eapply expr_complete in H. + eapply expr_complete in H0. + eexists. eexists. split. 1: eassumption. + eexists. 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. eexists. ssplit; intros; eauto using expr_complete; congruence. } + { eexists. eexists. ssplit; intros; eauto using expr_complete; congruence. } + { cbn. eapply weaken_cmd. + { eapply IHexec. } + cbv beta. intros. + eapply H1. eassumption. } + { cbn. eapply LeakageSemantics.exec.while_false; eauto. } + { rename IHexec into IH1, H3 into IH2. + cbn. eapply LeakageSemantics.exec.while_true; eassumption. } + { cbn. eexists. eexists. split. + { eapply complete_args. eassumption. } + unfold LeakageSemantics.call. do 4 eexists. 1: eassumption. do 2 eexists. 1: eassumption. + eapply LeakageSemantics.exec.weaken. + { eassumption. } + cbv beta. intros. + specialize H3 with (1 := H4). destruct H3 as (retvs & G & ? & ? & ?). eauto 8. } + { cbn. eexists. eexists. split. + { eapply complete_args. eassumption. } + eexists _, _. split. 1: eassumption. + eapply LeakageSemantics.ext_spec.weaken. 2: eassumption. + intros m0 args0 k0 Hmid. specialize H2 with (1 := Hmid). destruct H2 as (? & ? & ?). + eauto 8. } + Qed. + + Lemma start_func: forall e fname fimpl k t m args post, + map.get e fname = Some fimpl -> + LeakageWeakestPrecondition.func e fimpl k t m args post -> + LeakageWeakestPrecondition.call e fname k t m args post. + 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. eapply sound_getmany. assumption. + Qed. + + (** Ad-hoc lemmas here? *) + + Import bedrock2.Syntax bedrock2.LeakageSemantics bedrock2.LeakageWeakestPrecondition. + Lemma interact_nomem call action binds arges k t m l post + args k' (Hargs : dexprs m l k arges args k') + (Hext : ext_spec t map.empty binds args (fun mReceive (rets : list word) klist => + mReceive = map.empty /\ + exists l0 : locals, map.putmany_of_list_zip action rets l = Some l0 /\ + post (leak_list klist :: k')%list (cons (map.empty, binds, args, (map.empty, rets)) t) m l0)) + : LeakageWeakestPrecondition.cmd call (cmd.interact action binds arges) k t m l post. + Proof using word_ok mem_ok ext_spec_ok. + exists args, k'; 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 k (post1 post2: leakage -> word -> Prop), + LeakageWeakestPrecondition.expr m l k e post1 -> + LeakageWeakestPrecondition.expr m l k e post2 -> + LeakageWeakestPrecondition.expr m l k e (fun k v => post1 k v /\ post2 k v). + Proof using word_ok. + induction e; cbn; unfold literal, dlet.dlet, LeakageWeakestPrecondition.get; intros. + - eauto. + - decompose [and ex] H. decompose [and ex] H0. assert (x0 = x1) by congruence. subst. eauto. + - eapply Proper_expr. + 2: eapply IHe. + 2: eapply H. + 2: eapply H0. + unfold Morphisms.pointwise_relation, Basics.impl. + unfold load. intros. decompose [and ex] H1. assert (x0 = x) by congruence. subst. eauto. + - eapply Proper_expr. + 2: eapply IHe. + 2: eapply H. + 2: eapply H0. + unfold Morphisms.pointwise_relation, Basics.impl. + unfold load. intros. decompose [and ex] H1. assert (x0 = x) by congruence. subst. eauto. + - eapply Proper_expr. + 2: eapply IHe1. + 2: eapply H. + 2: eapply H0. + unfold Morphisms.pointwise_relation, Basics.impl. + unfold load. intros. decompose [and ex] H1. + eapply IHe2; eassumption. + - eapply Proper_expr. + 2: eapply IHe1. + 2: eapply H. + 2: eapply H0. + unfold Morphisms.pointwise_relation, Basics.impl. + intros ? ? [? ?]. Tactics.destruct_one_match; eauto using Proper_expr. + Qed. + + Lemma dexpr_expr (m : mem) l e k P + (H : LeakageWeakestPrecondition.expr m l k e P) + : exists v k', LeakageWeakestPrecondition.dexpr m l k e v k' /\ P k' v. + Proof using word_ok. + generalize dependent P; revert k; induction e; cbn. + { cbv [LeakageWeakestPrecondition.literal dlet.dlet]; cbn; eauto. } + { cbv [LeakageWeakestPrecondition.get]; intros ? ? (?&?&?). eauto 7. } + { intros k P H; case (IHe _ _ H) as (?&?&?&?&?&?); clear IHe H. + cbv [LeakageWeakestPrecondition.dexpr] in *. + eexists; eexists; split; [|eassumption]. + eapply Proper_expr; [|eauto]. + intros ? ? (?&?); subst. + eexists; eauto. } + { intros k P H; case (IHe _ _ H) as (?&?&?&?&?&?); clear IHe H. + cbv [LeakageWeakestPrecondition.dexpr] in *. + eexists; eexists; split; [|eassumption]. + eapply Proper_expr; [|eauto]. + intros ? ? (?&?); subst. + eexists; eauto. } + { intros k P H. + case (IHe1 _ _ H) as (?&?&?&H'); case (IHe2 _ _ H') as (?&?&?&?); + clear IHe1 IHe2 H H'. + cbv [LeakageWeakestPrecondition.dexpr] in *. + eexists; eexists; split; [|eassumption]. + eapply Proper_expr; [|eauto]; intros ? ? []. subst. + eapply Proper_expr; [|eauto]; intros ? ? []. subst. + auto. + } + { intros k P H. + case (IHe1 _ _ H) as (?&?&?&H'). Tactics.destruct_one_match_hyp. + { case (IHe3 _ _ H') as (?&?&?&?). + clear IHe1 IHe2 H H'. + cbv [LeakageWeakestPrecondition.dexpr] in *. + eexists; eexists; split; [|eassumption]. + eapply Proper_expr; [|eauto]; intros ? ? []. subst. + rewrite word.eqb_eq by reflexivity. assumption. } + { case (IHe2 _ _ H') as (?&?&?&?). + clear IHe1 IHe3 H H'. + cbv [LeakageWeakestPrecondition.dexpr] in *. + eexists; eexists; split; [|eassumption]. + eapply Proper_expr; [|eauto]; intros ? ? []. subst. + Tactics.destruct_one_match. 1: contradiction. assumption. } } + Qed. +End WeakestPrecondition. diff --git a/bedrock2/src/bedrock2/MetricLeakageSemantics.v b/bedrock2/src/bedrock2/MetricLeakageSemantics.v new file mode 100644 index 000000000..adfb2c376 --- /dev/null +++ b/bedrock2/src/bedrock2/MetricLeakageSemantics.v @@ -0,0 +1,718 @@ +Require Import coqutil.sanity coqutil.Byte. +Require Import coqutil.Tactics.fwd. +Require Import coqutil.Map.Properties. +Require coqutil.Map.SortedListString. +Require Import coqutil.Z.Lia. +Require Import bedrock2.Syntax coqutil.Map.Interface coqutil.Map.OfListWord. +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.MetricSemantics. +Require Import bedrock2.Semantics. +Require Import bedrock2.LeakageSemantics. +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}. + Context {ext_spec: ExtSpec}. + + Local Notation metrics := MetricLog. + + Section WithMemAndLocals. + Context (m : mem) (l : locals). + + Local Notation "' x <- a ; f" := (match a with Some x => f | None => None 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) (k : leakage) (mc : metrics) : option (word * leakage * metrics) := + match e with + | expr.literal v => Some (word.of_Z v, k, cost_lit isRegStr UNK mc) + | expr.var x => 'v <- map.get l x; Some (v, k, cost_set isRegStr UNK x mc) + | expr.inlinetable aSize t index => + '(index', k', mc') <- eval_expr index k mc; + 'v <- load aSize (map.of_list_word t) index'; + Some (v, leak_word index' :: k', cost_inlinetable isRegStr UNK UNK mc') + | expr.load aSize a => + '(a', k', mc') <- eval_expr a k mc; + 'v <- load aSize m a'; + Some (v, leak_word a' :: k', cost_load isRegStr UNK UNK mc') + | expr.op op e1 e2 => + '(v1, k', mc') <- eval_expr e1 k mc; + '(v2, k'', mc'') <- eval_expr e2 k' mc'; + Some (interp_binop op v1 v2, leak_binop op v1 v2 ++ k'', cost_op isRegStr UNK UNK UNK mc'') + | expr.ite c e1 e2 => + '(vc, k', mc') <- eval_expr c k mc; + let b := word.eqb vc (word.of_Z 0) in + eval_expr (if b then e2 else e1) (leak_bool (negb b) :: k') + (cost_if isRegStr UNK (Some UNK) mc') + end. + + Fixpoint eval_call_args (arges : list expr) (k : leakage) (mc : metrics) := + match arges with + | e :: tl => + '(v, k', mc') <- eval_expr e k mc; + '(args, k'', mc'') <- eval_call_args tl k' mc'; + Some (v :: args, k'', mc'') + | _ => Some (nil, k, 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 : leakage -> trace -> mem -> locals -> metrics -> Prop. (* COQBUG: unification finds Type instead of Prop and fails to downgrade *) + Inductive exec {pick_sp: PickSp} : + cmd -> leakage -> trace -> mem -> locals -> metrics -> + (leakage -> trace -> mem -> locals -> metrics -> Prop) -> Prop := + | skip + k t m l mc post + (_ : post k t m l mc) + : exec cmd.skip k t m l mc post + | set x e + k t m l mc post + v k' mc' (_ : eval_expr m l e k mc = Some (v, k', mc')) + (_ : post k' t m (map.put l x v) (cost_set isRegStr x UNK mc')) + : exec (cmd.set x e) k t m l mc post + | unset x + k t m l mc post + (_ : post k t m (map.remove l x) mc) + : exec (cmd.unset x) k t m l mc post + | store sz ea ev + k t m l mc post + a k' mc' (_ : eval_expr m l ea k mc = Some (a, k', mc')) + v k'' mc'' (_ : eval_expr m l ev k' mc' = Some (v, k'', mc'')) + m' (_ : store sz m a v = Some m') + (_ : post (leak_word a :: k'') t m' l (cost_store isRegStr UNK UNK mc'')) + : exec (cmd.store sz ea ev) k t m l mc post + | stackalloc x n body + k t mSmall l mc post + (_ : Z.modulo n (bytes_per_word width) = 0) + (_ : forall mStack mCombined, + let a := pick_sp k in + anybytes a n mStack -> + map.split mCombined mSmall mStack -> + exec body (leak_unit :: k) t mCombined (map.put l x a) (cost_stackalloc isRegStr x mc) + (fun k' t' mCombined' l' mc' => + exists mSmall' mStack', + anybytes a n mStack' /\ + map.split mCombined' mSmall' mStack' /\ + post k' t' mSmall' l' mc')) + : exec (cmd.stackalloc x n body) k t mSmall l mc post + | if_true k t m l mc e c1 c2 post + v k' mc' (_ : eval_expr m l e k mc = Some (v, k', mc')) + (_ : word.unsigned v <> 0) + (_ : exec c1 (leak_bool true :: k') t m l (cost_if isRegStr UNK (Some UNK) mc') post) + : exec (cmd.cond e c1 c2) k t m l mc post + | if_false e c1 c2 + k t m l mc post + v k' mc' (_ : eval_expr m l e k mc = Some (v, k', mc')) + (_ : word.unsigned v = 0) + (_ : exec c2 (leak_bool false :: k') t m l (cost_if isRegStr UNK (Some UNK) mc') post) + : exec (cmd.cond e c1 c2) k t m l mc post + | seq c1 c2 + k t m l mc post + mid (_ : exec c1 k t m l mc mid) + (_ : forall k' t' m' l' mc', mid k' t' m' l' mc' -> exec c2 k' t' m' l' mc' post) + : exec (cmd.seq c1 c2) k t m l mc post + | while_false e c + k t m l mc post + v k' mc' (_ : eval_expr m l e k mc = Some (v, k', mc')) + (_ : word.unsigned v = 0) + (_ : post (leak_bool false :: k') t m l (cost_loop_false isRegStr UNK (Some UNK) mc')) + : exec (cmd.while e c) k t m l mc post + | while_true e c + k t m l mc post + v k' mc' (_ : eval_expr m l e k mc = Some (v, k', mc')) + (_ : word.unsigned v <> 0) + mid (_ : exec c (leak_bool true :: k') t m l mc' mid) + (_ : forall k'' t' m' l' mc'', mid k'' t' m' l' mc'' -> + exec (cmd.while e c) k'' t' m' l' (cost_loop_true isRegStr UNK (Some UNK) mc'') post) + : exec (cmd.while e c) k t m l mc post + | call binds fname arges + k t m l mc post + params rets fbody (_ : map.get e fname = Some (params, rets, fbody)) + args k' mc' (_ : eval_call_args m l arges k mc = Some (args, k', mc')) + lf (_ : map.of_list_zip params args = Some lf) + mid (_ : exec fbody (leak_unit :: k') t m lf mc' mid) + (_ : forall k'' t' m' st1 mc'', mid k'' 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 k'' t' m' l' (cost_call PreSpill mc'')) + : exec (cmd.call binds fname arges) k t m l mc post + | interact binds action arges + k t m l mc post + mKeep mGive (_: map.split m mKeep mGive) + args k' mc' (_ : eval_call_args m l arges k mc = Some (args, k', mc')) + mid (_ : ext_spec t mGive action args mid) + (_ : forall mReceive resvals klist, mid mReceive resvals klist -> + exists l', map.putmany_of_list_zip binds resvals l = Some l' /\ + forall m', map.split m' mKeep mReceive -> + post (leak_list klist :: k') (cons ((mGive, action, args), (mReceive, resvals)) t) m' l' + (cost_interact PreSpill mc')) + : exec (cmd.interact binds action arges) k 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 {pick_sp: PickSp} : forall s k t m l mc post1, + exec s k t m l mc post1 -> + forall post2: _ -> _ -> _ -> _ -> _ -> Prop, + (forall k' t' m' l' mc', post1 k' t' m' l' mc' -> post2 k' t' m' l' mc') -> + exec s k 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 {pick_sp: PickSp} : forall k t l m mc s post1, + exec s k t m l mc post1 -> + forall post2, + exec s k t m l mc post2 -> + exec s k t m l mc (fun k' t' m' l' mc' => post1 k' t' m' l' mc' /\ post2 k' t' m' l' mc'). + Proof. + induction 1; + intros; + match goal with + | H: exec _ _ _ _ _ _ _ |- _ => inversion H; subst; clear H + end; + repeat 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, 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, H13 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 H17. (* not H2 *) + + simpl. intros *. intros [? ?]. + edestruct H3 as (? & ? & ? & ? & ?); [eassumption|]. + edestruct H18 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.mGive_unique as P. + specialize P with (1 := H) (2 := H7) (3 := H1) (4 := H15). + subst mGive0. + destruct (map.split_diff (map.same_domain_refl mGive) H H7) as (? & _). + subst mKeep0. + eapply interact. 1,2: eassumption. + + eapply ext_spec.intersect; [ exact H1 | exact H15 ]. + + simpl. intros *. intros [? ?]. + edestruct H2 as (? & ? & ?); [eassumption|]. + edestruct H16 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. + + Lemma eval_expr_extends_trace : + forall e0 m l mc k v mc' k', + eval_expr m l e0 k mc = Some (v, k', mc') -> + exists k'', k' = k'' ++ k. + Proof. + intros e0. induction e0; intros; simpl in *; + repeat match goal with + | H: (let (_, _) := ?x in _) = _ |- _ => + destruct x + | H: match ?x with + | Some _ => _ + | None => _ + end = Some (_, _, _) |- _ => + destruct x eqn:?; try congruence + | H: Some (?v1, ?mc1, ?t1) = Some (?v2, ?mc2, ?t2) |- _ => + injection H; intros; subst + end. + - eexists. align_trace. + - eexists. align_trace. + - specialize IHe0 with (1 := Heqo). fwd. eexists. align_trace. + - specialize IHe0 with (1 := Heqo). fwd. eexists. align_trace. + - specialize IHe0_1 with (1 := Heqo). specialize IHe0_2 with (1 := Heqo0). fwd. + eexists. align_trace. + - specialize IHe0_1 with (1 := Heqo). destruct (word.eqb _ _). + + specialize IHe0_3 with (1 := H). fwd. eexists. align_trace. + + specialize IHe0_2 with (1 := H). fwd. eexists. align_trace. + Qed. + + Lemma eval_call_args_extends_trace : + forall arges m l mc k args mc' k', + eval_call_args m l arges k mc = Some (args, k', mc') -> + exists k'', k' = k'' ++ k. + Proof. + intros arges. induction arges. + - simpl. intros. injection H. intros. subst. eexists. align_trace. + - simpl. intros. destruct (eval_expr _ _ _ _ _) eqn:E1; try congruence. + destruct p. destruct p. destruct (eval_call_args _ _ _ _ _) eqn:E2; try congruence. + destruct p. destruct p. injection H. intros. subst. + apply eval_expr_extends_trace in E1. specialize IHarges with (1 := E2). + fwd. eexists. align_trace. + Qed. + + Lemma expr_to_other_trace m l a mc mc' k1 k1' v : + eval_expr m l a k1 mc = Some (v, k1', mc') -> + exists k'', + k1' = k'' ++ k1 /\ + forall k2, + eval_expr m l a k2 mc = Some (v, k'' ++ k2, mc'). + Proof. + revert v. revert mc. revert k1. revert k1'. revert mc'. clear. + induction a; intros ? ? ? ? ? H; simpl in H; try (inversion H; subst; clear H). + - exists nil. auto. + - destruct (map.get l x) as [v0|] eqn:E; [|congruence]. inversion H1; subst; clear H1. + exists nil. simpl. rewrite E. auto. + - destruct (eval_expr _ _ _) as [v0|] eqn:E1; [|congruence]. + destruct v0. destruct p. destruct (load _ _ _) as [v0|] eqn:E2; [|congruence]. + inversion H1; subst; clear H1. eapply IHa in E1. destruct E1 as [k'' [E1 E3] ]. subst. + eexists (_ :: _). intuition. simpl. rewrite E3. rewrite E2. reflexivity. + - destruct (eval_expr _ _ _) as [v0|] eqn:E1; [|congruence]. + destruct v0. destruct p. destruct (load _ _ _) as [v0|] eqn:E2; [|congruence]. + inversion H1; subst; clear H1. eapply IHa in E1. destruct E1 as [k'' [E1 E3] ]. subst. + eexists (_ :: _). intuition. simpl. rewrite E3. rewrite E2. reflexivity. + - destruct (eval_expr _ _ a1 _ _) as [ [ [v0 mc0] p0]|] eqn:E1; [|congruence]. + destruct (eval_expr _ _ a2 _ _) as [ [ [v1 mc1] p1]|] eqn:E2; [|congruence]. + inversion H1; subst; clear H1. + eapply IHa1 in E1. destruct E1 as [k''1 [H1 H2] ]. eapply IHa2 in E2. + destruct E2 as [k''2 [H3 H4] ]. subst. + eexists (_ ++ _ ++ _). repeat rewrite <- (app_assoc _ _ k1). intuition. + simpl. rewrite H2. rewrite H4. f_equal. f_equal. repeat rewrite <- (app_assoc _ _ k2). + reflexivity. + - destruct (eval_expr _ _ a1 _ _) as [ [ [v0 mc0] p0]|] eqn:E1; [|congruence]. + eapply IHa1 in E1. destruct E1 as [k''1 [H2 H3] ]. subst. simpl. + destruct (word.eqb _ _) eqn:E. + + eapply IHa3 in H1. destruct H1 as [k''3 [H1 H2] ]. subst. + eexists (_ ++ _ :: _). repeat rewrite <- (app_assoc _ _ k1). + intuition. rewrite H3. rewrite E. rewrite H2. + repeat rewrite <- (app_assoc _ _ k2). reflexivity. + + eapply IHa2 in H1. destruct H1 as [k''2 [H1 H2] ]. subst. + eexists (_ ++ _ :: _). repeat rewrite <- (app_assoc _ _ k1). + intuition. rewrite H3. rewrite E. rewrite H2. + repeat rewrite <- (app_assoc _ _ k2). reflexivity. + Qed. + + Lemma call_args_to_other_trace (m : mem) l arges mc k1 vs mc' k1' : + eval_call_args m l arges k1 mc = Some (vs, k1', mc') -> + exists k'', + k1' = k'' ++ k1 /\ + forall k2, + eval_call_args m l arges k2 mc = Some (vs, k'' ++ k2, mc'). + Proof. + revert mc. revert k1. revert vs. revert mc'. revert k1'. induction arges; intros. + - cbn [eval_call_args] in H. inversion H. subst. + exists nil. intuition. + - cbn [eval_call_args] in *. + destruct (eval_expr _ _ _) as [ [ [v0 mc0] p0]|] eqn:E1; [|congruence]. + destruct (eval_call_args _ _ _) as [ [ [v1 mc1] p1]|] eqn:E2; [|congruence]. + apply expr_to_other_trace in E1. apply IHarges in E2. fwd. + eexists (_ ++ _). + repeat rewrite <- (app_assoc _ _ k1). intuition. repeat rewrite <- (app_assoc _ _ k2). + rewrite E1p1. rewrite E2p1. reflexivity. + Qed. + + Local Ltac subst_exprs := + repeat match goal with + | H : eval_expr _ _ _ _ _ = Some _ |- _ => + apply eval_expr_extends_trace in H; destruct H; subst + | H : eval_call_args _ _ _ _ _ = Some _ |- _ => + apply eval_call_args_extends_trace in H; destruct H; subst + end. + + Lemma exec_extends_trace {pick_sp: PickSp} s k t m l mc post : + exec s k t m l mc post -> + exec s k t m l mc (fun k' t' m' l' mc' => post k' t' m' l' mc' /\ exists k'', k' = k'' ++ k). + Proof. + intros H. induction H; try (econstructor; intuition eauto; subst_exprs; eexists; align_trace; fail). + - econstructor; intuition eauto. intros. eapply weaken. 1: eapply H1; eauto. + simpl. intros. fwd. eexists. eexists. intuition eauto. eexists. align_trace. + - eapply if_true; intuition eauto. eapply weaken. 1: eapply IHexec. + simpl. intros. fwd. intuition eauto. subst_exprs. eexists. align_trace. + - eapply if_false; intuition eauto. eapply weaken. 1: eapply IHexec. + simpl. intros. fwd. intuition eauto. subst_exprs. eexists. align_trace. + - econstructor; intuition eauto. fwd. eapply weaken. 1: eapply H1; eauto. + simpl. intros. fwd. intuition eauto. eexists. align_trace. + - eapply while_true; eauto. simpl. intros. fwd. eapply weaken. 1: eapply H3; eauto. + simpl. intros. fwd. intuition eauto. subst_exprs. eexists. align_trace. + - econstructor; intuition eauto. fwd. specialize H3 with (1 := H4p0). fwd. + eexists. intuition eauto. eexists. intuition eauto. subst_exprs. + eexists. align_trace. + - econstructor; intuition eauto. specialize H2 with (1 := H3). fwd. + eexists. intuition eauto. subst_exprs. eexists. align_trace. + Qed. + + Lemma exec_ext (pick_sp1: PickSp) s k t m l mc post : + exec (pick_sp := pick_sp1) s k t m l mc post -> + forall pick_sp2, + (forall k', pick_sp1 (k' ++ k) = pick_sp2 (k' ++ k)) -> + exec (pick_sp := pick_sp2) s k t m l mc post. + Proof. + intros H1 pick_sp2. induction H1; intros; try solve [econstructor; eauto]. + - econstructor; eauto. intros. replace (pick_sp1 k) with (pick_sp2 k) in *. + { subst a. eapply weaken. + { eapply H1; eauto. + intros. eassert (H2' := H2 (_ ++ _ :: nil)). rewrite <- app_assoc in H2'. eapply H2'. } + eauto. } + symmetry. apply H2 with (k' := nil). + - eapply if_true; eauto. eapply IHexec. subst_exprs. + intros. eassert (H2' := H2 (_ ++ _ :: _)). rewrite <- app_assoc in H2'. eapply H2'. + - eapply if_false; eauto. eapply IHexec. subst_exprs. + intros. eassert (H2' := H2 (_ ++ _ :: _)). rewrite <- app_assoc in H2'. eapply H2'. + - econstructor. 1: eapply exec_extends_trace; eauto. simpl. intros. fwd. + eapply H0; eauto. intros. repeat rewrite app_assoc. apply H2. + - eapply while_true; intuition eauto. + { eapply exec_extends_trace. eapply IHexec. subst_exprs. + intros. simpl. rewrite associate_one_left. rewrite app_assoc. apply H4. } + simpl in *. fwd. eapply H3; eauto. intros. subst_exprs. + rewrite associate_one_left. repeat rewrite app_assoc. auto. + - econstructor. 4: eapply exec_extends_trace. all: intuition eauto. + { eapply IHexec. subst_exprs. intros. + rewrite associate_one_left. repeat rewrite app_assoc. auto. } + fwd. specialize H3 with (1 := H5p0). fwd. intuition eauto. + Qed. + + Local Ltac solve_picksps_equal := + intros; cbv beta; f_equal; + repeat (rewrite rev_app_distr || cbn [rev app]); rewrite List.skipn_app_r; + [|repeat (rewrite app_length || rewrite rev_length || simpl); blia]; + repeat rewrite <- app_assoc; rewrite List.skipn_app_r; + [|rewrite rev_length; reflexivity]; + repeat (rewrite rev_app_distr || cbn [rev app] || rewrite rev_involutive); + repeat rewrite <- app_assoc; reflexivity. + + Lemma exec_to_other_trace (pick_sp: PickSp) s k1 k2 t m l mc post : + exec s k1 t m l mc post -> + exec (pick_sp := fun k => pick_sp (rev (skipn (length k2) (rev k)) ++ k1)) + s k2 t m l mc (fun k2' t' m' l' mc' => + exists k'', + k2' = k'' ++ k2 /\ + post (k'' ++ k1) t' m' l' mc'). + Proof. + intros H. generalize dependent k2. induction H; intros. + - econstructor. exists nil. eauto. + - apply expr_to_other_trace in H. destruct H as [k'' [H1 H2] ]. subst. + econstructor; intuition eauto. + - econstructor; intuition. exists nil. intuition. + - apply expr_to_other_trace in H. apply expr_to_other_trace in H0. + destruct H as [k''a [H3 H4] ]. subst. destruct H0 as [k''v [H5 H6] ]. subst. + econstructor; intuition eauto. eexists (_ :: _ ++ _). simpl. + repeat rewrite <- (app_assoc _ _ k2). repeat rewrite <- (app_assoc _ _ k). + intuition. + - econstructor; intuition. intros. + replace (rev k2) with (rev k2 ++ nil) in * by apply app_nil_r. Search (length (rev _)). + rewrite List.skipn_app_r in * by (rewrite rev_length; reflexivity). + simpl in *. eapply weaken. + { eapply exec_ext with (pick_sp1 := _). 1: eapply H1; eauto. solve_picksps_equal. } + simpl. intros. fwd. eexists _, _. intuition eauto. eexists (_ ++ _ :: nil). + rewrite <- app_assoc. simpl. rewrite <- (app_assoc _ _ k). simpl. eauto. + - apply expr_to_other_trace in H. fwd. eapply if_true; intuition eauto. + eapply weaken. + { eapply exec_ext with (pick_sp1 := _). 1: eapply IHexec. solve_picksps_equal. } + simpl. intros. fwd. eexists (_ ++ _ :: _). + repeat rewrite <- (app_assoc _ _ k2). repeat rewrite <- (app_assoc _ _ k). + intuition. + - apply expr_to_other_trace in H. fwd. eapply if_false; intuition. + eapply weaken. + { eapply exec_ext with (pick_sp1 := _). 1: eapply IHexec. solve_picksps_equal. } + simpl. intros. fwd. eexists (_ ++ _ :: _). + repeat rewrite <- (app_assoc _ _ k2). repeat rewrite <- (app_assoc _ _ k). + intuition. + - econstructor; intuition. fwd. eapply weaken. + { eapply exec_ext with (pick_sp1 := _). 1: eapply H1; eauto. solve_picksps_equal. } + simpl. intros. fwd. eexists (_ ++ _). + repeat rewrite <- (app_assoc _ _ k2). repeat rewrite <- (app_assoc _ _ k). + intuition. + - apply expr_to_other_trace in H. fwd. eapply while_false; intuition. + eexists (_ :: _). intuition. + - apply expr_to_other_trace in H. fwd. eapply while_true; intuition. + { eapply exec_ext with (pick_sp1 := _). 1: eapply IHexec. solve_picksps_equal. } + cbv beta in *. fwd. eapply weaken. + { eapply exec_ext with (pick_sp1 := _). 1: eapply H3; eauto. solve_picksps_equal. } + simpl. intros. fwd. eexists (_ ++ _ ++ _ :: _). + repeat rewrite <- (app_assoc _ _ k2). repeat rewrite <- (app_assoc _ _ k). + intuition. + - apply call_args_to_other_trace in H0. + fwd. econstructor; intuition eauto. + { eapply exec_ext with (pick_sp1 := _). 1: eapply IHexec; eauto. solve_picksps_equal. } + cbv beta in *. fwd. apply H3 in H0p2. + fwd. exists retvs. intuition. exists l'. intuition. eexists (_ ++ _ :: _). + repeat rewrite <- (app_assoc _ _ k2). repeat rewrite <- (app_assoc _ _ k). + intuition. + - apply call_args_to_other_trace in H0. fwd. econstructor; intuition eauto. + apply H2 in H0. fwd. exists l'. intuition. eexists (_ :: _). + intuition. + Qed. + + End WithEnv. + + Lemma extend_env {pick_sp: PickSp} : forall e1 e2, + map.extends e2 e1 -> + forall c k t m l mc post, + exec e1 c k t m l mc post -> + exec e2 c k 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} {pick_sp: PickSp}. + + Implicit Types (l: locals) (m: mem). + + Definition call e fname k 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 k t m l mc (fun k' t' m' l' mc' => exists rets, + map.getmany_of_list l' retnames = Some rets /\ post k' t' m' rets mc'). + + Lemma weaken_call: forall e fname k t m args mc post1, + call e fname k t m args mc post1 -> + forall (post2: leakage -> trace -> mem -> list word -> MetricLog -> Prop), + (forall k' t' m' rets mc', post1 k' t' m' rets mc' -> post2 k' t' m' rets mc') -> + call e fname k 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 k t m rets mc post, + call e1 f k t m rets mc post -> + call e2 f k t m rets mc post. + Proof. + unfold call. intros. fwd. repeat eexists. + - eapply H. eassumption. + - eassumption. + - eapply exec.extend_env; eassumption. + Qed. + + Lemma to_plain_eval_expr: forall m l e v, + Semantics.eval_expr m l e = Some v -> + forall k mc, exists k' mc', eval_expr m l e k mc = Some (v, k', mc'). + Proof. + induction e; cbn; intros; + repeat match goal with + | H: Some _ = Some _ |- _ => eapply Option.eq_of_eq_Some in H; subst + | IH: forall _, Some _ = Some _ -> _ |- _ => specialize IH with (1 := eq_refl) + | IH: forall _ _, Semantics.eval_expr _ _ _ = Some _ -> _ |- _ => + specialize (IH _ _ ltac:(eassumption)) + | IH: forall _: leakage, forall _: MetricLog, exists _: leakage, exists _: MetricLog, + eval_expr ?m ?l ?e _ _ = Some _ + |- context[eval_expr ?m ?l ?e ?k ?mc] => specialize (IH k mc) + | |- _ => progress fwd + | |- _ => Tactics.destruct_one_match + end; + eauto. + Qed. + + Lemma to_metric_eval_expr: forall m l e mc v mc', + MetricSemantics.eval_expr m l e mc = Some (v, mc') -> + forall k, exists k', eval_expr m l e k mc = Some (v, k', mc'). + Proof. + induction e; cbn; intros; + repeat match goal with + | H: Some _ = Some _ |- _ => eapply Option.eq_of_eq_Some in H; subst + | IH: forall _, Some _ = Some _ -> _ |- _ => specialize IH with (1 := eq_refl) + | IH: forall _ _ _, MetricSemantics.eval_expr _ _ _ _ = Some _ -> _ |- _ + => specialize (IH _ _ _ ltac:(eassumption)) + | IH: forall _: leakage, exists _: leakage, eval_expr ?m ?l ?e _ _ = Some _ + |- context[eval_expr ?m ?l ?e ?k ?mc] => specialize (IH k) + | |- _ => progress fwd + | |- _ => Tactics.destruct_one_match + end; + eauto. + Qed. + + Lemma to_leakage_eval_expr: forall m l e k v k', + LeakageSemantics.eval_expr m l e k = Some (v, k') -> + forall mc, exists mc', eval_expr m l e k mc = Some (v, k', mc'). + Proof. + induction e; cbn; intros; + repeat match goal with + | H: Some _ = Some _ |- _ => eapply Option.eq_of_eq_Some in H; subst + | IH: forall _, Some _ = Some _ -> _ |- _ => specialize IH with (1 := eq_refl) + | IH: forall _ _ _, LeakageSemantics.eval_expr _ _ _ _ = Some _ -> _ |- _ + => specialize (IH _ _ _ ltac:(eassumption)) + | IH: forall _: MetricLog, exists _: MetricLog, eval_expr ?m ?l ?e _ _ = Some _ + |- context[eval_expr ?m ?l ?e ?k ?mc] => specialize (IH mc) + | |- _ => progress fwd + | |- _ => Tactics.destruct_one_match + end; + eauto. + Qed. + + Lemma to_plain_eval_call_args: forall m l arges args, + Semantics.eval_call_args m l arges = Some args -> + forall k mc, exists k' mc', eval_call_args m l arges k mc = Some (args, k', mc'). + Proof. + induction arges; cbn; intros. + - eapply Option.eq_of_eq_Some in H. subst. eauto. + - fwd. + eapply to_plain_eval_expr in E. destruct E as (k' & mc' & E). rewrite E. + specialize IHarges with (1 := eq_refl). specialize (IHarges k' mc'). + destruct IHarges as (k'' & mc'' & IH). rewrite IH. eauto. + Qed. + + Lemma to_metric_eval_call_args: forall m l arges mc args mc', + MetricSemantics.eval_call_args m l arges mc = Some (args, mc') -> + forall k, exists k', eval_call_args m l arges k mc = Some (args, k', mc'). + Proof. + induction arges; cbn; intros. + - inversion H. subst. eauto. + - fwd. + eapply to_metric_eval_expr in E. destruct E as (k' & E). rewrite E. + specialize (IHarges _ _ _ ltac:(eassumption) k'). + destruct IHarges as (k'' & IH). rewrite IH. eauto. + Qed. + + Lemma to_leakage_eval_call_args: forall m l arges k args k', + LeakageSemantics.eval_call_args m l arges k = Some (args, k') -> + forall mc, exists mc', eval_call_args m l arges k mc = Some (args, k', mc'). + Proof. + induction arges; cbn; intros. + - inversion H. subst. eauto. + - fwd. + eapply to_leakage_eval_expr in E. destruct E as (mc' & E). rewrite E. + specialize (IHarges _ _ _ ltac:(eassumption) mc'). + destruct IHarges as (mc'' & IH). rewrite IH. eauto. + Qed. + + Definition deleakaged_ext_spec : Semantics.ExtSpec := + fun t mGive a args post => ext_spec t mGive a args (fun mReceive resvals klist => post mReceive resvals). + + Lemma deleakaged_ext_spec_ok {ext_spec_ok: ext_spec.ok ext_spec}: + Semantics.ext_spec.ok deleakaged_ext_spec. + Proof. + destruct ext_spec_ok. cbv [deleakaged_ext_spec]. constructor; eauto. + intros. cbv [Morphisms.Proper Morphisms.respectful Morphisms.pointwise_relation Basics.impl] in *. + intros. eapply weaken; eauto. intros. apply H. simpl in *. assumption. + Qed. + + Context (e: env). + Existing Instance deleakaged_ext_spec. + Existing Instance deleakaged_ext_spec_ok. + + Lemma to_plain_exec: forall c t m l post, + Semantics.exec e c t m l post -> + forall k mc, exec e c k t m l mc (fun _ t' m' l' _ => post t' m' l'). + Proof. + induction 1; intros; + repeat match reverse goal with + | H: Semantics.eval_expr _ _ _ = Some _ |- _ => + eapply to_plain_eval_expr in H; destruct H as (? & ? & H) + | H: Semantics.eval_call_args _ _ _ = Some _ |- _ => + eapply to_plain_eval_call_args in H; destruct H as (? & ? & H) + end; + try solve [econstructor; eauto]. + Qed. + + Lemma to_metric_exec: forall c t m l mc post, + MetricSemantics.exec e c t m l mc post -> + forall k, exec e c k t m l mc (fun _ t' m' l' mc' => post t' m' l' mc'). + Proof. + induction 1; intros; + repeat match reverse goal with + | H: MetricSemantics.eval_expr _ _ _ _ = Some _ |- _ => + eapply to_metric_eval_expr in H; destruct H as (? & H) + | H: MetricSemantics.eval_call_args _ _ _ _ = Some _ |- _ => + eapply to_metric_eval_call_args in H; destruct H as (? & H) + end; + try solve [econstructor; eauto]. + Qed. + + Lemma to_leakage_exec: forall c t m l k post, + LeakageSemantics.exec e c k t m l post -> + forall mc, exec e c k t m l mc (fun k' t' m' l' _ => post k' t' m' l'). + Proof. + induction 1; intros; + repeat match reverse goal with + | H: LeakageSemantics.eval_expr _ _ _ _ = Some _ |- _ => + eapply to_leakage_eval_expr in H; destruct H as (? & H) + | H: LeakageSemantics.eval_call_args _ _ _ _ = Some _ |- _ => + eapply to_leakage_eval_call_args in H; destruct H as (? & H) + end; + try solve [econstructor; eauto]. + Qed. + + Lemma to_plain_call: forall f t m args post, + Semantics.call e f t m args post -> + forall k mc, call e f k t m args mc (fun _ t' m' rets _ => post t' m' rets). + Proof. + unfold call, Semantics.call. intros. fwd. eauto 10 using to_plain_exec. + Qed. + + Lemma to_metric_call: forall f t m args mc post, + MetricSemantics.call e f t m args mc post -> + forall k, call e f k t m args mc (fun _ t' m' rets mc' => post t' m' rets mc'). + Proof. + unfold call, MetricSemantics.call. intros. fwd. eauto 10 using to_metric_exec. + Qed. + + Lemma to_leakage_call: forall f k t m args post, + LeakageSemantics.call e f k t m args post -> + forall mc, call e f k t m args mc (fun k' t' m' rets _ => post k' t' m' rets). + Proof. + unfold call, LeakageSemantics.call. intros. fwd. eauto 10 using to_leakage_exec. + Qed. +End WithParams. diff --git a/bedrock2/src/bedrock2/MetricSemantics.v b/bedrock2/src/bedrock2/MetricSemantics.v index 059ae32ff..4dbd8f0b2 100644 --- a/bedrock2/src/bedrock2/MetricSemantics.v +++ b/bedrock2/src/bedrock2/MetricSemantics.v @@ -23,7 +23,7 @@ Section semantics. Section WithMemAndLocals. Context (m : mem) (l : locals). - Local Notation "' x <- a | y ; f" := (match a with x => f | _ => y end) + Local Notation "' x <- a ; f" := (match a with Some x => f | None => None end) (right associativity, at level 70, x pattern). (* TODO XXX possibly be a bit smarter about whether things are registers, @@ -31,24 +31,21 @@ Section semantics. Fixpoint eval_expr (e : expr) (mc : metrics) : option (word * metrics) := match e with | 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, cost_set isRegStr UNK x mc) - | None => None - end + | expr.var x => 'v <- map.get l x; Some (v, cost_set isRegStr UNK x mc) | expr.inlinetable aSize t index => - 'Some (index', mc') <- eval_expr index mc | None; - 'Some v <- load aSize (map.of_list_word t) index' | None; + '(index', mc') <- eval_expr index mc; + 'v <- load aSize (map.of_list_word t) index'; 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; + '(a', mc') <- eval_expr a mc; + 'v <- load aSize m a'; 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; + '(v1, mc') <- eval_expr e1 mc; + '(v2, mc'') <- eval_expr e2 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; + '(vc, mc') <- eval_expr c mc; eval_expr (if word.eqb vc (word.of_Z 0) then e2 else e1) (cost_if isRegStr UNK (Some UNK) mc') end. @@ -56,8 +53,8 @@ Section semantics. Fixpoint eval_call_args (arges : list expr) (mc : metrics) := match arges with | e :: tl => - 'Some (v, mc') <- eval_expr e mc | None; - 'Some (args, mc'') <- eval_call_args tl mc' | None; + '(v, mc') <- eval_expr e mc; + '(args, mc'') <- eval_call_args tl mc'; Some (v :: args, mc'') | _ => Some (nil, mc) end. diff --git a/bedrock2/src/bedrock2/memory_mapped_ext_spec.v b/bedrock2/src/bedrock2/memory_mapped_ext_spec.v index 7e6266c5c..2067585d7 100644 --- a/bedrock2/src/bedrock2/memory_mapped_ext_spec.v +++ b/bedrock2/src/bedrock2/memory_mapped_ext_spec.v @@ -10,7 +10,7 @@ Require coqutil.Datatypes.String. Require Import coqutil.Map.Interface coqutil.Map.Domain. Require Import coqutil.Word.Interface coqutil.Word.Bitwidth. Require Import coqutil.Word.Properties. -Require Import bedrock2.Semantics. +Require Import bedrock2.Semantics bedrock2.LeakageSemantics bedrock2.MetricLeakageSemantics. Require Import bedrock2.TraceInspection. Local Open Scope string_scope. @@ -48,18 +48,20 @@ Section WithMem. (* Note: This ext_spec is crafted in such a way that no matter how liberal read_step and write_step are, all ext calls allowed by this ext_spec can be compiled to a RISC-V load or store instruction *) - Definition ext_spec{mmio_ext_calls: MemoryMappedExtCalls}: ExtSpec := + Definition leakage_ext_spec{mmio_ext_calls: MemoryMappedExtCalls}: ExtSpec := fun (t: trace) (mGive: mem) (action: string) (args: list word) - (post: mem -> list word -> Prop) => + (post: mem -> list word -> list word -> Prop) => exists n, (n = 1 \/ n = 2 \/ n = 4 \/ (n = 8 /\ width = 64%Z))%nat /\ ((action = "memory_mapped_extcall_read" ++ String.of_nat (n * 8) /\ exists addr, args = [addr] /\ mGive = map.empty /\ read_step n t addr (fun v mRcv => - post mRcv [word.of_Z (LittleEndian.combine n v)])) \/ + post mRcv [word.of_Z (LittleEndian.combine n v)] [addr])) \/ (action = "memory_mapped_extcall_write" ++ String.of_nat (n * 8) /\ exists addr v, args = [addr; word.of_Z (LittleEndian.combine n v)] /\ write_step n t addr v mGive /\ - post map.empty nil)). + post map.empty nil [addr])). + + Definition ext_spec{mmio_ext_calls: MemoryMappedExtCalls} : Semantics.ExtSpec := @deleakaged_ext_spec _ _ _ _ leakage_ext_spec. Definition footprint_list(addr: word)(n: nat): list word := List.unfoldn (word.add (word.of_Z 1)) n addr. @@ -97,22 +99,22 @@ Section WithMem. PropSet.subset (PropSet.of_list (footprint_list addr n)) (map.domain mExt); }. - Lemma weaken_ext_spec{mmio_ext_calls: MemoryMappedExtCalls} + Lemma weaken_leakage_ext_spec{mmio_ext_calls: MemoryMappedExtCalls} {mmio_ext_calls_ok: MemoryMappedExtCallsOk mmio_ext_calls}: forall t mGive a args post1 post2, - (forall mRcv rets, post1 mRcv rets -> post2 mRcv rets) -> - ext_spec t mGive a args post1 -> - ext_spec t mGive a args post2. + (forall mRcv rets klist, post1 mRcv rets klist -> post2 mRcv rets klist) -> + leakage_ext_spec t mGive a args post1 -> + leakage_ext_spec t mGive a args post2. Proof. - unfold ext_spec; intros; fwd; destruct H0p1; fwd; eauto 10 using weaken_read_step. + unfold leakage_ext_spec; intros; fwd; destruct H0p1; fwd; eauto 10 using weaken_read_step. Qed. - Instance ext_spec_ok(mmio_ext_calls: MemoryMappedExtCalls) - {mmio_ext_calls_ok: MemoryMappedExtCallsOk mmio_ext_calls}: ext_spec.ok ext_spec. + Instance leakage_ext_spec_ok(mmio_ext_calls: MemoryMappedExtCalls) + {mmio_ext_calls_ok: MemoryMappedExtCallsOk mmio_ext_calls}: ext_spec.ok leakage_ext_spec. Proof. constructor. - (* mGive unique *) - unfold ext_spec. intros. fwd. destruct H1p1; destruct H2p1; fwd; try congruence. + unfold leakage_ext_spec. intros. fwd. destruct H1p1; destruct H2p1; fwd; try congruence. inversion H1p1. fwd. subst n0. eapply (f_equal word.unsigned) in H2. pose proof (LittleEndian.combine_bound v). @@ -129,16 +131,20 @@ Section WithMem. eauto using write_step_unique_mGive. - (* weaken *) unfold Morphisms.Proper, Morphisms.respectful, Morphisms.pointwise_relation, - Basics.impl. eapply weaken_ext_spec. + Basics.impl. eapply weaken_leakage_ext_spec. - (* intersect *) - unfold ext_spec. intros. fwd. destruct Hp1; destruct H0p1; fwd; + unfold leakage_ext_spec. intros. fwd. destruct Hp1; destruct H0p1; fwd; match goal with | H: _ ++ _ = _ ++ _ |- _ => inversion H; clear H end; fwd; eauto 10 using intersect_read_step. Qed. + Definition ext_spec_ok(mmio_ext_calls: MemoryMappedExtCalls) + {mmio_ext_calls_ok: MemoryMappedExtCallsOk mmio_ext_calls}: Semantics.ext_spec.ok ext_spec := deleakaged_ext_spec_ok. End WithMem. +#[export] Existing Instance leakage_ext_spec. +#[export] Existing Instance leakage_ext_spec_ok. #[export] Existing Instance ext_spec. #[export] Existing Instance ext_spec_ok. diff --git a/bedrock2/src/bedrock2Examples/ct.v b/bedrock2/src/bedrock2Examples/ct.v new file mode 100644 index 000000000..d87afee6c --- /dev/null +++ b/bedrock2/src/bedrock2Examples/ct.v @@ -0,0 +1,407 @@ +Require Import Coq.ZArith.ZArith coqutil.Z.div_mod_to_equations. +Require Import bedrock2.NotationsCustomEntry. +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 div3329_vartime := func! (x) ~> ret { + ret = $(expr.op bopname.divu "x" 3329) +}. + +Definition div3329 := func! (x) ~> ret { + ret = (x * $989558401) >> $32; + ret = (ret + (x - ret >> $1)) >> $11 +}. + +From coqutil Require Import Word.Properties Word.Interface Tactics.letexists. +Import Interface Coq.Lists.List List.ListNotations. +From bedrock2 Require Import Semantics LeakageSemantics FE310CSemantics LeakageWeakestPrecondition LeakageProgramLogic. +Import LeakageProgramLogic.Coercions. +Section WithParameters. + Context {word: word.word 32} {mem: Interface.map.map word Byte.byte}. + Context {word_ok : word.ok word} {mem_ok : Interface.map.ok mem}. + Context {pick_sp: PickSp}. + +#[global] Instance ctspec_of_div3329 : spec_of "div3329" := + fnspec! exists f, "div3329" x ~> ret, + { requires k t m := True ; ensures k' t' m' := t' = t /\ m' = m /\ k' = f k }. + +Lemma div3329_ct : program_logic_goal_for_function! div3329. +Proof. + repeat (straightline || eexists). + { (* no leakag -- 3 minutes *) cbn. exact eq_refl. } +Qed. + +#[global] Instance vtspec_of_div3329 : spec_of "div3329_vartime" := + fnspec! "div3329_vartime" x ~> ret, + { requires k t m := True ; + ensures k' t' m' := t' = t /\ m' = m /\ + k' = [leak_word (word.of_Z 3329); leak_word x]++k }. + +Lemma div3329_vt : program_logic_goal_for_function! div3329_vartime. +Proof. + repeat (straightline || eexists). +Qed. + +(* Mon Jul 1 14:14:15 EDT 2024 *) + +Import Byte. +Definition getchar_event c : LogItem := + ((Interface.map.empty, "getchar", []), (Interface.map.empty, [word.of_Z (byte.unsigned c)])). +#[global] Instance ctspec_of_getchar : spec_of "getchar" := + fnspec! exists f, "getchar" ~> ret, + { requires k t m := True ; ensures k' t' m' := + exists c, ret = word.of_Z (byte.unsigned c) /\ + k' = f ++ k /\ m' = m /\ t' = cons (getchar_event c) t }. + +Definition getline := func! (dst, n) ~> n { + i = $0; + c = $0; + while (i < n) { + unpack! c = getchar(); + if (c == $0x0a) { n = i } + else { store1(dst + i, c); i = i + $1 } + } +}. + +Local Notation "xs $@ a" := (Array.array Separation.ptsto (word.of_Z 1) a xs) (at level 10, format "xs $@ a"). +Local Infix "*" := Separation.sep. +Local Infix "*" := Separation.sep : type_scope. + +Definition getline_io n bs := + let chars := map getchar_event (rev bs) in + let newline := if n =? length bs then nil else [getchar_event Byte.x0a] in + newline ++ chars. + +Local Fixpoint getline_leakage f dst n (bs : nat) (i : word) := + if i =? n then leak_bool false :: nil else + match bs with + | S bs => getline_leakage f dst n bs (word.add i (word.of_Z 1)) ++ (leak_word (word.add dst i) :: leak_bool false :: f ++ leak_unit :: leak_bool true :: nil) + | O => leak_bool false :: leak_bool true :: f ++ leak_unit :: leak_bool true :: nil + end. + +#[global] Instance ctspec_of_getline : spec_of "getline" := + fnspec! exists f, "getline" (dst n : word) / d R ~> l, + { requires k t m := (d$@dst * R) m /\ length d = n :> Z ; + ensures k' t' m' := exists bs es, + k' = f dst n l ++ k /\ + (bs$@dst * es$@(word.add dst l) * R) m' /\ + length bs = l :> Z /\ + length bs + length es = n :> Z /\ + t' = getline_io n bs ++ t + (* /\ ~ In Byte.x0a bs *) }. + +(* Mon Jul 1 14:24:28 EDT 2024 *) + +Require Coq.Program.Wf. +Import Separation SeparationLogic. + +Lemma getline_ct : program_logic_goal_for_function! getline. +Proof. + repeat straightline. + + refine ((LeakageLoops.tailrec_earlyout + (HList.polymorphic_list.cons (list byte) + (HList.polymorphic_list.cons (_ -> Prop) + HList.polymorphic_list.nil)) + ["dst";"n";"i";"c"]) + (fun (v:Z) es R k t m dst_ n_ i c => PrimitivePair.pair.mk ( + n_ = n /\ dst_ = dst /\ v = i :> Z /\ + i <= n /\ + (es$@(word.add dst i) * R) m /\ length es = word.sub n i :> Z + ) + (fun K T M DST N I C => DST = dst /\ + exists bs ES, (bs$@(word.add dst i) * ES$@(word.add dst I) * R) M /\ I = N /\ + length bs = word.sub I i :> Z /\ + length ES = word.sub n I :> Z /\ + i <= N <= n /\ + T = getline_io (n-i) bs ++ t /\ + K = getline_leakage f dst n (length bs) i ++ k + )) + (fun i j => j < i <= n) + _ _ _ _ _ _ _); + 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. + { eapply Z.gt_wf. } + { split. { subst i. rewrite word.unsigned_of_Z_0. blia. } + subst i; rewrite word.add_0_r; split; [ecancel_assumption|]. rewrite word.sub_0_r; auto. } + + { + pose proof word.unsigned_range n. + pose proof word.unsigned_range x3 as Hx3. + subst br. rewrite word.unsigned_ltu in H2; case Z.ltb eqn:? in H2; + rewrite ?word.unsigned_of_Z_1, ?word.unsigned_of_Z_0, ?word.unsigned_sub_nowrap in *; try blia; []. + eapply LeakageWeakestPreconditionProperties.Proper_call; repeat intro; cycle 1. + { eapply H. exact I. } + repeat straightline. + eexists _, _; repeat straightline. + split; repeat straightline. + split; repeat straightline. + { left; repeat straightline. + { subst br. rewrite word.unsigned_ltu, Z.ltb_irrefl. apply word.unsigned_of_Z_0. } + eexists _, _; repeat straightline. + eapply word.if_nonzero, word.eqb_true in H4. + instantiate (1:=nil); cbn [Array.array]; split. + { ecancel_assumption. } + { + split; trivial. + split. { rewrite word.unsigned_sub_nowrap; simpl length; blia. } + split. { rewrite word.unsigned_sub_nowrap; blia. } + split. { blia. } + split. { (* I/O *) + cbv [getline_io]. cbn [map rev List.app length]. case (Z.eqb_spec (n'0-x3) 0%nat) as []; try blia. + rewrite app_nil_r. subst a0. simpl. + eapply f_equal2; f_equal; trivial. + progress change 10 with (byte.unsigned Byte.x0a) in H4. + pose proof byte.unsigned_range x2. + pose proof byte.unsigned_range Byte.x0a. + eapply word.of_Z_inj_small, byte.unsigned_inj in H4; trivial; blia. } + (* leakage *) + subst k'''. cbn [getline_leakage leak_binop "++" length]. + rewrite (proj2 (Z.eqb_neq _ _)) by blia; trivial. simpl. rewrite <- app_assoc. reflexivity. } } + + (* store *) + destruct x as [|x_0 x]. { cbn [length] in *; blia. } + cbn [Array.array] in *. + repeat straightline. + right; repeat straightline. + eexists _, _, _; repeat straightline. + { instantiate (1:=x). + subst i. + rewrite word.add_assoc. + split. { rewrite word.unsigned_add_nowrap; rewrite ?word.unsigned_of_Z_1; try blia. } + split; [ecancel_assumption|]. + cbn [length] in *. + pose proof word.unsigned_of_Z_1. + pose proof word.unsigned_add_nowrap x3 (word.of_Z 1). + pose proof word.unsigned_sub_nowrap n (word.add x3 (word.of_Z 1)). + blia. } + { split. + { subst i. + pose proof word.unsigned_of_Z_1. + pose proof word.unsigned_add_nowrap x3 (word.of_Z 1). + pose proof word.unsigned_sub_nowrap n (word.add x3 (word.of_Z 1)). + blia. } + repeat straightline. + (* subroutine return *) + subst i. + + rename x9 into bs. + rename x10 into es. + rename x6 into I. + rename x3 into _i. + rewrite word.add_assoc in H10. + + eexists (byte.of_Z x1 :: bs), (es). + cbn ["$@" "++"]. + split. { ecancel_assumption. } + split; trivial. + split. { cbn [length]. rewrite Nat2Z.inj_succ, H15. + pose proof word.unsigned_of_Z_1. + pose proof word.unsigned_add_nowrap _i (word.of_Z 1) ltac:(blia). + rewrite 2 word.unsigned_sub_nowrap; blia. } + split; trivial. + split. { + pose proof word.unsigned_of_Z_1. + pose proof word.unsigned_add_nowrap _i (word.of_Z 1) ltac:(blia). + blia. } + split. { (* I/O *) + subst T a0. + cbv [getline_io]; cbn [rev List.map]. + repeat rewrite ?map_app, <-?app_comm_cons, <-?app_assoc; f_equal. + { pose proof word.unsigned_of_Z_1 as H_1. + rewrite (word.unsigned_add_nowrap _i (word.of_Z 1) ltac:(blia)), H_1; cbn [length]. + case Z.eqb eqn:? at 1; case Z.eqb eqn:? at 1; trivial; try blia. + { (* WHY manual? does zify do a bad job here? *) eapply Z.eqb_neq in Heqb1. blia. } + { (* WHY manual? does zify do a bad job here? *) eapply Z.eqb_eq in Heqb1. blia. } } + f_equal. + cbn [map List.app]. + f_equal. + f_equal. + subst x1. + pose proof byte.unsigned_range x2. + rewrite word.unsigned_of_Z_nowrap, byte.of_Z_unsigned; trivial; blia. } + (* leakage *) + subst K a1; cbn [getline_leakage leak_binop length]. + rewrite (proj2 (Z.eqb_neq _ _)) by blia; trivial. + repeat (simpl || rewrite <- app_assoc). reflexivity. } } + + { (* buffer full *) + replace x3 with n in *; cycle 1. + { subst br; rewrite word.unsigned_ltu in *; eapply word.if_zero, Z.ltb_nlt in H2. + apply word.unsigned_inj. blia. } + exists x, nil; cbn [Array.array]. + split. { ecancel_assumption. } + split. { trivial. } + split. { trivial. } + rewrite word.unsigned_sub_nowrap, Z.sub_diag in H7 by blia. + split. { rewrite word.unsigned_sub_nowrap, Z.sub_diag by blia; trivial. } + split. { blia. } + split. { destruct x; cbn [length] in *; try blia; cbn. + rewrite Z.sub_diag; reflexivity. } + destruct x; try (cbn in *; blia). + cbn [getline_leakage length]; rewrite Z.eqb_refl; trivial. } + + do 2 eexists. split. { subst k0 i. rewrite word.sub_0_r in *. + assert (length x3 = Z.to_nat (word.unsigned x0)) as -> by blia. reflexivity. } + subst i. + rewrite word.add_0_r in *. + split. + { ecancel_assumption. } + split. + { rewrite H5. rewrite word.sub_0_r. trivial. } + split. + { rewrite H5, H6, word.sub_0_r, word.unsigned_sub_nowrap; blia. } + subst t0. + rewrite word.unsigned_of_Z_0, Z.sub_0_r. + trivial. + +(* Tue Jul 2 14:26:41 EDT 2024 *) +Qed. +(* Mon Jul 8 17:03:59 EDT 2024 *) + +Require Import bedrock2Examples.memequal. + +Definition password_checker := func! (password) ~> ret { + stackalloc 8 as x; (*password is 8 characters*) + unpack! n = getline(x, $8); + unpack! ok = memequal(x, password, $8); + ret = (n == $8) & ok + }. + +#[global] Instance ctspec_of_password_checker : spec_of "password_checker" := + fnspec! exists f, "password_checker" password_addr / R password ~> ret (*bs =? password*), + { requires k t m := length password = 8 :> Z /\ (password$@password_addr * R) m ; + ensures k' t' m' := exists bs (l : word), + (password$@password_addr * R) m' /\ + t' = getline_io 8 bs ++ t /\ + length bs = l :> Z /\ + (k' = f k password_addr l) /\ + (word.unsigned ret = 1 <-> bs = password) }. + +Fail Lemma password_checker_ct : program_logic_goal_for_function! password_checker. (*Why*) +Global Instance spec_of_memequal : spec_of "memequal" := spec_of_memequal. + +Require Import coqutil.Word.SimplWordExpr. + +Lemma password_checker_ct : program_logic_goal_for_function! password_checker. +Proof. + repeat straightline. + eapply LeakageWeakestPreconditionProperties.Proper_call; repeat intro; cycle 1. + { eapply H. split. 2: rewrite word.unsigned_of_Z; eassumption. ecancel_assumption. } + repeat straightline. + seprewrite_in_by @Array.bytearray_index_merge H9 ltac:(blia). + eapply LeakageWeakestPreconditionProperties.Proper_call; repeat intro; cycle 1. + { eapply H0. split. + { ecancel_assumption. } split. + { ecancel_assumption. } + split. + { rewrite ?app_length; blia. } + { rewrite H1. rewrite word.unsigned_of_Z. reflexivity. } } + assert (length ((x0 ++ x1)) = 8%nat). + { rewrite ?app_length. rewrite word.unsigned_of_Z_nowrap in H11; blia. } + repeat straightline. + do 2 eexists. split. { ecancel_assumption. } + split. { subst a0. rewrite word.unsigned_of_Z. exact eq_refl. } + split. { eassumption. } + split. { (* leakage *) + subst a0. subst a. subst a2. instantiate (1 := fun _ _ => _). simpl. reflexivity. } + { (* functional correctness *) + subst ret. + destruct (word.eqb_spec x (word.of_Z 8)) as [->|?]; cycle 1. + { rewrite word.unsigned_and_nowrap, word.unsigned_of_Z_0, Z.land_0_l; split; try discriminate. + intros X%(f_equal (@length _)). case H13; clear H13; apply word.unsigned_inj. + rewrite <-H10, X, word.unsigned_of_Z_nowrap; blia. } + rewrite word.unsigned_and_nowrap, word.unsigned_of_Z_1. + destruct x1; cycle 1. + { cbn [length] in *. blia. } { rewrite ?app_nil_r in *. rewrite <-H16. + case H15 as [->|]; intuition try congruence. rewrite H15. trivial. } } +Qed. + +Definition output_event x : LogItem := + ((Interface.map.empty, "output", [x]), (Interface.map.empty, [])). +#[global] Instance ctspec_of_output : spec_of "output" := + fnspec! exists f, "output" x, + { requires k t m := True ; + ensures k' t' m' := k' = f ++ k /\ m' = m /\ t' = cons (output_event x) t }. + +Definition getprime_event p : LogItem := + ((Interface.map.empty, "getprime", []), (Interface.map.empty, [p])). +#[global] Instance ctspec_of_getprime : spec_of "getprime" := + fnspec! exists f, "getprime" ~> p, + { requires k t m := True ; + ensures k' t' m' := k' = f ++ k /\ m' = m /\ t' = cons (getprime_event p) t }. + +Definition semiprime := func! () ~> (p, q) { + unpack! p = getprime(); + unpack! q = getprime(); + n = p * q; + output(n) +}. + +#[global] Instance ctspec_of_semiprime : spec_of "semiprime" := + fnspec! exists f, "semiprime" ~> p q, + { requires k t m := True ; + ensures k' t' m' := + k' = f ++ k /\ m' = m + /\ t' = [output_event (word.mul p q); getprime_event q; getprime_event p]++ t }. + +Lemma semiprime_ct : program_logic_goal_for_function! semiprime. +Proof. + repeat straightline. + eapply LeakageWeakestPreconditionProperties.Proper_call; repeat intro; [|eapply H]; repeat straightline. + eapply LeakageWeakestPreconditionProperties.Proper_call; repeat intro; [|eapply H]; repeat straightline. + eapply LeakageWeakestPreconditionProperties.Proper_call; repeat intro; [|eapply H0]; repeat straightline. + Tactics.ssplit; trivial; simpl in *; align_trace. +Qed. + +Definition maskloop := func! (a) { + i = $0; + while (i < $2) { + mask = $0-((load1(a+i)>>i)&$1); + store1(a+i, mask & $123); + i = i + $1 + } +}. + +Require Import coqutil.Map.Interface bedrock2.Map.Separation bedrock2.Map.SeparationLogic. + +(*#[global] Instance ctspec_of_maskloop : spec_of "maskloop" := + fun functions => forall k a, exists k_, forall a0 a1 R t m, + m =* ptsto a a0 * ptsto (word.add a (word.of_Z 1)) a1 * R -> + LeakageWeakestPrecondition.call functions "maskloop" k t m [a] + (fun k' t' m' rets => rets = [] /\ k' = k_). + +Lemma maskloop_ct : program_logic_goal_for_function! maskloop. +Proof. + enter maskloop. eapply LeakageWeakestPreconditionProperties.start_func. + enter maskloop. + + match goal with + | H:map.get ?functions ?fname = Some _ + |- _ => + eapply LeakageWeakestPreconditionProperties.start_func; + [ exact H | clear H ] + end; cbv beta match delta [func]. + intros. eexists. intros. + repeat straightline. + eexists nat, lt, (fun i k _ (m : mem) (l : locals) => + map.get l "a" = Some a /\ + map.get l "i" = Some (word.of_Z (Z.of_nat i)) /\ ( + i = 0%nat /\ m =* ptsto a a0 * ptsto (word.add a (word.of_Z 1)) a1 * R \/ + i = 1%nat /\ False \/ + i = 2%nat /\ False)). + Tactics.ssplit; auto using lt_wf. + { exists 0%nat; Tactics.ssplit. + { subst l0 l; rewrite map.get_put_diff, map.get_put_same; congruence. } + { subst l0 l v. rewrite map.get_put_same; trivial. } + left; Tactics.ssplit; trivial; ecancel_assumption. } + intuition subst. +Abort.*) +End WithParameters. diff --git a/bedrock2/src/bedrock2Examples/kyberslash.v b/bedrock2/src/bedrock2Examples/kyberslash.v new file mode 100644 index 000000000..6b8a4d71c --- /dev/null +++ b/bedrock2/src/bedrock2Examples/kyberslash.v @@ -0,0 +1,493 @@ + +(*https://github.com/pq-crystals/kyber/commit/dda29cc63af721981ee2c831cf00822e69be3220*) +(* +typedef struct{ + int16_t coeffs[KYBER_N]; +} poly; + +void poly_tomsg(uint8_t msg[KYBER_INDCPA_MSGBYTES], const poly *a) +{ + unsigned int i,j; + uint32_t t; + + for(i=0;icoeffs[8*i+j]; + // t += ((int16_t)t >> 15) & KYBER_Q; + // t = (((t << 1) + KYBER_Q/2)/KYBER_Q) & 1; + t <<= 1; + t += 1665; + t *= 80635; + t >>= 28; + t &= 1; + msg[i] |= t << j; + } + } +} + *) +Require Import Coq.Lists.List. +Require Import Coq.Strings.String Coq.ZArith.ZArith. +Require Import coqutil.Z.Lia. +From bedrock2 Require Import NotationsCustomEntry LeakageProgramLogic Map.Separation Map.SeparationLogic Array Scalars LeakageLoops. +From coqutil Require Import Datatypes.List Word.Bitwidth Word.Interface Map.Interface. +From bedrock2 Require Import Semantics LeakageSemantics Syntax. +Import coqutil.Word.Properties coqutil.Map.Properties. +Require Import bedrock2.AbsintWordToZ. +Require Import bedrock2.ZnWords. +From coqutil.Macros Require Import symmetry. +Require Import bedrock2.LeakageWeakestPrecondition bedrock2.LeakageWeakestPreconditionProperties bedrock2.LeakageProgramLogic. +From coqutil.Tactics Require Import Tactics letexists. + +Local Open Scope string_scope. Local Open Scope Z_scope. Local Open Scope list_scope. + +Infix "/" := (expr.op bopname.divu) (in custom bedrock_expr at level 5, left associativity). + +Section WithWord. + Context {width: Z} {BW: Bitwidth width}. + Context {word: word.word width}. + Context {mem: map.map word Byte.byte}. + Context {locals: map.map string word}. + Context {env : map.map string (list string * list string * Syntax.cmd)}. + Context {word_ok: word.ok word} {mem_ok: map.ok mem} {locals_ok: map.ok locals} {env_ok : map.ok env}. + Context {ext_spec: ExtSpec} {ext_spec_ok : ext_spec.ok ext_spec}. + Context {pick_sp: PickSp}. + Context (width_big: 4 <= width). (*to avoid division by zero*) + Context (KYBER_N KYBER_Q KYBER_INDCPA_MSGBYTES : Z). + (* global constants in bedrock2? *) + + Definition poly_tomsg := + func! (msg, a_coeffs) { + i = $0; + while (i < coq:(expr.literal KYBER_N) / $8) { + store1(msg + i, $0); + j = $0; + while (j < $8) { + t = load2(a_coeffs + $2 (* <- because 2-byte array *) * ($8 * i + j)); + t = t << $1; + t = t + $1665; + t = t * $80635; + t = t >> $28; + t = t & $1; + store1(msg + i, load1(msg + i) | (t << j)); + j = j + $1; + coq:(cmd.unset "t") + }; + i = i + $1; + coq:(cmd.unset "j") + } + }. + + Fixpoint get_inner_leakage start f fi fj (i j : word) (vj : nat) : leakage := + match vj with + | S vj' => get_inner_leakage start f fi fj (fi i j) (fj i j) vj' ++ f i j + | O => start + end%list. + + Fixpoint get_outer_leakage start f fi (i : word) (vi : nat) : leakage := + match vi with + | S vi' => get_outer_leakage start f fi (fi i) vi' ++ f i + | O => start + end%list. + + + (*Definition bad_poly_tomsg := + func! (msg, a_coeffs) { + i = $0; + while (i < KYBER_N / $8) { + store1(msg + i, $0); + j = $0; + while (j < $8) { + t = load2(a_coeffs + $2 (* <- because 2-byte array *) * ($8 * i + j)); + (* t += ((int16_t)t >> 15) & KYBER_Q; + ^ want to implement that. t is a uint16_t. + apparently uint -> int casts are implementation-defined when not in range. + how confusing. + so we should assume that t is in int16_t range. + But then ((int16_t)t >> 15) = 0, and the whole thing is a no-op. + So what? + I suppose we just assume the cast just does nothing (aside from changing the type), + regardless of the value of t. That's the only thing that makes that line of code + reasonable. + *) + j = j + $1 + }; + i = i + $1 + } + }.*) + + Instance poly_tomsg_ct : spec_of "poly_tomsg" := + fnspec! exists f : word -> word -> Z -> leakage, "poly_tomsg" msg a_coeffs / (msg_vals : list Byte.byte) (a_coeffs_vals : list word) (R : mem -> Prop), + { requires k t m := + ((array scalar8 (word.of_Z 1) msg msg_vals) ⋆ (array scalar16 (word.of_Z 2) a_coeffs a_coeffs_vals) ⋆ R)%sep m /\ + KYBER_N = Z.of_nat (List.length a_coeffs_vals) /\ + KYBER_INDCPA_MSGBYTES = Z.of_nat (List.length msg_vals) /\ + @word.unsigned _ word (word.divu (word.of_Z KYBER_N) (word.of_Z 8)) <= KYBER_INDCPA_MSGBYTES ; + ensures K T M := + T = t /\ K = f msg a_coeffs KYBER_N ++ k }. + + Lemma poly_tomsg_ok : program_logic_goal_for_function! poly_tomsg. + Proof. + repeat straightline. + refine ((LeakageLoops.tailrec + (* types of ghost variables*) (let c := HList.polymorphic_list.cons in c _ (c _ HList.polymorphic_list.nil)) + (* program variables *) ("i" :: "a_coeffs" :: "msg" :: nil))%string + (fun vi msg_vals a_coeffs_vals k t m i a_coeffs msg => + PrimitivePair.pair.mk + (KYBER_N = Z.of_nat (List.length a_coeffs_vals) /\ + KYBER_INDCPA_MSGBYTES = Z.of_nat (List.length msg_vals) /\ + ((array scalar8 (word.of_Z 1) msg msg_vals) * (array scalar16 (word.of_Z 2) a_coeffs a_coeffs_vals) * R)%sep m + /\ vi = @word.unsigned _ word (word.divu (word.of_Z KYBER_N) (word.of_Z 8)) - word.unsigned i) (* precondition *) + (fun K T M I A_COEFFS MSG => (*postcondition*) + T = t /\ A_COEFFS = a_coeffs /\ MSG = msg /\ + exists MSG_VALS A_COEFFS_VALS, + KYBER_N = Z.of_nat (List.length A_COEFFS_VALS) /\ + KYBER_INDCPA_MSGBYTES = Z.of_nat (List.length MSG_VALS) /\ + ((array scalar8 (word.of_Z 1) msg MSG_VALS) * (array scalar16 (word.of_Z 2) a_coeffs A_COEFFS_VALS) * R)%sep M /\ + K = (get_outer_leakage _ _ _ i (Z.to_nat vi) ++ k)%list + )) + (fun n m => 0 <= n < m) (* well_founded relation *) + _ _ _ _ _ _ _); + 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 *. + { cbv [enforce]; cbn. + subst l l0. + repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec; cbn). split. + { exact eq_refl. } + { eapply map.map_ext; intros k0. + repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec, ?map.get_empty; cbn -[String.eqb]). + repeat (destruct String.eqb; trivial). } } + { exact (Z.lt_wf _). } + { repeat (straightline || intuition eauto). } + { repeat straightline. cbn in localsmap. + cbv [dexpr expr expr_body localsmap get literal dlet.dlet]. + (*would be nice if straightline did the following stuff... but the non-leakage + straightline doesn't do it either, so i'll just leave it alone*) + eexists. eexists. split. + { rewrite ?Properties.map.get_put_dec. cbn. eauto. } + split. + { repeat straightline. + eapply dexpr_expr. repeat straightline. letexists; split. + { rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. letexists; split. + { rewrite ?Properties.map.get_put_dec; exact eq_refl. } + eapply dexpr_expr. repeat straightline. + (*finally we do something interesting*) + destruct (word.ltu x1 _) eqn:E. + 2: { rewrite word.unsigned_of_Z_0 in H4. exfalso. auto. } + rewrite word.unsigned_ltu in E. apply Z.ltb_lt in E. + assert (nsmall: (0 <= Z.to_nat (word.unsigned x1) < Datatypes.length x)%nat) by ZnWords. + assert (Ex1: x1 = @word.of_Z _ word (@word.unsigned _ word (word.of_Z 1) * Z.of_nat (Z.to_nat (word.unsigned x1)))). + { rewrite Z2Nat.id. + 2: { assert (Hnonneg := word.unsigned_range x1 ). blia. } + apply word.unsigned_inj. repeat rewrite word.unsigned_of_Z. cbv [word.wrap]. + rewrite Z.mul_mod_idemp_l. + 2: { apply word.modulus_nonzero. } + assert (Hx1 := word.unsigned_range x1). + rewrite Z.mod_small; blia. } + eapply Scalars.store_one_of_sep. + { seprewrite_in (@array_index_nat_inbounds _ _ _ _ _ _ _ ptsto (word.of_Z 1) Byte.x00 x x3 (Z.to_nat (word.unsigned x1))) H3. + { ZnWords. } + rewrite <- Ex1 in H3. + ecancel_assumption. } + repeat straightline. (* neat, why did that work now? *) + refine ((LeakageLoops.tailrec + (* types of ghost variables*) (let c := HList.polymorphic_list.cons in c _ (c _ HList.polymorphic_list.nil)) + (* program variables *) ("j" :: "i" :: "a_coeffs" :: "msg" :: nil))%string + (fun vj msg_vals a_coeffs_vals k t m j i a_coeffs msg => + PrimitivePair.pair.mk + (KYBER_N = Z.of_nat (List.length a_coeffs_vals) /\ + KYBER_INDCPA_MSGBYTES = Z.of_nat (List.length msg_vals) /\ + i = x1(*value of i before we enter loop*) /\ + ((array scalar8 (word.of_Z 1) msg msg_vals) * (array scalar16 (word.of_Z 2) a_coeffs a_coeffs_vals) * R)%sep m + /\ vj = word.wrap 8 - word.unsigned j) (* precondition *) + (fun K T M J I A_COEFFS MSG => (*postcondition*) + T = t /\ A_COEFFS = a_coeffs /\ MSG = msg /\ + exists MSG_VALS A_COEFFS_VALS, + KYBER_N = Z.of_nat (List.length A_COEFFS_VALS) /\ + KYBER_INDCPA_MSGBYTES = Z.of_nat (List.length MSG_VALS) /\ + I = x1 /\ + ((array scalar8 (word.of_Z 1) msg MSG_VALS) * (array scalar16 (word.of_Z 2) a_coeffs A_COEFFS_VALS) * R)%sep M /\ + K = (get_inner_leakage _ _ _ _ i j (Z.to_nat vj) ++ k)%list + )) + (fun n m => 0 <= n < m) (* well_founded relation *) + _ _ _ _ _ _ _); + 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 *. + { cbv [LeakageLoops.enforce]; cbn. + subst l. + repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec; cbn). split. + { exact eq_refl. } + { eapply map.map_ext; intros k0'. + repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec, ?map.get_empty; cbn -[String.eqb]). + repeat (destruct String.eqb; trivial). } } + { exact (Z.lt_wf _). } + { seprewrite_in (symmetry! @array_cons) H5. + seprewrite_in @sep_comm H5. + remember (Z.to_nat (word.unsigned x1)) as n eqn:En. + rewrite Ex1 in H5. + replace (Z.of_nat n) with (Z.of_nat (List.length (List.firstn n x))) in H5. + 2: { rewrite List.firstn_length. blia. } + seprewrite_in (symmetry! @array_append) H5. subst. + split; [|split; [|split; [|split] ] ]. + 4: ecancel_assumption. + { assumption. } + { repeat rewrite List.app_length. cbn [List.length]. + rewrite List.firstn_length. rewrite List.skipn_length. blia. } + { reflexivity. } + { reflexivity. } } + { repeat straightline. cbn in localsmap. + cbv [dexpr expr expr_body localsmap get literal dlet.dlet]. + (*would be nice if straightline did the following stuff... but the non-leakage + straightline doesn't do it either, so i'll just leave it alone*) + eexists. eexists. split. + { rewrite ?Properties.map.get_put_dec. cbn. eauto. } + split. + { repeat straightline. + eapply dexpr_expr. repeat straightline. letexists; split. + { rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. letexists; split. + { rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. letexists; split. + { rewrite ?Properties.map.get_put_dec; exact eq_refl. } + destruct (word.ltu _ _) eqn:Ex6. + 2: { rewrite word.unsigned_of_Z_0 in H10. exfalso. auto. } + rewrite word.unsigned_ltu in Ex6. apply Z.ltb_lt in Ex6. + eexists. split. + { eapply load_two_of_sep. repeat straightline. + remember (word.add (word.mul v2 x1) x6) as n eqn:En. + seprewrite_in (@array_index_nat_inbounds _ _ _ _ _ _ _ scalar16 (word.of_Z 2) (word.of_Z 0) x5 x8 (Z.to_nat (word.unsigned n))) H11. + 2: { repeat straightline. use_sep_assumption. cancel. + cancel_seps_at_indices 1%nat 0%nat. + { f_equal. f_equal. subst v1 n. Fail Timeout 1 ZnWords. + rewrite Z2Nat.id. + 2: { assert (Hnonneg:= word.unsigned_range (word.add (word.mul v2 x1) x6)). + blia. } + ZnWords. } + ecancel_done. } + Fail Timeout 1 ZnWords. + subst. subst v1. subst v0. subst v2. + assert (Hnonneg := word.unsigned_range (word.add (word.mul (word.of_Z 8) x1) x6)). + enough ((word.unsigned (word.add (word.mul (word.of_Z 8) x1) x6)) < KYBER_N). + { subst KYBER_N. blia. } + assert (0 < word.unsigned (word:=word) (word.of_Z 8)). + { rewrite word.unsigned_of_Z. cbv [word.wrap]. + rewrite Z.mod_small; try split; try blia. + assert (X := Z.pow_le_mono_r 2 4 width). specialize (X ltac:(blia) ltac:(blia)). + blia. } + assert (0 < 2 ^ width). + { apply Z.pow_pos_nonneg; blia. } + rewrite word.unsigned_add, word.unsigned_mul, word.unsigned_divu in * by blia. + rewrite word.unsigned_of_Z in E. cbv [word.wrap] in *. + + rewrite Z.add_mod_idemp_l by blia. rewrite word.unsigned_of_Z in *. + assert (word.unsigned x1 < KYBER_N mod 2 ^ width / word.wrap 8). + { eapply Z.lt_le_trans. 1: eassumption. + apply Z.mod_le; try blia. + apply Z_div_nonneg_nonneg; try blia. + apply Z_mod_nonneg_nonneg; blia. } + enough ((word.wrap 8 * word.unsigned x1 + word.unsigned x6) < KYBER_N). + { eapply Z.le_lt_trans. 2: eassumption. apply Z.mod_le; try blia. + assert (Hx6 := word.unsigned_range x6). assert (Hx1 := word.unsigned_range x1). + blia. } + assert (word.unsigned x1 < KYBER_N / word.wrap 8). + { eapply Z.lt_le_trans. 1: eassumption. + apply Z.div_le_mono; try blia. apply Z.mod_le; blia. } + enough (word.wrap 8 * (word.unsigned x1 + 1) <= KYBER_N). + { blia. } + assert (word.unsigned x1 + 1 <= KYBER_N / word.wrap 8) by blia. + apply Zmult_le_compat_l with (p := word.wrap 8) in H16; try blia. + eapply Z.le_trans. 1: eassumption. + apply Z.mul_div_le. blia. } + repeat straightline. eapply dexpr_expr. repeat straightline. letexists; split. + { cbv [l]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. eapply dexpr_expr. repeat straightline. letexists; split. + { cbv [l0]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. eapply dexpr_expr. repeat straightline. letexists; split. + { cbv [l1]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. eapply dexpr_expr. repeat straightline. letexists; split. + { cbv [l2]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. eapply dexpr_expr. repeat straightline. letexists; split. + { cbv [l3]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. eapply dexpr_expr. repeat straightline. letexists; split. + { cbv [l4 l3 l2 l1 l0 l]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. letexists; split. + { cbv [l4 l3 l2 l1 l0 l]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. eapply dexpr_expr. repeat straightline. letexists; split. + { cbv [l4 l3 l2 l1 l0 l]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. letexists; split. + { cbv [l4 l3 l2 l1 l0 l]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. letexists; split. + { eapply Scalars.load_one_of_sep. + seprewrite_in_by (@array_index_nat_inbounds _ _ _ _ _ _ _ ptsto (word.of_Z 1) Byte.x00 x4 x9 (Z.to_nat (word.unsigned x1))) H11 ZnWords. + rewrite <- Ex1 in H11. + ecancel_assumption. } + repeat straightline. letexists; split. + { cbv [l4 l3 l2 l1 l0 l]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. letexists; split. + { cbv [l4 l3 l2 l1 l0 l]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + eapply Scalars.store_one_of_sep. + { seprewrite_in_by (@array_index_nat_inbounds _ _ _ _ _ _ _ ptsto (word.of_Z 1) Byte.x00 x4 x9 (Z.to_nat (word.unsigned x1))) H11 ZnWords. + rewrite <- Ex1 in H11. ecancel_assumption. } + repeat straightline. eapply dexpr_expr. repeat straightline. letexists; split. + { cbv [l4 l3 l2 l1 l0 l]. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. eexists _, _, _, _. split. + { cbv [enforce l6 l5 l4 l3 l2 l1 l0 l]. + repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec; cbn). split. + { exact eq_refl. } + { eapply map.map_ext; intros kk. + repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec, ?map.get_empty; cbn -[String.eqb]). + repeat (destruct String.eqb; trivial). } } + repeat straightline. + + seprewrite_in (symmetry! @array_cons) H12. + remember (Byte.byte.of_Z (word.unsigned (word.or _ _))) as something. + seprewrite_in @sep_comm H12. + remember (Z.to_nat (word.unsigned x1)) as n eqn:En. + rewrite Ex1 in H12. + replace (Z.of_nat n) with (Z.of_nat (List.length (List.firstn n x4))) in H12. + 2: { rewrite List.firstn_length. blia. } + seprewrite_in (symmetry! @array_append) H12. subst. + assert (8 < 2 ^ width). + { assert (X := Z.pow_le_mono_r 2 4 width). specialize (X ltac:(blia) ltac:(blia)). + blia. } + rewrite word.unsigned_of_Z in Ex6. cbv [word.wrap] in *. + rewrite Z.mod_small in * by blia. + + eexists. eexists. eexists. split. + { ssplit. 4: ecancel_assumption. + all: intuition eauto. + repeat rewrite List.app_length. cbn [List.length]. + rewrite List.firstn_length. rewrite List.skipn_length. blia. } + split. + { clear H12. subst v15. subst v. + rewrite word.unsigned_add. rewrite word.unsigned_of_Z. cbv [word.wrap]. + rewrite (Z.mod_small 8) by blia. rewrite (Z.mod_small 1) by blia. + pose proof (word.unsigned_range x6). rewrite Z.mod_small by blia. + blia. } + (*postcondition*) + intros. intuition. + destruct H18 as [MSG_VALS [A_COEFFS_VALS [H18 [H19 [H20 [H21 H22] ] ] ] ] ]. + eexists. eexists. ssplit. + 4: ecancel_assumption. + 1,2,3: solve [auto]. + subst v. replace (Z.to_nat (8 mod 2 ^ width - word.unsigned x6)) with + (S (Z.to_nat (8 - word.unsigned (word.add x6 (word.of_Z 1))))). + { cbn [get_inner_leakage]. + rewrite H22. + assert (app_one_cons : forall A (a : A) l, (a :: l = (cons a nil) ++ l)%list). + { reflexivity. } + clear H22. + repeat rewrite <- List.app_assoc. f_equal. + { f_equal. + { instantiate (1 := fun _ _ => _). simpl. reflexivity. } + { instantiate (1 := fun _ _ => _). simpl. reflexivity. } } + instantiate (1 := fun _ _ => _). simpl. align_trace. } + clear H22. rewrite word.unsigned_add. clear H12. cbv [word.wrap]. + rewrite word.unsigned_of_Z. cbv [word.wrap]. rewrite (Z.mod_small 1) by blia. + rewrite (Z.mod_small 8) by blia. rewrite Z.mod_small. + { blia. } + pose proof (word.unsigned_range x6). blia. } + intros. intuition. eexists. eexists. split; [|split; [|split; [|split] ] ]. + 4: ecancel_assumption. + all: auto. + f_equal. + replace (Z.to_nat v) with O. + { simpl. instantiate (1 := (_ :: nil)%list). reflexivity. } + subst v. destruct (word.ltu _ _) eqn:Ex6; try congruence. + rewrite word.unsigned_ltu in Ex6. apply Z.ltb_nlt in Ex6. + rewrite word.unsigned_of_Z in Ex6. cbv [word.wrap] in *. + assert (8 < 2 ^ width). + { assert (X := Z.pow_le_mono_r 2 4 width). specialize (X ltac:(blia) ltac:(blia)). + blia. } + rewrite (Z.mod_small 8) in * by blia. blia. } + { cbn in l. repeat straightline. + eapply dexpr_expr. repeat straightline. letexists; split. + { cbn. rewrite ?Properties.map.get_put_dec; exact eq_refl. } + repeat straightline. eexists. eexists. eexists. split. + { cbv [LeakageLoops.enforce]; cbn. + subst l l0. + repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec; cbn). split. + { exact eq_refl. } + { eapply map.map_ext; intros K0. + repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec, ?map.get_empty; cbn -[String.eqb]). + repeat (destruct String.eqb; trivial). } } + (*postcondition*) + repeat straightline. + eexists. eexists. eexists. split. + { ssplit. 3: ecancel_assumption. all: eauto. } + split. + (*the following block, 'block X', is copied and pasted down below*) + { subst v v1. + assert (8 < 2 ^ width). + { assert (X := Z.pow_le_mono_r 2 4 width). specialize (X ltac:(blia) ltac:(blia)). + blia. } + assert (0 < word.unsigned (word:=word) (word.of_Z 8)). + { rewrite word.unsigned_of_Z. cbv [word.wrap]. + rewrite Z.mod_small by blia. blia. } + remember (word.divu _ _) as cow. + rewrite word.unsigned_add. rewrite word.unsigned_of_Z. + cbv [word.wrap]. rewrite (Z.mod_small 1) by blia. + rewrite (Z.mod_small (word.unsigned x1 + 1)). + { blia. } + pose proof (word.unsigned_range x1). split; try blia. + pose proof (word.unsigned_range cow). blia. } + repeat straightline. eexists. eexists. ssplit. + 3: ecancel_assumption. + 1,2: assumption. + simpl. subst k1. + replace (Z.to_nat v) with (S (Z.to_nat + (word.unsigned (word:=word) (word.divu (word.of_Z KYBER_N) (word.of_Z 8)) - + word.unsigned (word.add x1 (word.of_Z 1))))). + { cbn [get_outer_leakage]. + rewrite H17. clear H17. + assert (app_one_cons : forall A (a : A) l, (a :: l = (cons a nil) ++ l)%list). + { reflexivity. } + simpl. + repeat (rewrite List.app_assoc || rewrite (app_one_cons _ _ (_ ++ k0)%list) || rewrite (app_one_cons _ _ k0)). + f_equal. repeat rewrite <- List.app_assoc. f_equal. + 2: { instantiate (1 := fun _ => _). cbv beta. simpl. reflexivity. } + f_equal. + { instantiate (1 := fun _ => _). simpl. reflexivity. } } + (*block X again*) + { subst v v1. + assert (8 < 2 ^ width). + { assert (X := Z.pow_le_mono_r 2 4 width). specialize (X ltac:(blia) ltac:(blia)). + blia. } + assert (0 < word.unsigned (word:=word) (word.of_Z 8)). + { rewrite word.unsigned_of_Z. cbv [word.wrap]. + rewrite Z.mod_small by blia. blia. } + remember (word.divu _ _) as cow. + rewrite word.unsigned_add. rewrite word.unsigned_of_Z. + cbv [word.wrap]. rewrite (Z.mod_small 1) by blia. + rewrite (Z.mod_small (word.unsigned x1 + 1)). + { blia. } + pose proof (word.unsigned_range x1). split; try blia. + pose proof (word.unsigned_range cow). blia. } } } + intros. intuition. eexists. eexists. ssplit. + 3: ecancel_assumption. + 1,2: assumption. + simpl. replace (Z.to_nat v) with 0%nat. + { cbn [get_outer_leakage]. instantiate (1 := (_ :: _ :: _ :: nil)%list). reflexivity. } + destruct (word.ltu x1 _) eqn:E. + { rewrite word.unsigned_of_Z_1 in H4. congruence. } + rewrite word.unsigned_ltu in E. apply Z.ltb_nlt in E. + blia. } + repeat straightline. + subst k0. + assert (app_one_cons : forall A (a : A) l, (a :: l = (cons a nil) ++ l)%list). + { reflexivity. } + repeat (rewrite List.app_assoc || rewrite (app_one_cons _ _ (_ ++ k)%list)). + instantiate (1 := fun _ _ _ => _). reflexivity. + Qed. +End WithWord. diff --git a/bedrock2/src/bedrock2Examples/memequal.v b/bedrock2/src/bedrock2Examples/memequal.v index 41629648b..145addde6 100644 --- a/bedrock2/src/bedrock2Examples/memequal.v +++ b/bedrock2/src/bedrock2Examples/memequal.v @@ -14,7 +14,8 @@ Definition memequal := func! (x,y,n) ~> r { r = r == $0 }. -Require Import bedrock2.WeakestPrecondition bedrock2.Semantics bedrock2.ProgramLogic. +Require Import bedrock2.LeakageWeakestPrecondition bedrock2.LeakageSemantics. +Require Import bedrock2.LeakageLoops bedrock2.LeakageProgramLogic. Require Import coqutil.Word.Interface coqutil.Word.Bitwidth. Require Import coqutil.Map.Interface bedrock2.Map.SeparationLogic. Require Import bedrock2.ZnWords. @@ -27,14 +28,14 @@ Local Notation "xs $@ a" := (Array.array ptsto (word.of_Z 1) a xs) (at level 10, Section WithParameters. Context {width} {BW: Bitwidth width}. Context {word: word.word width} {mem: map.map word byte} {locals: map.map string word}. - Context {ext_spec: ExtSpec}. - Import ProgramLogic.Coercions. + Context {ext_spec: ExtSpec} {pick_sp: PickSp}. + Import LeakageProgramLogic.Coercions. Global Instance spec_of_memequal : spec_of "memequal" := - fnspec! "memequal" (x y n : word) / (xs ys : list byte) (Rx Ry : mem -> Prop) ~> r, - { requires t m := m =* xs$@x * Rx /\ m =* ys$@y * Ry /\ + fnspec! exists f, "memequal" (x y n : word) / (xs ys : list byte) (Rx Ry : mem -> Prop) ~> r, + { requires k t m := m =* xs$@x * Rx /\ m =* ys$@y * Ry /\ length xs = n :>Z /\ length ys = n :>Z; - ensures t' m' := m=m' /\ t=t' /\ (r = 0 :>Z \/ r = 1 :>Z) /\ + ensures k' t' m' := f x y n ++ k = k' /\ m=m' /\ t=t' /\ (r = 0 :>Z \/ r = 1 :>Z) /\ (r = 1 :>Z <-> xs = ys) }. Context {word_ok: word.ok word} {mem_ok: map.ok mem} {locals_ok : map.ok locals} @@ -43,23 +44,30 @@ Section WithParameters. Import coqutil.Tactics.letexists coqutil.Tactics.Tactics coqutil.Tactics.autoforward. Import coqutil.Word.Properties coqutil.Map.Properties. + Local Fixpoint newtrace x y n := + match n with + | S n' => newtrace (word.add x (word.of_Z 1)) (word.add y (word.of_Z 1)) n' ++ + [leak_word y; leak_word x; leak_bool true] + | O => [] + end. + Local Ltac ZnWords := destruct width_cases; bedrock2.ZnWords.ZnWords. Lemma memequal_ok : program_logic_goal_for_function! memequal. Proof. repeat straightline. - refine ((Loops.tailrec + refine ((LeakageLoops.tailrec (HList.polymorphic_list.cons _ (HList.polymorphic_list.cons _ (HList.polymorphic_list.cons _ (HList.polymorphic_list.cons _ HList.polymorphic_list.nil)))) ["x";"y";"n";"r"]) - (fun (v:nat) xs Rx ys Ry t m x y n r => PrimitivePair.pair.mk ( + (fun (v:nat) xs Rx ys Ry k t m x y n r => PrimitivePair.pair.mk ( m =* xs$@x * Rx /\ m =* ys$@y * Ry /\ v=n :> Z /\ length xs = n :> Z /\ length ys = n :> Z ) - (fun T M (X Y N R : word) => m = M /\ t = T /\ + (fun K T M (X Y N R : word) => leak_bool false :: newtrace x y (Z.to_nat (word.unsigned n)) ++ k = K /\ t = T /\ m = M /\ exists z, R = Z.lor r z :> Z /\ (z = 0 :>Z <-> xs = ys) )) lt @@ -72,19 +80,19 @@ Section WithParameters. List.repeat Datatypes.length HList.polymorphic_list.repeat HList.polymorphic_list.length PrimitivePair.pair._1 PrimitivePair.pair._2] in *. - { cbv [Loops.enforce]; cbn. + { cbv [LeakageLoops.enforce]; cbn. subst l l0. repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec; cbn); split. { exact eq_refl. } - { eapply map.map_ext; intros k. + { eapply map.map_ext; intros kk. repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec, ?map.get_empty; cbn -[String.eqb]). repeat (destruct String.eqb; trivial). } } { eapply Wf_nat.lt_wf. } { cbn; ssplit; eauto. } - { intros ?v ?xs ?Rx ?ys ?Ry ?t ?m ?x ?y ?n ?r. + { intros ?v ?xs ?Rx ?ys ?Ry ?k ?t ?m ?x ?y ?n ?r. repeat straightline. cbn in localsmap. - eexists n0; split; cbv [expr expr_body localsmap get]. + eexists n0; exists k0; split; cbv [dexpr expr expr_body localsmap get]. { rewrite ?Properties.map.get_put_dec. exists n0; cbn. auto. } split; cycle 1. { intros Ht; rewrite Ht in *. @@ -96,7 +104,7 @@ Section WithParameters. cbn [length Array.array] in *; try (cbn in *; congruence); []; repeat straightline. - eapply WeakestPreconditionProperties.dexpr_expr. + eapply LeakageWeakestPreconditionProperties.dexpr_expr. repeat straightline. letexists; split. @@ -113,17 +121,17 @@ Section WithParameters. repeat straightline. repeat straightline. - eapply WeakestPreconditionProperties.dexpr_expr. + eapply LeakageWeakestPreconditionProperties.dexpr_expr. letexists; split. { subst l. rewrite ?Properties.map.get_put_dec; cbn. exact eq_refl. } repeat straightline. - eapply WeakestPreconditionProperties.dexpr_expr. + eapply LeakageWeakestPreconditionProperties.dexpr_expr. letexists; split. { subst l l0. rewrite ?Properties.map.get_put_dec; cbn. exact eq_refl. } repeat straightline. - eapply WeakestPreconditionProperties.dexpr_expr. + eapply LeakageWeakestPreconditionProperties.dexpr_expr. letexists; split. { subst l l0 l1. rewrite ?Properties.map.get_put_dec; cbn. exact eq_refl. } @@ -131,10 +139,10 @@ Section WithParameters. eexists _, _, _, _. split. - { cbv [Loops.enforce l l0 l1 l2]; cbn. + { cbv [LeakageLoops.enforce l l0 l1 l2]; cbn. repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec; cbn); split. { exact eq_refl. } - { eapply map.map_ext; intros k. + { eapply map.map_ext; intros kk. repeat (rewrite ?map.get_put_dec, ?map.get_remove_dec, ?map.get_empty; cbn -[String.eqb]). repeat (destruct String.eqb; trivial). } } eexists _, _, _, _, (length xs0); split; ssplit. @@ -146,11 +154,14 @@ Section WithParameters. split. { cbn in *; ZnWords. } intuition idtac; repeat straightline_cleanup. - rewrite H10, word.unsigned_or_nowrap, <-Z.lor_assoc. + { replace (Z.to_nat n0) with (S (Z.to_nat (word.sub n0 v4))) by ZnWords. + cbn [newtrace List.app]. rewrite <- List.app_assoc. cbn [List.app]. + subst v3 v4. apply H9. } + rewrite H11, word.unsigned_or_nowrap, <-Z.lor_assoc. eexists; split; trivial. transitivity (hxs = hys /\ xs0 = ys0); [|intuition congruence]. - rewrite <-H11. rewrite Z.lor_eq_0_iff. eapply and_iff_compat_r. - subst v0 v1. rewrite word.unsigned_xor_nowrap, Z.lxor_eq_0_iff. + rewrite <-H12. rewrite Z.lor_eq_0_iff. eapply and_iff_compat_r. + subst v1 v2. rewrite word.unsigned_xor_nowrap, Z.lxor_eq_0_iff. split; [|intros;subst;trivial]. intro HH. pose proof byte.unsigned_range hxs; @@ -158,16 +169,19 @@ Section WithParameters. eapply word.unsigned_inj in HH; eapply word.of_Z_inj_small in HH; try ZnWords. eapply byte.unsigned_inj in HH; trivial. } - intuition idtac. case H6 as (?&?&?). subst. subst r. - eapply WeakestPreconditionProperties.dexpr_expr. + intuition idtac. case H7 as (?&?&?). subst. subst r. + eapply LeakageWeakestPreconditionProperties.dexpr_expr. letexists; split; cbn. { rewrite ?Properties.map.get_put_dec; cbn; exact eq_refl. } eexists; split; cbn. { rewrite ?Properties.map.get_put_dec; cbn; exact eq_refl. } - rewrite word.unsigned_of_Z_0, Z.lor_0_l in H5; subst x4 v. + rewrite word.unsigned_of_Z_0, Z.lor_0_l in H6; subst x4 v. setoid_rewrite word.unsigned_eqb; setoid_rewrite word.unsigned_of_Z_0. - eexists; ssplit; eauto; destr Z.eqb; autoforward with typeclass_instances in E; + repeat straightline. split. + { instantiate (1 := fun _ _ _ => _ :: _). simpl. reflexivity. } + repeat straightline. split. + all: destr Z.eqb; autoforward with typeclass_instances in E; rewrite ?word.unsigned_of_Z_1, ?word.unsigned_of_Z_0; eauto. all : intuition eauto; discriminate. Qed. diff --git a/bedrock2/src/bedrock2Examples/stackalloc.v b/bedrock2/src/bedrock2Examples/stackalloc.v index 2920be2d2..5bae4cf91 100644 --- a/bedrock2/src/bedrock2Examples/stackalloc.v +++ b/bedrock2/src/bedrock2Examples/stackalloc.v @@ -18,24 +18,66 @@ Definition stackdisj := func! () ~> (a,b) { /*skip*/ }. -Require bedrock2.WeakestPrecondition. -Require Import bedrock2.Semantics bedrock2.FE310CSemantics. +Definition stackswap := func! (a, b) ~> (b, a) { + stackalloc 4 as x; + store(x, a); + stackalloc 4 as y; + store(y, b); + swap(y, x); + a = load(x); + b = load(y) +}. + +Require Import bedrock2.LeakageWeakestPrecondition. +Require Import bedrock2.Semantics bedrock2.LeakageSemantics bedrock2.FE310CSemantics. Require Import coqutil.Map.Interface bedrock2.Map.Separation bedrock2.Map.SeparationLogic. -Require bedrock2.WeakestPreconditionProperties. +Require Import bedrock2.LeakageWeakestPreconditionProperties. From coqutil.Tactics Require Import letexists eabstract. -Require Import bedrock2.ProgramLogic bedrock2.Scalars. +Require Import bedrock2.LeakageProgramLogic bedrock2.Scalars. Require Import coqutil.Word.Interface. From coqutil.Tactics Require Import reference_to_string . From bedrock2 Require ToCString PrintListByte. +(* where to put all of this? *) +Require Import coqutil.Z.Lia. +Section aLemmaThatDoesntBelongHere. + Context {width: Z} {word: word.word width} {word_ok : word.ok word}. + Lemma word_to_bytes (a : word) : + a = word.of_Z (LittleEndianList.le_combine (LittleEndianList.le_split (Z.to_nat ((width + 7) / 8)) (word.unsigned a))). + Proof. + rewrite LittleEndianList.le_combine_split. rewrite Z.mod_small. + - symmetry. apply word.of_Z_unsigned. + - assert (H := Properties.word.unsigned_range a). destruct H as [H1 H2]. + + split; try apply H1. clear H1. + eapply Z.lt_le_trans; try apply H2. clear H2. + apply Z.pow_le_mono_r; try blia. + rewrite Znat.Z2Nat.id. + + replace ((width + 7) / 8 * 8) with (width + 7 - (width + 7) mod 8). + -- assert (H := Z.mod_pos_bound (width + 7) 8). blia. + -- rewrite Zdiv.Zmod_eq_full; blia. + + apply Z.div_pos; try blia. destruct word_ok. blia. + Qed. + + Lemma word_to_bytes' (a : word) : + exists l, length l = (Z.to_nat ((width + 7) / 8)) /\ + a = word.of_Z (LittleEndianList.le_combine l). + Proof. + eexists. split; try apply word_to_bytes. apply LittleEndianList.length_le_split. + Qed. +End aLemmaThatDoesntBelongHere. + Section WithParameters. Context {word: word.word 32} {mem: map.map word Byte.byte}. Context {word_ok: word.ok word} {mem_ok: map.ok mem}. + Context {pick_sp: PickSp}. + Local Open Scope string_scope. Local Open Scope Z_scope. Local Open Scope list_scope. - Instance spec_of_stacktrivial : spec_of "stacktrivial" := fun functions => forall m t, - WeakestPrecondition.call functions - "stacktrivial" t m [] (fun t' m' rets => rets = [] /\ m'=m /\ t'=t). + Instance spec_of_stacktrivial : spec_of "stacktrivial" := + fun functions => + forall k t m, LeakageWeakestPrecondition.call functions + "stacktrivial" k t m [] (fun k' t' m' rets => rets = [] /\ m'=m /\ t'=t). Lemma stacktrivial_ok : program_logic_goal_for_function! stacktrivial. Proof. @@ -56,9 +98,9 @@ Section WithParameters. intuition congruence. Qed. - Instance spec_of_stacknondet : spec_of "stacknondet" := fun functions => forall m t, - WeakestPrecondition.call functions - "stacknondet" t m [] (fun t' m' rets => exists a b, rets = [a;b] /\ a = b /\ m'=m/\t'=t). + Instance spec_of_stacknondet : spec_of "stacknondet" := fun functions => forall m t k, + LeakageWeakestPrecondition.call functions + "stacknondet" k t m [] (fun k' t' m' rets => exists a b, rets = [a;b] /\ a = b /\ m'=m/\t'=t). Add Ring wring : (Properties.word.ring_theory (word := word)) (preprocess [autorewrite with rew_word_morphism], @@ -83,6 +125,7 @@ Section WithParameters. assert ((Array.array ptsto (word.of_Z 1) a [(Byte.byte.of_Z (word.unsigned v0)); b0; b1; b2] ⋆ R)%sep m1). { cbn [Array.array]. use_sep_assumption; cancel; Morphisms.f_equiv; f_equal; f_equal; ring. } + subst a. seprewrite_in_by @scalar32_of_bytes H0 reflexivity. repeat straightline. seprewrite_in_by (symmetry! @scalar32_of_bytes) H0 reflexivity. @@ -111,9 +154,9 @@ Section WithParameters. Definition stacknondet_c := String.list_byte_of_string (c_module (("main",stacknondet_main)::("stacknondet",stacknondet)::nil)). (* Goal True. print_list_byte stacknondet_c. Abort. *) - Instance spec_of_stackdisj : spec_of "stackdisj" := fun functions => forall m t, - WeakestPrecondition.call functions - "stackdisj" t m [] (fun t' m' rets => exists a b, rets = [a;b] /\ a <> b /\ m'=m/\t'=t). + Instance spec_of_stackdisj : spec_of "stackdisj" := fun functions => forall m t k, + LeakageWeakestPrecondition.call functions + "stackdisj" k t m [] (fun k' t' m' rets => exists a b, rets = [a;b] /\ a <> b /\ m'=m/\t'=t). Lemma stackdisj_ok : program_logic_goal_for_function! stackdisj. Proof. @@ -125,4 +168,48 @@ Section WithParameters. all : try intuition congruence. match goal with |- _ <> _ => idtac end. Abort. + + Instance spec_of_stackswap : spec_of "stackswap" := + fnspec! exists f, "stackswap" a b ~> B A, + { requires k t m := True ; + ensures k' t' m' := k' = f k ++ k }. + + Require Import bedrock2Examples.swap. + + Lemma stackswap_ok : + let swapspec := spec_of_swap in + program_logic_goal_for_function! stackswap. + Proof. + repeat straightline. + set (R := eq m). + pose proof (eq_refl : R m) as Hm. + repeat straightline. + repeat (destruct stack as [|?b stack]; try solve [cbn in H2; Lia.lia]; []). + clear H2. clear length_stack. clear H1. + seprewrite_in @scalar_of_bytes Hm. + { rewrite H3. reflexivity. } + repeat straightline. + repeat (destruct stack as [|?b stack]; try solve [cbn in length_stack0; Lia.lia]; []). + clear H5 H3. + seprewrite_in @scalar_of_bytes H1. + { rewrite H6. reflexivity. } + repeat straightline. + assert (HToBytesa := word_to_bytes' a). + destruct HToBytesa as [la [length_la HToBytesa]]. + repeat (destruct la as [|? la]; try solve [cbn in length_la; Lia.lia]; []). + assert (HToBytesb := word_to_bytes' b). + destruct HToBytesb as [lb [length_lb HToBytesb]]. + repeat (destruct lb as [|? lb]; try solve [cbn in length_lb; Lia.lia]; []). + subst a b. + straightline_call. + { apply sep_assoc. eassumption. } + repeat straightline. + Import symmetry. + seprewrite_in_by (symmetry! @scalar_of_bytes) H7 reflexivity. + straightline_stackdealloc. + seprewrite_in_by (symmetry! @scalar_of_bytes) H7 reflexivity. + straightline_stackdealloc. + repeat straightline. align_trace. + Qed. End WithParameters. + diff --git a/bedrock2/src/bedrock2Examples/swap.v b/bedrock2/src/bedrock2Examples/swap.v index 36878febc..071173fac 100644 --- a/bedrock2/src/bedrock2Examples/swap.v +++ b/bedrock2/src/bedrock2Examples/swap.v @@ -19,49 +19,53 @@ Definition swap_swap := func! (a, b) { swap(a, b) }. -Require Import bedrock2.WeakestPrecondition. -Require Import bedrock2.Semantics bedrock2.FE310CSemantics. +Require Import bedrock2.LeakageWeakestPrecondition. +Require Import bedrock2.Semantics bedrock2.LeakageSemantics bedrock2.FE310CSemantics. Require Import coqutil.Map.Interface bedrock2.Map.Separation bedrock2.Map.SeparationLogic. -Require bedrock2.WeakestPreconditionProperties. +Require bedrock2.LeakageWeakestPreconditionProperties. From coqutil.Tactics Require Import Tactics letexists eabstract. -Require Import bedrock2.ProgramLogic bedrock2.Scalars. +Require Import bedrock2.LeakageProgramLogic bedrock2.Scalars. Require Import coqutil.Word.Interface. Require Import coqutil.Tactics.rdelta. Section WithParameters. Context {word: word.word 32} {mem: map.map word Byte.byte}. Context {word_ok: word.ok word} {mem_ok: map.ok mem}. + Context {pick_sp: PickSp}. Instance spec_of_swap : spec_of "swap" := - fnspec! "swap" a_addr b_addr / a b R, - { requires t m := m =* scalar a_addr a * scalar b_addr b * R; - ensures T M := M =* scalar a_addr b * scalar b_addr a * R /\ T = t }. + fnspec! exists f, "swap" a_addr b_addr / a b R, + { requires k t m := m =* scalar a_addr a * scalar b_addr b * R; + ensures K T M := M =* scalar a_addr b * scalar b_addr a * R /\ T = t /\ + K = f a_addr b_addr ++ k }. Lemma swap_ok : program_logic_goal_for_function! swap. - Proof. repeat straightline; eauto. Qed. + Proof. repeat straightline. intuition eauto. align_trace. Qed. Instance spec_of_bad_swap : spec_of "bad_swap" := - fnspec! "bad_swap" a_addr b_addr / a b R, - { requires t m := m =* scalar a_addr a * scalar b_addr b * R; - ensures T M := M =* scalar a_addr b * scalar b_addr a * R /\ T = t }. + fnspec! exists f, "bad_swap" a_addr b_addr / a b R, + { requires k t m := m =* scalar a_addr a * scalar b_addr b * R; + ensures K T M := M =* scalar a_addr b * scalar b_addr a * R /\ T = t /\ + K = f a_addr b_addr ++ k }. + Lemma bad_swap_ok : program_logic_goal_for_function! bad_swap. - Proof. repeat straightline; eauto. Abort. + Proof. repeat straightline. intuition eauto; align_trace. Abort. Definition spec_of_swap_same : spec_of "swap" := - fnspec! "swap" a_addr b_addr / a R, - { requires t m := m =* scalar a_addr a * R /\ b_addr = a_addr; - ensures T M := M =* scalar a_addr a * R /\ T = t }. + fnspec! exists f, "swap" a_addr b_addr / a R, + { requires k t m := m =* scalar a_addr a * R /\ b_addr = a_addr; + ensures K T M := M =* scalar a_addr a * R /\ T = t /\ K = f a_addr ++ k }. Lemma swap_same_ok : let spec_of_swap := spec_of_swap_same in program_logic_goal_for_function! swap. - Proof. repeat straightline; eauto. Qed. + Proof. repeat straightline. intuition eauto. align_trace. Qed. Instance spec_of_swap_swap : spec_of "swap_swap" := fnspec! "swap_swap" a_addr b_addr / a b R, - { requires t m := m =* scalar a_addr a * scalar b_addr b * R; - ensures T M := M =* scalar a_addr a * scalar b_addr b * R /\ T = t}. + { requires k t m := m =* scalar a_addr a * scalar b_addr b * R; + ensures K T M := M =* scalar a_addr a * scalar b_addr b * R /\ T = t }. Lemma swap_swap_ok : program_logic_goal_for_function! swap_swap. Proof. repeat (straightline || straightline_call); eauto. Qed. diff --git a/compiler/src/compiler/CompilerInvariant.v b/compiler/src/compiler/CompilerInvariant.v index 409847301..786a870d7 100644 --- a/compiler/src/compiler/CompilerInvariant.v +++ b/compiler/src/compiler/CompilerInvariant.v @@ -3,6 +3,7 @@ Require Import coqutil.Map.Interface coqutil.Map.Properties. Require Import coqutil.Word.Interface coqutil.Word.Properties. Require Import coqutil.Byte. Require Import riscv.Utility.bverify. +Require Import riscv.Spec.LeakageOfInstr. Require Import bedrock2.Array. Require Import bedrock2.Map.SeparationLogic. Require Import compiler.SeparationLogic. @@ -35,10 +36,10 @@ Section Pipeline1. Context {mem: map.map word byte}. Context {Registers: map.map Z word}. Context {string_keyed_map: forall T: Type, map.map string T}. (* abstract T for better reusability *) - Context {ext_spec: Semantics.ExtSpec}. + Context {ext_spec: LeakageSemantics.ExtSpec}. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {word_riscv_ok: RiscvWordProperties.word.riscv_ok word}. Context {string_keyed_map_ok: forall T, map.ok (string_keyed_map T)}. @@ -47,10 +48,11 @@ Section Pipeline1. Context {iset: InstructionSet}. Context {BWM: bitwidth_iset width iset}. Context {mem_ok: map.ok mem}. - Context {ext_spec_ok: Semantics.ext_spec.ok ext_spec}. + Context {ext_spec_ok: LeakageSemantics.ext_spec.ok ext_spec}. Context (compile_ext_call : string_keyed_map Z -> Z -> Z -> FlatImp.stmt Z -> list Instruction). + Context (leak_ext_call: word -> string_keyed_map Z -> Z -> Z -> FlatImp.stmt Z -> list word -> list LeakageEvent). Context (compile_ext_call_correct: forall resvars extcall argvars, - compiles_FlatToRiscv_correctly compile_ext_call compile_ext_call + compiles_FlatToRiscv_correctly compile_ext_call leak_ext_call compile_ext_call (FlatImp.SInteract resvars extcall argvars)). Context (compile_ext_call_length_ignores_positions: forall stackoffset posmap1 posmap2 c pos1 pos2, List.length (compile_ext_call posmap1 pos1 stackoffset c) = @@ -191,6 +193,7 @@ Section Pipeline1. initial.(getNextPc) = word.add initial.(getPc) (word.of_Z 4) /\ regs_initialized.regs_initialized initial.(getRegs) /\ initial.(getLog) = nil /\ + initial.(getTrace) = Some nil /\ valid_machine initial. Lemma compiler_invariant_proofs: @@ -201,7 +204,7 @@ Section Pipeline1. exists suff, spec.(goodTrace) (suff ++ st.(getLog))). Proof. ssplit; intros. - - eapply (establish_ll_inv _ compile_ext_call_correct compile_ext_call_length_ignores_positions). + - eapply (establish_ll_inv _ _ compile_ext_call_correct compile_ext_call_length_ignores_positions). 1: assumption. unfold initial_conditions, ToplevelLoop.initial_conditions in *. simp. diff --git a/compiler/src/compiler/DeadCodeElim.v b/compiler/src/compiler/DeadCodeElim.v index eeb9c86ec..7b61bb2e7 100644 --- a/compiler/src/compiler/DeadCodeElim.v +++ b/compiler/src/compiler/DeadCodeElim.v @@ -1,3 +1,4 @@ +Require Import bedrock2.LeakageSemantics. Require Import compiler.FlatImp. Require Import Coq.Lists.List. Import ListNotations. Require Import bedrock2.Syntax. @@ -14,17 +15,17 @@ Require Import bedrock2.MetricCosts. (* below only for of_list_list_diff *) Require Import compiler.DeadCodeElimDef. -Local Notation exec := (exec PreSpill isRegStr). +Local Notation exec e pick_sp := (@exec _ _ _ _ _ _ _ _ PreSpill isRegStr pick_sp e). Section WithArguments1. Context {width: Z}. - Context {BW: Bitwidth.Bitwidth width }. - Context {word : word width } { word_ok : word.ok word }. - Context {env: map.map string (list var * list var * stmt var) } { env_ok : map.ok env }. - 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 } . - + Context {BW: Bitwidth.Bitwidth width}. + Context {word : word width} {word_ok : word.ok word}. + Context {env: map.map string (list var * list var * stmt var)} {env_ok : map.ok env}. + 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: LeakageSemantics.ExtSpec} {ext_spec_ok: LeakageSemantics.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 @@ -109,16 +110,20 @@ Section WithArguments1. rewrite ListSet.of_list_removeb end. - Ltac mcsolve := eexists; split; [|split; cycle 1; [eauto|FlatImp.scost_hammer]]; try assumption. - + Ltac solve_compile_post := + do 5 eexists; ssplit; + [eauto | repeat listset_to_set; agree_on_solve | scost_hammer | align_trace | align_trace | + intros; rewrite dfix_step; simpl_rev; cbv beta; try reflexivity ]. + Lemma dce_correct_aux : - forall eH eL, + forall eH eL pick_spL, dce_functions eH = Success eL -> - forall sH t m mcH lH postH, - exec eH sH t m lH mcH postH -> - forall used_after lL mcL, + forall pick_spH sH kH t m mcH lH postH, + exec pick_spH eH sH kH t m lH mcH postH -> + forall f used_after kL 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 mcH mcL used_after postH). + (forall k, pick_spH (k ++ kH) = pick_spL (rev kL ++ stmt_leakage eH (rev k, sH, used_after, f k))) -> + exec (fun k => pick_spL (rev k)) eL (dce sH used_after) kL t m lL mcL (compile_post eH sH kH kL mcH mcL used_after postH). Proof. induction 2; match goal with @@ -131,20 +136,17 @@ Section WithArguments1. * eapply H1. * repeat listset_to_set. agree_on_solve. + intros. - eapply H3 in H5. + eapply H3 in H6. fwd. let Heq := fresh in - pose proof H5p0 as Heq; + pose proof H6p0 as Heq; eapply map.putmany_of_list_zip_sameLength, map.sameLength_putmany_of_list in Heq. fwd. eexists. split. - * eapply H5. - * intros. - unfold compile_post. - exists l'. mcsolve. - agree_on_solve. repeat listset_to_set. - subset_union_solve. + * eapply H6. + * intros. solve_compile_post. + repeat listset_to_set. subset_union_solve. - intros. eapply @exec.call; try solve [ eassumption ]. + unfold dce_functions, dce_function in *. @@ -155,37 +157,33 @@ Section WithArguments1. * eapply H1. * listset_to_set. agree_on_solve. + eapply IHexec. - eapply agree_on_refl. + -- eapply agree_on_refl. + -- intros. rewrite associate_one_left. rewrite H6. rewrite dfix_step. + simpl. rewrite rev_app_distr. simpl. rewrite H0. + rewrite <- app_assoc. reflexivity. + intros. unfold compile_post in *. - fwd. eapply H4 in H6p2. fwd. + fwd. eapply H4 in H7p0. fwd. let Heq := fresh in - pose proof H6p2p1 as Heq; - eapply map.putmany_of_list_zip_sameLength, map.sameLength_putmany_of_list in H6p2p1. fwd. + pose proof H7p0p1 as Heq; + eapply map.putmany_of_list_zip_sameLength, map.sameLength_putmany_of_list in H7p0p1. fwd. exists retvs. eexists. repeat split. * erewrite agree_on_getmany. - -- eapply H6p2p0. + -- eapply H7p0p0. -- listset_to_set. agree_on_solve. - * eapply H6p2p1. - * eexists. mcsolve. - agree_on_solve. - repeat listset_to_set. - subset_union_solve. + * eapply H7p0p1. + * solve_compile_post. + -- agree_on_solve. repeat listset_to_set. subset_union_solve. + -- simpl. rewrite H0. rewrite H7p5. reflexivity. - intros. eapply agree_on_find in H3; fwd. destr (existsb (eqb x) used_after); fwd. + eapply @exec.load. * rewrite <- H3p1. eassumption. * eauto. - * unfold compile_post. - exists (map.put l x v); mcsolve. - repeat listset_to_set. - agree_on_solve. + * solve_compile_post. simpl. rewrite E. reflexivity. + eapply @exec.skip. - * unfold compile_post. - exists (map.put l x v); mcsolve. - repeat listset_to_set. - agree_on_solve. + * solve_compile_post. simpl. rewrite E. reflexivity. - intros. repeat listset_to_set. eapply agree_on_union in H4; fwd. all: try solve [ eauto using String.eqb_spec ]. @@ -195,43 +193,34 @@ 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; mcsolve. + + solve_compile_post. - 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; eexists; mcsolve. - repeat listset_to_set; agree_on_solve. + * solve_compile_post. simpl. rewrite E. reflexivity. + eapply @exec.skip; eauto. - unfold compile_post. - eexists; mcsolve. - repeat listset_to_set; agree_on_solve. + solve_compile_post. simpl. rewrite E. reflexivity. - intros. repeat listset_to_set. eapply @exec.stackalloc. * eassumption. - * intros. eapply H2 with (used_after := used_after) (lL := (map.put lL x a)) in H4. - 2: eapply H5. - 2: { agree_on_solve. } - eapply @exec.weaken. - -- eapply H4. - -- unfold compile_post. intros. fwd. exists mSmall'. exists mStack'. split. - ++ eassumption. - ++ split. - ** eassumption. - ** eexists; mcsolve; eauto. + * intros. assert (H4' := H4 nil). rewrite dfix_step in H4'. simpl in H4'. + rewrite app_nil_r in H4'. rewrite H4' in H2. + eapply H2 with (used_after := used_after) (lL := (map.put lL x a)) in H5; subst a. + 2: eapply H6. + -- eapply exec.weaken. 1: eapply H5. intros. + unfold compile_post in H7. fwd. eexists. eexists. + split; [eassumption|]. split; [eassumption|]. + solve_compile_post. simpl. rewrite H7p5. reflexivity. + -- agree_on_solve. + -- intros. rewrite associate_one_left. rewrite H4. rewrite dfix_step. simpl. + rewrite rev_app_distr. simpl. repeat Tactics.destruct_one_match. + repeat rewrite <- app_assoc. reflexivity. - intros. destr (existsb (eqb x) used_after). - + eapply @exec.lit. - unfold compile_post. - repeat listset_to_set. - eexists; mcsolve. - agree_on_solve. - + eapply @exec.skip. - unfold compile_post. - repeat listset_to_set. - eexists; mcsolve. - agree_on_solve. + + eapply @exec.lit. solve_compile_post. + + eapply @exec.skip. solve_compile_post. - destr z. + intros. repeat listset_to_set. eapply agree_on_union in H3; try solve [ eauto using String.eqb_spec ]. @@ -247,39 +236,28 @@ Section WithArguments1. unfold elem_of; destr ((y =? v)%string). ++ eapply in_eq. ++ eapply in_cons, in_eq. - -- unfold compile_post. - eexists; mcsolve. - agree_on_solve. - * eapply @exec.skip. - unfold compile_post. - eexists; mcsolve. - agree_on_solve. + -- solve_compile_post. simpl. rewrite E. + Tactics.destruct_one_match; reflexivity. + * eapply @exec.skip. solve_compile_post. simpl. rewrite E. + Tactics.destruct_one_match; reflexivity. + intros. eapply agree_on_find in H3; fwd. destr (existsb (eqb x) used_after). * eapply @exec.op. -- rewrite <- H3p1. eassumption. -- simpl. constructor. - -- unfold compile_post. simpl in *. inversion H1. fwd. eexists; mcsolve. - repeat listset_to_set. - agree_on_solve. - * eapply @exec.skip. unfold compile_post. - eexists; mcsolve. - repeat listset_to_set. - agree_on_solve. + -- simpl in *. inversion H1. subst. solve_compile_post. + simpl. rewrite E. Tactics.destruct_one_match; reflexivity. + * eapply @exec.skip. solve_compile_post. + simpl. rewrite E. Tactics.destruct_one_match; reflexivity. - intros. eapply agree_on_find in H2; fwd. repeat listset_to_set. destr (existsb (eqb x) used_after). { eapply @exec.set. - rewrite <- H2p1; eassumption. - - unfold compile_post. eexists; mcsolve. - agree_on_solve. - } - { eapply @exec.skip. - - unfold compile_post. eexists; mcsolve. - agree_on_solve. - } + - solve_compile_post. } + { eapply @exec.skip. solve_compile_post. } - intros. repeat listset_to_set. eapply agree_on_union in H2; fwd. @@ -287,10 +265,12 @@ Section WithArguments1. eapply @exec.if_true. + erewrite agree_on_eval_bcond; [ eassumption | ]. pose agree_on_comm; eauto. - + eapply @exec.weaken; [eauto|]. - unfold compile_post. - intros * (?&?&?&?&?). - eexists. mcsolve. + + eapply @exec.weaken. + -- eapply IHexec; eauto. intros. rewrite associate_one_left. rewrite H3. + rewrite dfix_step. rewrite rev_app_distr. simpl. + repeat Tactics.destruct_one_match. rewrite <- app_assoc. reflexivity. + -- unfold compile_post. intros. fwd. solve_compile_post. + simpl. rewrite H2p6. reflexivity. - intros. repeat listset_to_set. eapply agree_on_union in H2; fwd. @@ -298,89 +278,88 @@ Section WithArguments1. eapply @exec.if_false. + erewrite agree_on_eval_bcond; [ eassumption | ]. pose agree_on_comm; eauto. - + eapply @exec.weaken; [eauto|]. - unfold compile_post. - intros * (?&?&?&?&?). - eexists. mcsolve. + + eapply @exec.weaken. + -- eapply IHexec; eauto. intros. rewrite associate_one_left. rewrite H3. + rewrite dfix_step. rewrite rev_app_distr. simpl. + repeat rewrite <- app_assoc. repeat Tactics.destruct_one_match. + reflexivity. + -- unfold compile_post. intros. fwd. solve_compile_post. + simpl. rewrite H2p6. reflexivity. - 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 mc mcL (live (SLoop body1 cond body2) used_after) mid2). + eapply exec.loop_cps. + eapply exec.weaken. { eapply IH1. - eapply agree_on_subset. - - let Heq := fresh in - specialize (live_while body1 cond body2 used_after) as Heq; cbn zeta in Heq. - eapply H4. - - eapply H7. - } + - eapply agree_on_subset. + + let Heq := fresh in + specialize (live_while body1 cond body2 used_after) as Heq; cbn zeta in Heq. + eapply H4. + + eapply H7. + - intros. rewrite H8. rewrite dfix_step. reflexivity. } + intros. unfold compile_post in *. fwd. + specialize H1 with (1 := H4p0). specialize H2 with (1 := H4p0). + assert (H4: eval_bcond lH' cond = eval_bcond l' cond). { intros. unfold compile_post in *. - repeat destr H4. destr H6. - eapply H1 in H8. - erewrite agree_on_eval_bcond; [ eassumption | ]. - eapply agree_on_comm. + fwd. + apply agree_on_eval_bcond. repeat listset_to_set. - eapply agree_on_union in H4. + apply agree_on_union in H4p1. fwd. - eapply agree_on_subset; [ | eapply H4p1 ]. + eapply agree_on_subset; [ | eapply H4p1p1 ]. subset_union_solve. } - { intros. - unfold compile_post in *. - 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 ]. - subset_union_solve. - } - { - intros. - unfold compile_post in *. - 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. } - 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. destr H6. - eapply @exec.weaken; [eapply IH12|]. - - eapply H8. - - eapply H4. - - cbv beta. intros * (?&?&?&?&?). - eexists. mcsolve. - } + rewrite -> H4 in *. + destruct (eval_bcond l' cond) as [b|] eqn:E; try congruence. exists b. + split; [reflexivity|]. split. + { intros. subst. solve_compile_post. rewrite <- app_assoc. simpl. rewrite H4p5. + cbv [FixEq.Let_In_pf_nd]. rewrite List.skipn_app_r by reflexivity. + simpl. rewrite <- app_assoc. reflexivity. } + { intros. subst. + eapply exec.weaken. + { eapply IH2; eauto. + - repeat listset_to_set. + eapply agree_on_subset; [ | eapply H4p1 ]. + subset_union_solve. + - intros. rewrite associate_one_left. rewrite app_assoc. rewrite H8. + repeat (rewrite rev_app_distr || cbn [rev List.app] || rewrite <- app_assoc). + rewrite dfix_step. cbn [stmt_leakage_body]. rewrite H4p5. + cbv [FixEq.Let_In_pf_nd]. rewrite List.skipn_app_r by reflexivity. + reflexivity. + } + cbv beta. intros. fwd. + eapply exec.weaken. + { eapply IH12; [eassumption|eassumption|]. + intros. rewrite associate_one_left. repeat rewrite app_assoc. rewrite H8. + repeat (rewrite rev_app_distr || cbn [rev List.app] || rewrite <- app_assoc). + f_equal. f_equal. rewrite dfix_step. cbn [stmt_leakage_body]. rewrite H4p5. + cbv [FixEq.Let_In_pf_nd]. rewrite List.skipn_app_r by reflexivity. + f_equal. f_equal. rewrite H6p5. rewrite List.skipn_app_r by reflexivity. + reflexivity. } + cbv beta. intros. fwd. solve_compile_post. cbn [stmt_leakage_body]. + repeat (rewrite rev_app_distr || cbn [rev List.app] || rewrite <- app_assoc). + rewrite H4p5. f_equal. cbv [FixEq.Let_In_pf_nd]. + rewrite List.skipn_app_r by reflexivity. f_equal. rewrite H6p5. + rewrite List.skipn_app_r by reflexivity. rewrite H6p9. reflexivity. } - intros. eapply @exec.seq. - + eapply IHexec. eassumption. + + eapply IHexec; [eassumption|]. + intros. rewrite H4. rewrite dfix_step. reflexivity. + unfold compile_post. intros. fwd. eapply @exec.weaken. - * eapply H2. - -- eassumption. - -- eassumption. - * unfold compile_post. intros. fwd. - eexists. mcsolve. - - intros. - eapply @exec.skip. - unfold compile_post. eexists. mcsolve. + * eapply H2; [eassumption|eassumption|]. + intros. rewrite app_assoc. rewrite H4. + repeat rewrite rev_app_distr. rewrite dfix_step. simpl. + rewrite H5p5. rewrite List.skipn_app_r by reflexivity. + rewrite <- app_assoc. reflexivity. + * unfold compile_post. intros. fwd. solve_compile_post. + simpl. rewrite <- app_assoc. rewrite H5p5. + rewrite List.skipn_app_r by reflexivity. rewrite H5p9. + repeat rewrite <- app_assoc. reflexivity. + - intros. eapply @exec.skip. solve_compile_post. Qed. End WithArguments1. diff --git a/compiler/src/compiler/DeadCodeElimDef.v b/compiler/src/compiler/DeadCodeElimDef.v index 5e905c4c1..bd8c9c526 100644 --- a/compiler/src/compiler/DeadCodeElimDef.v +++ b/compiler/src/compiler/DeadCodeElimDef.v @@ -1,3 +1,6 @@ +Require Import Coq.Init.Wf Wellfounded. +Require Import compiler.FixEq. +Require Import bedrock2.LeakageSemantics. Require Import compiler.FlatImp. Require Import Coq.Lists.List. Import ListNotations. Require Import bedrock2.Syntax. @@ -22,7 +25,8 @@ Section WithArguments1. Context {env: map.map string (list var * list var * stmt var) } { env_ok : map.ok env }. 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 } . + Context {ext_spec : LeakageSemantics.ExtSpec } {ext_spec_ok: LeakageSemantics.ext_spec.ok ext_spec } . + Context {pick_sp : LeakageSemantics.PickSp}. Ltac subset_union_solve := match goal with | |- subset (union _ _) _ => eapply subset_union_l; subset_union_solve @@ -465,12 +469,13 @@ Section WithArguments1. SSkip | SStackalloc x n body => SStackalloc x n (dce body u) - (* The below optimization probably can be made to work; - on past attempt, got stuck at goals about `Memory.anybytes n a map.empty` *) - (* if (existsb (eqb x) (live body u)) then - SStackalloc x n (dce body u) - else - (dce body u) *) + (* Finite address spaces are subtle, so the below optimization is unsound. See + https://github.com/mit-plv/bedrock2/pull/393#discussion_r1685645179*) + + (* if (existsb (eqb x) (live body u)) then + SStackalloc x n (dce body u) + else + (dce body u) *) | SLit x _ => if (existsb (eqb x) u) then s @@ -502,6 +507,191 @@ Section WithArguments1. | SSkip => SSkip end. + Definition tuple : Type := leakage * stmt var * list var * (leakage (*skip*) -> leakage (*everything*)). + Definition project_tuple (tup : tuple) := + let '(kH, s, u, f) := tup in (length kH, s). + Definition lt_tuple (x y : tuple) := + lt_tuple' (project_tuple x) (project_tuple y). + + Lemma lt_tuple_wf : well_founded lt_tuple. + Proof. + cbv [lt_tuple]. apply wf_inverse_image. apply lt_tuple'_wf. + Defined. + + Definition stmt_leakage_body + (e: env) + (tup : leakage * stmt var * list var * (leakage -> leakage)) + (dtransform_stmt_trace : forall othertup, lt_tuple othertup tup -> leakage) + : leakage. + refine ( + match tup as x return tup = x -> _ with + | (kH, s, u, f) => + fun _ => + match s as x return s = x -> _ with + | SInteract _ _ _ => + fun _ => + match kH with + | leak_list l :: kH' => leak_list l :: f [leak_list l] + | _ => nil + end + | SCall _ fname _ => + fun _ => + match kH as x return kH = x -> _ with + | leak_unit :: kH' => + fun _ => + match map.get e fname with + | Some (params, rets, fbody) => + leak_unit :: dtransform_stmt_trace (kH', fbody, rets, + fun skip => f (leak_unit :: skip)) _ + | None => nil + end + | _ => fun _ => nil + end eq_refl + | SLoad _ x _ _ => + fun _ => + match kH with + | leak_word addr :: kH' => + if (existsb (eqb x) u) then + leak_word addr :: f [leak_word addr] + else + f [leak_word addr] + | _ => nil + end + | SStore _ _ _ _ => + fun _ => + match kH with + | leak_word addr :: kH' => leak_word addr :: f [leak_word addr] + | _ => nil + end + | SInlinetable _ x _ _ => + fun _ => + match kH with + | leak_word i :: kH' => + if (existsb (eqb x) u) then + leak_word i :: f [leak_word i] + else + f [leak_word i] + | _ => nil + end + | SStackalloc x n body => + fun _ => + match kH as x return kH = x -> _ with + | leak_unit :: kH' => + fun _ => leak_unit :: dtransform_stmt_trace (kH', body, u, + fun skip => f (leak_unit :: skip)) _ + | _ => fun _ => nil + end eq_refl + | SLit x _ => fun _ => f nil + | SOp x op _ _ => + fun _ => + (*copied from spilling. + Maybe I should do a leak_list for ops (or just make every op leak two words) + to avoid this awkwardness*) + let skip := + match op with + | Syntax.bopname.divu + | Syntax.bopname.remu => + match kH with + | leak_word x1 :: leak_word x2 :: k' => [leak_word x1; leak_word x2] + | _ => nil + end + | Syntax.bopname.slu + | Syntax.bopname.sru + | Syntax.bopname.srs => + match kH with + | leak_word x2 :: kH' => [leak_word x2] + | _ => nil + end + | _ => nil + end + in + if (existsb (eqb x) u) then + skip ++ f skip + else + f skip + | SSet _ _ => fun _ => f nil + | SIf _ thn els => + fun _ => + match kH as x return kH = x -> _ with + | leak_bool b :: kH' => + fun _ => leak_bool b :: + dtransform_stmt_trace (kH', if b then thn else els, u, + fun skip => f (leak_bool b :: skip)) _ + | _ => fun _ => nil + end eq_refl + | SLoop s1 c s2 => + fun _ => + let live_before := live (SLoop s1 c s2) u in + dtransform_stmt_trace (kH, s1, (list_union String.eqb + (live s2 live_before) + (list_union eqb (accessed_vars_bcond c) u)), + fun skip1 => + Let_In_pf_nd (List.skipn (length skip1) kH) + (fun kH' _ => + match kH' as x return kH' = x -> _ with + | leak_bool true :: kH'' => + fun _ => leak_bool true :: + dtransform_stmt_trace (kH'', s2, live_before, + fun skip2 => + let kH''' := List.skipn (length skip2) kH'' in + dtransform_stmt_trace (kH''', s, u, + fun skip3 => + f (skip1 ++ [leak_bool true] ++ skip2 ++ skip3)) _) _ + | leak_bool false :: kH'' => + fun _ => leak_bool false :: f (skip1 ++ [leak_bool false]) + | _ => fun _ => nil + end eq_refl)) _ + | SSeq s1 s2 => + fun _ => dtransform_stmt_trace (kH, s1, live s2 u, + fun skip1 => + let kH' := List.skipn (length skip1) kH in + dtransform_stmt_trace (kH', s2, u, fun skip2 => f (skip1 ++ skip2)) _) _ + | SSkip => fun _ => f nil + end eq_refl + end eq_refl). + Proof. + Unshelve. + all: cbv [lt_tuple project_tuple]. + all: subst. + all: repeat match goal with + | t := List.skipn ?n ?k |- _ => + let H := fresh "H" in + assert (H := List.skipn_length n k); subst t end. + all: try (left; simpl; blia). + all: try (right; constructor; constructor). + - assert (H' := skipn_length (length skip1) kH). + rewrite e3 in *. simpl in *. left. blia. + - assert (H' := skipn_length (length skip1) kH). + rewrite e3 in *. simpl in *. left. blia. + - destruct (length (List.skipn (length skip1) kH) =? length kH)%nat eqn:E. + + apply Nat.eqb_eq in E. rewrite E. right. constructor. constructor. + + apply Nat.eqb_neq in E. left. blia. + Defined. + + Definition stmt_leakage e := + Fix lt_tuple_wf _ (stmt_leakage_body e). + + Definition Equiv (x1 x2 : tuple) : Prop := + let '(k1, s1, u1, f1) := x1 in + let '(k2, s2, u2, f2) := x2 in + k1 = k2 /\ s1 = s2 /\ u1 = u2 /\ (forall y, f1 y = f2 y). + + Lemma Equiv_refl x : Equiv x x. + Proof. unfold Equiv. repeat Tactics.destruct_one_match. auto. Qed. + + Lemma dfix_step e tup : stmt_leakage e tup = stmt_leakage_body e tup (fun y _ => stmt_leakage e y). + Proof. + cbv [stmt_leakage]. + apply (@Fix_eq'_nondep _ _ lt_tuple_wf _ (stmt_leakage_body e) Equiv eq). + 2: { apply Equiv_refl. } + intros. clear tup. cbv [stmt_leakage_body]. cbv beta. + destruct x1 as [[[k1 s1] u1] c1]. destruct x2 as [[[k2 s2] u2] c2]. + cbv [Equiv] in H. fwd. subst. rename Hp3 into H. repeat rewrite H. + repeat (Tactics.destruct_one_match || rewrite H || apply H0 || cbv [Equiv] || intuition auto || match goal with | |- _ :: _ = _ :: _ => f_equal end || intuition auto(*why does putting this here make this work*)). + apply Let_In_pf_nd_ext. + repeat (Tactics.destruct_one_match || rewrite H || apply H0 || cbv [Equiv] || intuition auto || match goal with | |- _ :: _ = _ :: _ => f_equal end || intuition auto). + Qed. + Definition dce_function: (list string * list string * stmt string ) -> @@ -518,16 +708,19 @@ Section WithArguments1. Definition compile_post - mcH mcL used_after - (postH: Semantics.trace -> mem -> locals -> MetricLog -> Prop) + e s kH kL mcH mcL used_after + (postH: leakage -> Semantics.trace -> mem -> locals -> MetricLog -> Prop) : - Semantics.trace -> mem -> locals -> MetricLog -> Prop + leakage -> Semantics.trace -> mem -> locals -> MetricLog -> Prop := - (fun t' m' lL' mcL' => - exists lH' mcH', - map.agree_on (PropSet.of_list used_after) lH' lL' + (fun k' t' m' lL' mcL' => + exists kH' kH'' kL'' lH' mcH', + postH kH' t' m' lH' mcH' + /\ map.agree_on (PropSet.of_list used_after) lH' lL' /\ metricsLeq (mcL' - mcL) (mcH' - mcH) - /\ postH t' m' lH' mcH'). + /\ k' = kL'' ++ kL + /\ kH' = kH'' ++ kH + /\ forall kH''' f, stmt_leakage e (rev kH'' ++ kH''', s, used_after, f) = rev kL'' ++ f (rev kH'')). Lemma agree_on_eval_bcond: forall cond (m1 m2: locals), diff --git a/compiler/src/compiler/ExprImpEventLoopSpec.v b/compiler/src/compiler/ExprImpEventLoopSpec.v index 9fd056417..790dd1ca6 100644 --- a/compiler/src/compiler/ExprImpEventLoopSpec.v +++ b/compiler/src/compiler/ExprImpEventLoopSpec.v @@ -10,7 +10,7 @@ Require Import compiler.Pipeline. Section Params1. 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}. + Context {ext_spec: LeakageSemantics.ExtSpec}. Set Implicit Arguments. @@ -24,8 +24,8 @@ Section Params1. isReady: Semantics.trace -> mem -> Prop; }. - Definition hl_inv(spec: ProgramSpec)(t: Semantics.trace)(m: mem) - (l: locals)(mc: bedrock2.MetricLogging.MetricLog) + Definition hl_inv(spec: ProgramSpec)(k: LeakageSemantics.leakage)(t: Semantics.trace) + (m: mem)(l: locals)(mc: bedrock2.MetricLogging.MetricLog) : Prop := (* Restriction: no locals can be shared between loop iterations *) spec.(isReady) t m /\ spec.(goodTrace) t. @@ -40,12 +40,12 @@ Section Params1. get_init_code: map.get (map.of_list e : env) init_f = Some (nil, nil, init_code); init_code_correct: forall m0 mc0, mem_available spec.(datamem_start) spec.(datamem_pastend) m0 -> - MetricSemantics.exec (map.of_list e) init_code nil m0 map.empty mc0 (hl_inv spec); + (forall pick_spH : LeakageSemantics.PickSp, MetricLeakageSemantics.exec (map.of_list e) init_code nil nil m0 map.empty mc0 (hl_inv spec)); loop_body: Syntax.cmd.cmd; get_loop_body: map.get (map.of_list e : env) loop_f = Some (nil, nil, loop_body); - loop_body_correct: forall t m l mc, - hl_inv spec t m l mc -> - MetricSemantics.exec (map.of_list e) loop_body t m l mc (hl_inv spec); + loop_body_correct: forall k t m l mc, + hl_inv spec k t m l mc -> + (forall pick_spH : LeakageSemantics.PickSp, MetricLeakageSemantics.exec (map.of_list e) loop_body k t m l mc (hl_inv spec)); }. End Params1. diff --git a/compiler/src/compiler/FixEq.v b/compiler/src/compiler/FixEq.v new file mode 100644 index 000000000..2b0f5dab7 --- /dev/null +++ b/compiler/src/compiler/FixEq.v @@ -0,0 +1,145 @@ +(* In this file, we prove various strengthenings of Fix_eq. + Most of the file is nearly-verbatim copied from the stdlib file where Fix_eq is proven. + The main goal is to get general recursion even when some arguments to the recursive function is a function (without using funext). *) + +(*Here we have, supposedly, maximum generality: dependent types and arbitrary relations*) +Section GeneralFixPoint. + Variable A : Type. + Variable R : A -> A -> Prop. + Hypothesis Rwf : well_founded R. + Variable P : A -> Type. + Variable F : forall x:A, (forall y:A, R y x -> P y) -> P x. + Variable E1 : A -> A -> Prop. + Variable E2 : forall {a1 a2}, P a1 -> P a2 -> Prop. + + Let Fix_F := @Fix_F A R P F. + + Lemma Fix_F_eq' (x:A) (r:Acc R x) : + F x (fun (y:A) (p:R y x) => Fix_F y (Acc_inv r p)) = Fix_F x r. + Proof. + destruct r using Acc_inv_dep; auto. + Qed. + + Let Fix := @Fix A R Rwf P F. + + Hypothesis + F_ext : + forall (x1 x2:A) (f1:forall y:A, R y x1 -> P y) (f2:forall y:A, R y x2 -> P y), + E1 x1 x2 -> + (forall (y1 y2:A) (p1:R y1 x1) (p2:R y2 x2), + E1 y1 y2 -> E2 (f1 y1 p1) (f2 y2 p2)) -> E2 (F x1 f1) (F x2 f2). + + Lemma Fix_F_inv' : forall (x1 x2:A) (r1:Acc R x1) (r2:Acc R x2), + E1 x1 x2 -> E2 (Fix_F x1 r1) (Fix_F x2 r2). + Proof. + intro x1; induction (Rwf x1); intros x2 r1 r2. + rewrite <- (Fix_F_eq' x r1); rewrite <- (Fix_F_eq' x2 r2); intros. + apply F_ext; auto. + Qed. + + Lemma Fix_eq' : forall (x1 x2:A), + E1 x1 x2 -> E2 (Fix x1) (F x2 (fun (y:A) (p:R y x2) => Fix y)). + Proof. + intro x; unfold Fix, Wf.Fix. + rewrite <- Fix_F_eq'. + intros. apply F_ext; intros. + - assumption. + - apply Fix_F_inv'. assumption. + Qed. +End GeneralFixPoint. + +(*Here we have arbitrary relations but no dependent types.*) +Section NonDepFixPoint. + Variable A : Type. + Variable R : A -> A -> Prop. + Hypothesis Rwf : well_founded R. + Variable P : Type. + Variable F : forall x:A, (forall y:A, R y x -> P) -> P. + Variable E1 : A -> A -> Prop. + Variable E2 : P -> P -> Prop. + + Hypothesis + F_ext : + forall (x1 x2:A) (f1:forall y:A, R y x1 -> P) (f2:forall y:A, R y x2 -> P), + E1 x1 x2 -> + (forall (y1 y2:A) (p1:R y1 x1) (p2:R y2 x2), + E1 y1 y2 -> E2 (f1 y1 p1) (f2 y2 p2)) -> E2 (F x1 f1) (F x2 f2). + + Definition Fix_eq'_nondep := Fix_eq' A R Rwf (fun _ => P) F E1 (fun _ _ => E2) F_ext. +End NonDepFixPoint. + +(*Here we have dependent types but the relations are eq. + Sadly we can't simply deduce these lemmas from the supposedly-general case, + since there's no good way of instantiating E2 to get eq.*) +Section EqualityFixPoint. + Variable A : Type. + Variable R : A -> A -> Prop. + Hypothesis Rwf : well_founded R. + Variable P : A -> Type. + Variable F : forall x:A, (forall y:A, R y x -> P y) -> P x. + + Hypothesis F_ext : + (forall (x : A) (f g : forall y : A, R y x -> P y), + (forall (y : A) (p1 p2 : R y x), f y p1 = g y p2) -> F x f = F x g). + + Lemma Fix_F_inv'_eq : + forall (x : A) (r s : Acc R x), Fix_F P F r = Fix_F P F s. + Proof. + intros x. induction (Rwf x). intros r s. + rewrite <- Fix_F_eq. rewrite <- Fix_F_eq. apply F_ext. auto. + Qed. + + (*Fix_eq'_eq has the same generality as Fix_eq, but it's a bit stronger + (and often more convenient). + For instance, auto can prove that Fix_eq'_eq implies Fix_eq, but it can't prove + the converse.*) + Lemma Fix_eq'_eq : + forall x : A, Fix Rwf P F x = F x (fun (y : A) (_ : R y x) => Fix Rwf P F y). + Proof. + intros. unfold Fix. rewrite <- Fix_F_eq. + apply F_ext. intros. apply Fix_F_inv'_eq. + Qed. +End EqualityFixPoint. + +Section EqualityNonDepFixPoint. + Variable A : Type. + Variable R : A -> A -> Prop. + Hypothesis Rwf : well_founded R. + Variable P : Type. + Variable F : forall x:A, (forall y:A, R y x -> P) -> P. + + Hypothesis F_ext : + (forall (x : A) (f g : forall y : A, R y x -> P), + (forall (y : A) (p1 p2 : R y x), f y p1 = g y p2) -> F x f = F x g). + + Definition Fix_eq'_eq_nondep := Fix_eq'_eq A R Rwf (fun _ => P) F F_ext. +End EqualityNonDepFixPoint. + +(*the rest of this file is random things that are useful sometimes + for writing recursive functions*) + +Definition Let_In_pf_nd {A B} (x : A) (f : forall a : A, a = x -> B) : B := let y := x in f y eq_refl. + +Lemma Let_In_pf_nd_ext {A B} (E : B -> B -> Prop) (x : A) (f1 f2 : forall a : A, a = x -> B) : + (forall x1 x2, E (f1 x1 x2) (f2 x1 x2)) -> + E (Let_In_pf_nd x f1) (Let_In_pf_nd x f2). +Proof. intros H. cbv [Let_In_pf_nd]. apply H. Qed. + +Require Import bedrock2.Semantics. +Require Import compiler.FlatImp. +Require Import Coq.ZArith.ZArith. +Require Import Coq.Init.Wf Relation_Operators Wellfounded. +Require Import riscv.Utility.Utility. + +Section WithWord. + Context {var : Type} {width} {BW: Bitwidth width} {word: word.word width}. + + Definition lt_tuple' : nat * stmt var -> nat * stmt var -> Prop := slexprod _ _ lt stmt_lt. + + Lemma lt_tuple'_wf : well_founded lt_tuple'. + Proof. + apply wf_slexprod. + - apply lt_wf. + - apply wf_stmt_lt. + Defined. +End WithWord. diff --git a/compiler/src/compiler/FlatImp.v b/compiler/src/compiler/FlatImp.v index 403c4443a..3283a2ed6 100644 --- a/compiler/src/compiler/FlatImp.v +++ b/compiler/src/compiler/FlatImp.v @@ -1,4 +1,6 @@ Require Import Coq.Bool.Bool. +Require Import Coq.Relations.Relation_Operators. +Require Import Coq.Wellfounded.Transitive_Closure. Require Import Coq.ZArith.ZArith. Require Import Coq.Lists.List. Import ListNotations. Require Import bedrock2.MetricLogging. @@ -13,10 +15,13 @@ Require Import bedrock2.Syntax. Require Import coqutil.Z.Lia. Require Import coqutil.Tactics.Simp. Require Import bedrock2.Semantics. +Require Import bedrock2.LeakageSemantics. +Require Import bedrock2.MetricLeakageSemantics. Require Import coqutil.Datatypes.ListSet. Require Import coqutil.Map.OfListWord. Require Import coqutil.Word.Bitwidth. Require Import coqutil.Word.Interface. +Require Import coqutil.Tactics.fwd. Local Hint Mode Word.Interface.word - : typeclass_instances. Inductive bbinop: Set := @@ -87,6 +92,30 @@ Section Syntax. try apply Z.div_pos; blia. Qed. + Inductive subexpression : stmt -> stmt -> Prop := + | SStackalloc_subexp : forall x1 x2 s, subexpression s (SStackalloc x1 x2 s) + | SIf_then_subexp : forall x1 x2 s, subexpression s (SIf x1 s x2) + | SIf_else_subexp : forall x1 x2 s, subexpression s (SIf x1 x2 s) + | SLoop_body1_subexp : forall x1 x2 s, subexpression s (SLoop s x1 x2) + | SLoop_body2_subexp : forall x1 x2 s, subexpression s (SLoop x1 x2 s) + | SSeq_body1_subexp : forall x1 s, subexpression s (SSeq s x1) + | SSeq_body2_subexp : forall x1 s, subexpression s (SSeq x1 s). + + Lemma wf_subexpression : well_founded subexpression. + Proof. + cbv [well_founded]. intros a. induction a. + all: constructor; intros ? H; inversion H; subst; assumption. + Defined. + + Definition stmt_lt := + clos_trans _ subexpression. + + Lemma wf_stmt_lt : well_founded stmt_lt. + Proof. + cbv [stmt_lt]. apply Transitive_Closure.wf_clos_trans. + apply wf_subexpression. + Defined. + Fixpoint modVars_as_list(veq: varname -> varname -> bool)(s: stmt): list varname := match s with | SSkip | SStore _ _ _ _ => [] @@ -206,7 +235,6 @@ Section FlatImp1. Context {width: Z} {BW: Bitwidth width} {word: word.word width}. Context {mem: map.map word byte} {locals: map.map varname word} {env: map.map String.string (list varname * list varname * stmt varname)}. - Context {ext_spec: ExtSpec}. Section WithEnv. Variable (e: env). @@ -289,7 +317,7 @@ Module exec. Local Notation metrics := MetricLog. (* COQBUG(unification finds Type instead of Prop and fails to downgrade *) - Implicit Types post : trace -> mem -> locals -> metrics -> Prop. + Implicit Types post : leakage -> trace -> mem -> locals -> metrics -> Prop. Definition lookup_op_locals (l: locals) (o: operand) := match o with @@ -322,117 +350,118 @@ Module exec. end mc. (* alternative semantics which allow non-determinism *) - Inductive exec: + Inductive exec {pick_sp: PickSp} : stmt varname -> - trace -> mem -> locals -> metrics -> - (trace -> mem -> locals -> metrics -> Prop) + leakage -> trace -> mem -> locals -> metrics -> + (leakage -> trace -> mem -> locals -> metrics -> Prop) -> Prop := - | interact: forall t m mKeep mGive l mc action argvars argvals resvars outcome post, + | interact: forall k t m mKeep mGive l mc action argvars argvals resvars outcome post, map.split m mKeep mGive -> map.getmany_of_list l argvars = Some argvals -> ext_spec t mGive action argvals outcome -> - (forall mReceive resvals, - outcome mReceive resvals -> + (forall mReceive resvals klist, + outcome mReceive resvals klist -> 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' + post (leak_list klist :: k) (((mGive, action, argvals), (mReceive, resvals)) :: t) m' l' (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, + exec (SInteract resvars action argvars) k t m l mc post + | call: forall k 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 mc outcome -> - (forall t' m' mc' st1, - outcome t' m' st1 mc' -> + exec fbody (leak_unit :: k) t m st0 mc outcome -> + (forall k' t' m' mc' st1, + outcome k' 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 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, + post k' t' m' l' (cost_call phase mc')) -> + exec (SCall binds fname args) k t m l mc post + | load: forall k 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_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, + post (leak_word (word.add addr (word.of_Z o)) :: k) t m (map.put l x v) (cost_load isReg x a mc)-> + exec (SLoad sz x a o) k t m l mc post + | store: forall k 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_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, + post (leak_word (word.add addr (word.of_Z o)) :: k) t m' l (cost_store isReg a v mc) -> + exec (SStore sz a v o) k t m l mc post + | inlinetable: forall sz x table i v index k 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_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, + post (leak_word index :: k) t m (map.put l x v) (cost_inlinetable isReg x i mc) -> + exec (SInlinetable sz x table i) k t m l mc post + | stackalloc: forall k t mSmall l mc x n body post, n mod (bytes_per_word width) = 0 -> - (forall a mStack mCombined, + (forall mStack mCombined, + let a := pick_sp k in anybytes a n mStack -> map.split mCombined mSmall mStack -> - exec body t mCombined (map.put l x a) mc - (fun t' mCombined' l' mc' => + exec body (leak_unit :: k) t mCombined (map.put l x a) mc + (fun k' t' mCombined' l' mc' => exists mSmall' mStack', anybytes a n mStack' /\ map.split mCombined' mSmall' mStack' /\ - 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_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, + post k' t' mSmall' l' (cost_stackalloc isReg x mc'))) -> + exec (SStackalloc x n body) k t mSmall l mc post + | lit: forall k t m l mc x v post, + post k t m (map.put l x (word.of_Z v)) (cost_lit isReg x mc) -> + exec (SLit x v) k t m l mc post + | op: forall k 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')) (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, + post (leak_binop op y' z' ++ k) t m (map.put l x (interp_binop op y' z')) (cost_SOp x y z mc) -> + exec (SOp x op y z) k t m l mc post + | set: forall k t m l mc x y y' post, map.get l y = Some y' -> - 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, + post k t m (map.put l x y') (cost_set isReg x y mc) -> + exec (SSet x y) k t m l mc post + | if_true: forall k t m l mc cond bThen bElse post, eval_bcond l cond = Some true -> - 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, + exec bThen (leak_bool true :: k) t m l (cost_SIf cond mc) post -> + exec (SIf cond bThen bElse) k t m l mc post + | if_false: forall k t m l mc cond bThen bElse post, eval_bcond l cond = Some false -> - 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, + exec bElse (leak_bool false :: k) t m l (cost_SIf cond mc) post -> + exec (SIf cond bThen bElse) k t m l mc post + | loop: forall k t m l mc cond body1 body2 mid1 mid2 post, (* This case is carefully crafted in such a way that recursive uses of exec only appear under forall and ->, but not under exists, /\, \/, to make sure the auto-generated induction principle contains an IH for all recursive uses. *) - exec body1 t m l mc mid1 -> - (forall t' m' l' mc', - mid1 t' m' l' mc' -> + exec body1 k t m l mc mid1 -> + (forall k' t' m' l' mc', + mid1 k' t' m' l' mc' -> eval_bcond l' cond <> None) -> - (forall t' m' l' mc', - mid1 t' m' l' mc' -> + (forall k' t' m' l' mc', + mid1 k' t' m' l' mc' -> eval_bcond l' cond = Some false -> - post t' m' l' (cost_SLoop_false cond mc')) -> - (forall t' m' l' mc', - mid1 t' m' l' mc' -> + post (leak_bool false :: k') t' m' l' (cost_SLoop_false cond mc')) -> + (forall k' t' m' l' mc', + mid1 k' t' m' l' mc' -> eval_bcond l' cond = Some true -> - exec body2 t' m' l' mc' mid2) -> - (forall t'' m'' l'' mc'', - mid2 t'' m'' l'' mc'' -> - exec (SLoop body1 cond body2) t'' m'' l'' + exec body2 (leak_bool true :: k') t' m' l' mc' mid2) -> + (forall k'' t'' m'' l'' mc'', + mid2 k'' t'' m'' l'' mc'' -> + exec (SLoop body1 cond body2) k'' t'' m'' l'' (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 -> - (forall t' m' l' mc', mid t' m' l' mc' -> exec s2 t' m' l' mc' post) -> - exec (SSeq s1 s2) t m l mc post - | skip: forall t m l mc post, - post t m l mc -> - exec SSkip t m l mc post. - - Lemma det_step: forall t0 m0 l0 mc0 s1 s2 t1 m1 l1 mc1 post, - exec s1 t0 m0 l0 mc0 (fun t1' m1' l1' mc1' => t1' = t1 /\ m1' = m1 /\ l1' = l1 /\ mc1 = mc1') -> - exec s2 t1 m1 l1 mc1 post -> - exec (SSeq s1 s2) t0 m0 l0 mc0 post. + exec (SLoop body1 cond body2) k t m l mc post + | seq: forall k t m l mc s1 s2 mid post, + exec s1 k t m l mc mid -> + (forall k' t' m' l' mc', mid k' t' m' l' mc' -> exec s2 k' t' m' l' mc' post) -> + exec (SSeq s1 s2) k t m l mc post + | skip: forall k t m l mc post, + post k t m l mc -> + exec SSkip k t m l mc post. + + Lemma det_step {pick_sp: PickSp} : forall k0 t0 m0 l0 mc0 s1 s2 k1 t1 m1 l1 mc1 post, + exec s1 k0 t0 m0 l0 mc0 (fun k1' t1' m1' l1' mc1' => k1' = k1 /\ t1' = t1 /\ m1' = m1 /\ l1' = l1 /\ mc1 = mc1') -> + exec s2 k1 t1 m1 l1 mc1 post -> + exec (SSeq s1 s2) k0 t0 m0 l0 mc0 post. Proof. intros. eapply seq; [eassumption|]. @@ -440,37 +469,37 @@ Module exec. assumption. Qed. - Lemma seq_cps: forall s1 s2 t m (l: locals) mc post, - exec s1 t m l mc (fun t' m' l' mc' => exec s2 t' m' l' mc' post) -> - exec (SSeq s1 s2) t m l mc post. + Lemma seq_cps {pick_sp: PickSp} : forall s1 s2 k t m (l: locals) mc post, + exec s1 k t m l mc (fun k' t' m' l' mc' => exec s2 k' t' m' l' mc' post) -> + exec (SSeq s1 s2) k t m l mc post. Proof. intros. eapply seq. 1: eassumption. simpl. clear. auto. Qed. - Lemma call_cps: forall fname params rets binds args fbody argvs t (l: locals) m mc st post, + Lemma call_cps {pick_sp: PickSp} : forall fname params rets binds args fbody argvs k t (l: locals) m mc st post, 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 mc - (fun t' m' st' mc' => + exec fbody (leak_unit :: k) t m st mc + (fun k' 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 phase mc')) -> - exec (SCall binds fname args) t m l mc post. + post k' t' m' l' (cost_call phase mc')) -> + exec (SCall binds fname args) k t m l mc post. Proof. intros. eapply call; try eassumption. cbv beta. intros *. exact id. Qed. - 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, + Lemma loop_cps {pick_sp: PickSp} : forall body1 cond body2 k t m l mc post, + exec body1 k t m l mc (fun k t m l mc => exists b, eval_bcond l cond = Some b /\ - (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 + (b = false -> post (leak_bool false :: k) t m l (cost_SLoop_false cond mc)) /\ + (b = true -> exec body2 (leak_bool true :: k) t m l mc (fun k t m l mc => + exec (SLoop body1 cond body2) k t m l (cost_SLoop_true cond mc) post))) -> - exec (SLoop body1 cond body2) t m l mc post. + exec (SLoop body1 cond body2) k t m l mc post. Proof. intros. eapply loop. 1: eapply H. all: cbv beta; intros; simp. - congruence. @@ -479,11 +508,11 @@ Module exec. - assumption. Qed. - Lemma weaken: forall t l m mc s post1, - exec s t m l mc post1 -> + Lemma weaken {pick_sp: PickSp} : forall s k t m l mc post1, + exec s k 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. + (forall k' t' m' l' mc', post1 k' t' m' l' mc' -> post2 k' t' m' l' mc') -> + exec s k t m l mc post2. Proof. induction 1; intros; try solve [econstructor; eauto]. - eapply interact; try eassumption. @@ -502,22 +531,22 @@ Module exec. intros. simp. eauto 10. Qed. - Lemma seq_assoc: forall s1 s2 s3 t m l mc post, - exec (SSeq s1 (SSeq s2 s3)) t m l mc post -> - exec (SSeq (SSeq s1 s2) s3) t m l mc post. + Lemma seq_assoc {pick_sp: PickSp} : forall s1 s2 s3 k t m l mc post, + exec (SSeq s1 (SSeq s2 s3)) k t m l mc post -> + exec (SSeq (SSeq s1 s2) s3) k t m l mc post. Proof. intros. simp. eapply seq_cps. eapply seq_cps. eapply weaken. 1: eassumption. intros. - specialize H8 with (1 := H). simp. + specialize H9 with (1 := H). simp. eapply weaken. 1: eassumption. intros. eauto. Qed. - Lemma seq_assoc_bw: forall s1 s2 s3 t m l mc post, - exec (SSeq (SSeq s1 s2) s3) t m l mc post -> - exec (SSeq s1 (SSeq s2 s3)) t m l mc post. + Lemma seq_assoc_bw {pick_sp: PickSp} : forall s1 s2 s3 k t m l mc post, + exec (SSeq (SSeq s1 s2) s3) k t m l mc post -> + exec (SSeq s1 (SSeq s2 s3)) k t m l mc post. Proof. intros. simp. eauto 10 using seq. Qed. Ltac equalities := @@ -528,42 +557,42 @@ Module exec. end; simp. - Lemma intersect: forall t l m mc s post1, - exec s t m l mc post1 -> + Lemma intersect {pick_sp: PickSp} : forall k t m l mc s post1, + exec s k 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'). + exec s k t m l mc post2 -> + exec s k t m l mc (fun k' t' m' l' mc' => post1 k' t' m' l' mc' /\ post2 k' t' m' l' mc'). Proof. induction 1; intros; match goal with - | H: exec _ _ _ _ _ _ |- _ => inversion H; subst; clear H + | H: exec _ _ _ _ _ _ _ |- _ => inversion H; subst; clear H end; equalities; try solve [econstructor; eauto | exfalso; congruence]. - (* SInteract *) pose proof ext_spec.mGive_unique as P. - specialize P with (1 := H) (2 := H7) (3 := H1) (4 := H14). + specialize P with (1 := H) (2 := H7) (3 := H1) (4 := H15). subst mGive0. destruct (map.split_diff (map.same_domain_refl mGive) H H7) as (? & _). subst mKeep0. eapply @interact. + eassumption. + eassumption. - + eapply ext_spec.intersect; [exact H1|exact H14]. + + eapply ext_spec.intersect; [exact H1|exact H15]. + simpl. intros. simp. edestruct H2 as (? & ? & ?); [eassumption|]. - edestruct H15 as (? & ? & ?); [eassumption|]. + edestruct H16 as (? & ? & ?); [eassumption|]. simp. equalities. eauto 10. - (* SCall *) rename IHexec into IH. - specialize IH with (1 := H16). + specialize IH with (1 := H17). eapply @call; [..|exact IH|]; eauto. rename H3 into Ex1. - rename H17 into Ex2. + rename H18 into Ex2. move Ex1 before Ex2. intros. simpl in *. simp. edestruct Ex1; [eassumption|]. @@ -575,7 +604,7 @@ Module exec. - (* SStackalloc *) eapply @stackalloc. 1: eassumption. intros. - rename H0 into Ex1, H12 into Ex2. + rename H0 into Ex1, H13 into Ex2. eapply weaken. 1: eapply H1. 1,2: eassumption. 1: eapply Ex2. 1,2: eassumption. cbv beta. @@ -593,9 +622,9 @@ Module exec. + simpl. intros. simp. eauto. + simpl. intros. simp. eauto. + simpl. intros. simp. eapply H3; [eassumption..|]. (* also an IH *) - eapply H18; eassumption. - + simpl. intros. simp. eapply H5; [eassumption..|]. (* also an IH *) eapply H19; eassumption. + + simpl. intros. simp. eapply H5; [eassumption..|]. (* also an IH *) + eapply H20; eassumption. - (* SSeq *) pose proof IHexec as IH1. @@ -606,6 +635,139 @@ Module exec. eauto. Qed. + Lemma exec_extends_trace {pick_sp: PickSp} s k t m l mc post : + exec s k t m l mc post -> + exec s k t m l mc (fun k' t' m' l' mc' => post k' t' m' l' mc' /\ exists k'', k' = k'' ++ k). + Proof. + intros H. induction H; try (econstructor; intuition eauto; eexists; align_trace; fail). + - econstructor; intuition eauto. specialize H2 with (1 := H3). fwd. + eexists. intuition eauto. eexists. align_trace. + - econstructor; intuition eauto. fwd. specialize H3 with (1 := H4p0). fwd. + eexists. intuition eauto. eexists. intuition eauto. + eexists. align_trace. + - econstructor; intuition eauto. intros. eapply weaken. 1: eapply H1; eauto. + simpl. intros. fwd. eexists. eexists. intuition eauto. eexists. align_trace. + - eapply if_true; intuition eauto. eapply weaken. 1: eapply IHexec. + simpl. intros. fwd. intuition eauto. eexists. align_trace. + - eapply if_false; intuition eauto. eapply weaken. 1: eapply IHexec. + simpl. intros. fwd. intuition eauto. eexists. align_trace. + - clear H2 H4. econstructor; intuition eauto; fwd; eauto. + { eexists. align_trace. } + { eapply weaken. 1: eapply H3; eauto. simpl. intros. fwd. + instantiate (1 := fun k'0 t'0 m'0 l'0 mc'0 => + mid2 k'0 t'0 m'0 l'0 mc'0 /\ exists k'', k'0 = k'' ++ k). + simpl. intuition. eexists. align_trace. } + simpl in *. fwd. eapply weaken. 1: eapply H5; eauto. + simpl. intros. fwd. intuition. eexists. align_trace. + - econstructor; intuition eauto. fwd. eapply weaken. 1: eapply H1; eauto. + simpl. intros. fwd. intuition eauto. eexists. align_trace. + Qed. + + Lemma exec_ext (pick_sp1: PickSp) s k t m l mc post : + exec (pick_sp := pick_sp1) s k t m l mc post -> + forall pick_sp2, + (forall k', pick_sp1 (k' ++ k) = pick_sp2 (k' ++ k)) -> + exec (pick_sp := pick_sp2) s k t m l mc post. + Proof. + intros H1 pick_sp2. induction H1; intros; try solve [econstructor; eauto]. + - econstructor. 4: eapply exec_extends_trace. all: intuition eauto. + { eapply IHexec. intros. rewrite associate_one_left. + repeat rewrite app_assoc. auto. } + fwd. specialize H3 with (1 := H5p0). fwd. intuition eauto. + - econstructor; eauto. intros. replace (pick_sp1 k) with (pick_sp2 k) in *. + { subst a. eapply weaken. + { eapply H1; eauto. + intros. eassert (H2' := H2 (_ ++ _ :: nil)). rewrite <- app_assoc in H2'. eapply H2'. } + eauto. } + symmetry. apply H2 with (k' := nil). + - eapply if_true; eauto. eapply IHexec. + intros. rewrite associate_one_left. repeat rewrite app_assoc. auto. + - eapply if_false; intuition eauto. eapply IHexec. + intros. rewrite associate_one_left. repeat rewrite app_assoc. auto. + - clear H2 H4. eapply loop. 1: eapply exec_extends_trace. all: intuition eauto; fwd; eauto. + { eapply weaken. 1: eapply exec_extends_trace. + { eapply H3; eauto. + intros. rewrite associate_one_left. repeat rewrite app_assoc. auto. } + simpl. intros. fwd. + instantiate (1 := fun k'0 t'0 m'0 l'0 mc'0 => + mid2 k'0 t'0 m'0 l'0 mc'0 /\ exists k'', k'0 = k'' ++ k). + simpl. intuition eauto. eexists. align_trace. } + simpl in *. fwd. eapply H5; eauto. intros. + repeat (rewrite app_assoc || rewrite (app_one_l _ (_ ++ k))). auto. + - econstructor. 1: eapply exec_extends_trace; eauto. simpl. intros. fwd. + eapply H0; eauto. intros. repeat rewrite app_assoc. apply H2. + Qed. + + Local Ltac solve_picksps_equal := + intros; cbv beta; f_equal; + repeat (rewrite rev_app_distr || cbn [rev app]); rewrite List.skipn_app_r; + [|repeat (rewrite app_length || rewrite rev_length || simpl); blia]; + repeat rewrite <- app_assoc; rewrite List.skipn_app_r; + [|rewrite rev_length; reflexivity]; + repeat (rewrite rev_app_distr || cbn [rev app] || rewrite rev_involutive); + repeat rewrite <- app_assoc; reflexivity. + + Lemma exec_to_other_trace (pick_sp: PickSp) s k1 k2 t m l mc post : + exec s k1 t m l mc post -> + exec (pick_sp := fun k => pick_sp (rev (skipn (length k2) (rev k)) ++ k1)) + s k2 t m l mc (fun k2' t' m' l' mc' => + exists k'', + k2' = k'' ++ k2 /\ + post (k'' ++ k1) t' m' l' mc'). + Proof. + intros H. generalize dependent k2. induction H; intros. + - econstructor; intuition eauto. apply H2 in H3. fwd. + eexists. intuition eauto. eexists. intuition eauto. 1: align_trace. + auto. + - econstructor; intuition eauto. + { eapply exec_ext with (pick_sp1 := _). 1: eapply IHexec; eauto. solve_picksps_equal. } + cbv beta in *. fwd. apply H3 in H4p1. + fwd. eexists. intuition eauto. eexists. intuition eauto. eexists. + split; [align_trace|]. repeat rewrite <- app_assoc. auto. + - econstructor; intuition eauto. eexists. split; [align_trace|]. auto. + - econstructor; intuition eauto. eexists. split; [align_trace|]. auto. + - econstructor; intuition eauto. eexists. split; [align_trace|]. auto. + - econstructor; intuition eauto. intros. + replace (rev k2) with (rev k2 ++ nil) in * by apply app_nil_r. + rewrite List.skipn_app_r in * by (rewrite rev_length; reflexivity). + simpl in *. eapply weaken. + { eapply exec_ext with (pick_sp1 := _). 1: eapply H1; eauto. solve_picksps_equal. } + simpl. intros. fwd. eexists _, _. intuition eauto. eexists (_ ++ _ :: nil). + rewrite <- app_assoc. simpl. rewrite <- (app_assoc _ _ k). simpl. eauto. + - econstructor; intuition eauto. eexists. split; [align_trace|]. auto. + - econstructor; intuition eauto. + - econstructor; intuition eauto. eexists. split; [align_trace|]. auto. + - eapply if_true; intuition eauto. eapply weaken. + { eapply exec_ext with (pick_sp1 := _). 1: eapply IHexec. solve_picksps_equal. } + simpl. intros. fwd. eexists. split; [align_trace|]. + repeat rewrite <- app_assoc. auto. + - eapply if_false; intuition eauto. eapply weaken. + { eapply exec_ext with (pick_sp1 := _). 1: eapply IHexec. solve_picksps_equal. } + simpl. intros. fwd. eexists. split; [align_trace|]. + repeat rewrite <- app_assoc. auto. + - eapply loop. 1: eapply IHexec. all: intuition eauto; fwd; eauto. + { eexists. split; [align_trace|]. simpl. auto. } + { eapply exec_ext with (pick_sp1 := _). + { eapply weaken. 1: eapply H3; eauto. simpl. intros. + instantiate (1 := fun k'0 t'0 m'0 l'0 mc'0 => + exists k''0, + k'0 = k''0 ++ k2 /\ + mid2 (k''0 ++ k) t'0 m'0 l'0 mc'0). + fwd. eexists. split; [align_trace|]. + repeat rewrite <- app_assoc. simpl. auto. } + solve_picksps_equal. } + simpl in *. fwd. eapply exec_ext with (pick_sp1 := _). + { eapply weaken. 1: eapply H5; eauto. simpl. intros. fwd. + eexists. split; [align_trace|]. + repeat rewrite <- app_assoc. auto. } + solve_picksps_equal. + - econstructor; intuition. fwd. eapply weaken. + { eapply exec_ext with (pick_sp1 := _). 1: eapply H1; eauto. solve_picksps_equal. } + simpl. intros. fwd. eexists. split; [align_trace|]. + repeat rewrite <- app_assoc. auto. + - econstructor. eexists. split; [align_trace|]. assumption. + Qed. + End FlatImpExec. End exec. Notation exec := exec.exec. @@ -616,7 +778,7 @@ Section FlatImp2. Context {width: Z} {BW: Bitwidth width} {word: word.word width}. Context {mem: map.map word byte} {locals: map.map varname word} {env: map.map String.string (list varname * list varname * stmt varname)}. - Context {ext_spec: ExtSpec}. + Context {ext_spec: ExtSpec} {pick_sp: PickSp}. Context {varname_eq_spec: EqDecider varname_eqb} {word_ok: word.ok word} {mem_ok: map.ok mem} @@ -627,15 +789,15 @@ Section FlatImp2. Variable (phase: compphase). Variable (isReg: varname -> bool). - Definition SimState: Type := trace * mem * locals * MetricLog. + Definition SimState: Type := leakage * trace * mem * locals * MetricLog. Definition SimExec(e: env)(c: stmt varname): SimState -> (SimState -> Prop) -> Prop := - fun '(t, m, l, mc) post => - exec phase isReg e c t m l mc (fun t' m' l' mc' => post (t', m', l', mc')). + fun '(k, t, m, l, mc) post => + exec phase isReg e c k t m l mc (fun k' t' m' l' mc' => post (k', t', m', l', mc')). - Lemma modVarsSound: forall e s initialT (initialSt: locals) initialM (initialMc: MetricLog) post, - 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). + Lemma modVarsSound: forall e s initialK initialT (initialSt: locals) initialM (initialMc: MetricLog) post, + exec phase isReg e s initialK initialT initialM initialSt initialMc post -> + exec phase isReg e s initialK initialT initialM initialSt initialMc + (fun initialK finalT finalM finalSt _ => map.only_differ initialSt (modVars s) finalSt). Proof. induction 1; try solve [ econstructor; [eassumption..|simpl; map_solver locals_ok] ]. @@ -666,9 +828,9 @@ Section FlatImp2. eapply exec.weaken; [eassumption|]. simpl; intros. map_solver locals_ok. - eapply @exec.loop with - (mid1 := fun t' m' l' mc' => mid1 t' m' l' mc' /\ + (mid1 := fun k' t' m' l' mc' => mid1 k' t' m' l' mc' /\ map.only_differ l (modVars body1) l') - (mid2 := fun t' m' l' mc' => mid2 t' m' l' mc' /\ + (mid2 := fun k' t' m' l' mc' => mid2 k' t' m' l' mc' /\ map.only_differ l (modVars (SLoop body1 cond body2)) l'). + eapply exec.intersect; eassumption. + intros. simp. eauto. @@ -683,7 +845,7 @@ Section FlatImp2. * eapply H5; eassumption. * simpl. intros. map_solver locals_ok. - eapply @exec.seq with - (mid := fun t' m' l' mc' => mid t' m' l' mc' /\ map.only_differ l (modVars s1) l'). + (mid := fun k' t' m' l' mc' => mid k' t' m' l' mc' /\ map.only_differ l (modVars s1) l'). + eapply exec.intersect; eassumption. + simpl; intros. simp. eapply exec.weaken; [eapply H1; eauto|]. diff --git a/compiler/src/compiler/FlatToRiscvCommon.v b/compiler/src/compiler/FlatToRiscvCommon.v index 7cdb7c49b..70676fd2d 100644 --- a/compiler/src/compiler/FlatToRiscvCommon.v +++ b/compiler/src/compiler/FlatToRiscvCommon.v @@ -1,5 +1,6 @@ Require Import riscv.Utility.Monads. Require Import riscv.Utility.MonadNotations. Require Import coqutil.Macros.unique. +Require Import bedrock2.LeakageSemantics. Require Import compiler.FlatImp. Require Import Coq.Lists.List. Import ListNotations. @@ -46,6 +47,7 @@ Require Import compiler.DivisibleBy4. Require Import compiler.MetricsToRiscv. Require Export compiler.regs_initialized. Require Import bedrock2.MetricCosts. +Require Import riscv.Spec.LeakageOfInstr. Require Import coqutil.Word.Interface. Local Hint Mode Word.Interface.word - : typeclass_instances. @@ -79,14 +81,15 @@ Section WithParameters. Context {pos_map: map.map String.string Z}. Context (compile_ext_call: pos_map -> Z -> Z -> stmt Z -> list Instruction). + Context (leak_ext_call: word -> pos_map -> Z -> Z -> stmt Z -> list word -> list LeakageEvent). Context {word_ok: word.ok word}. Context {mem: map.map word byte}. Context {env: map.map String.string (list Z * list Z * stmt Z)}. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. - Context {ext_spec: Semantics.ExtSpec}. + Context {ext_spec: LeakageSemantics.ExtSpec}. Definition runsTo{BWM: bitwidth_iset width iset}: (* BWM is unused, but makes iset inferrable *) MetricRiscvMachine -> (MetricRiscvMachine -> Prop) -> Prop := @@ -289,7 +292,7 @@ Section WithParameters. exists pos, map.get finfo f = Some pos /\ pos mod 4 = 0. Local Notation stmt := (stmt Z). - Local Notation exec := (exec PostSpill isRegZ). + Local Notation exec pick_sp e := (@exec _ _ _ _ _ _ _ _ PostSpill isRegZ e pick_sp). (* 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 @@ -298,9 +301,9 @@ Section WithParameters. Definition compiles_FlatToRiscv_correctly{BWM: bitwidth_iset width iset} (f: pos_map -> Z -> Z -> stmt -> list Instruction) (s: stmt): Prop := - forall e_impl_full initialTrace initialMH initialRegsH initialMetricsH postH, - exec e_impl_full s initialTrace (initialMH: mem) initialRegsH initialMetricsH postH -> - forall g e_impl e_pos program_base insts xframe (initialL: MetricRiscvMachine) pos, + forall e_impl_full pick_sp1 initialK initialTrace initialMH initialRegsH initialMetricsH postH, + exec pick_sp1 e_impl_full s initialK initialTrace (initialMH: mem) initialRegsH initialMetricsH postH -> + forall g e_impl e_pos program_base insts xframe (initialL: MetricRiscvMachine) pos initialKL cont, map.extends e_impl_full e_impl -> good_e_impl e_impl e_pos -> fits_stack g.(rem_framewords) g.(rem_stackwords) e_impl s -> @@ -310,20 +313,30 @@ Section WithParameters. pos mod 4 = 0 -> word.unsigned program_base mod 4 = 0 -> initialL.(getPc) = word.add program_base (word.of_Z pos) -> + initialL.(getTrace) = Some initialKL -> iff1 g.(allx) (xframe * program iset (word.add program_base (word.of_Z pos)) insts * functions program_base e_pos e_impl)%sep -> goodMachine initialTrace initialMH initialRegsH g initialL -> - runsTo initialL (fun finalL => exists finalTrace finalMH finalRegsH finalMetricsH, - postH finalTrace finalMH finalRegsH finalMetricsH /\ + (forall k, pick_sp1 (k ++ initialK) = snd (stmt_leakage iset compile_ext_call leak_ext_call e_pos e_impl_full program_base + (s, rev k, rev initialKL, pos, g.(p_sp), bytes_per_word * rem_framewords g, cont k))) -> + runsTo initialL (fun finalL => exists finalK finalTrace finalMH finalRegsH finalMetricsH, + postH finalK finalTrace finalMH finalRegsH finalMetricsH /\ finalL.(getPc) = word.add initialL.(getPc) (word.of_Z (4 * Z.of_nat (List.length insts))) /\ map.only_differ initialL.(getRegs) (union (of_list (modVars_as_list Z.eqb s)) (singleton_set RegisterNames.ra)) finalL.(getRegs) /\ (finalL.(getMetrics) - initialL.(getMetrics) <= - lowerMetrics (finalMetricsH - initialMetricsH))%metricsL /\ - goodMachine finalTrace finalMH finalRegsH g finalL). + lowerMetrics (finalMetricsH - initialMetricsH))%metricsL /\ + (exists kH'' finalKL, + finalK = kH'' ++ initialK /\ + finalL.(getTrace) = Some finalKL /\ + forall k cont, + stmt_leakage iset compile_ext_call leak_ext_call e_pos e_impl_full program_base + (s, rev kH'' ++ k, rev initialKL, pos, g.(p_sp), bytes_per_word * rem_framewords g, cont) = + cont (rev kH'') (rev finalKL)) /\ + goodMachine finalTrace finalMH finalRegsH g finalL). End WithParameters. @@ -351,15 +364,16 @@ Ltac simpl_bools := Section FlatToRiscv1. Context {iset: InstructionSet}. Context {fun_info: map.map String.string Z}. - Context (compile_ext_call: fun_info -> Z -> Z -> stmt Z -> list Instruction). Context {width: Z} {BW: Bitwidth width} {word: word.word width}. + Context (compile_ext_call: fun_info -> Z -> Z -> stmt Z -> list Instruction). + Context (leak_ext_call: fun_info -> Z -> Z -> stmt Z -> list word -> list LeakageEvent). Context {word_ok: word.ok word}. Context {locals: map.map Z word}. Context {mem: map.map word byte}. Context {env: map.map String.string (list Z * list Z * stmt Z)}. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {ext_spec: Semantics.ExtSpec}. Context {word_riscv_ok: word.riscv_ok word}. @@ -451,7 +465,25 @@ Section FlatToRiscv1. Memory.load, Memory.load_Z in *; simp; simulate''; simpl; simpl_word_exprs word_ok; destruct initialL; try eassumption]. - Qed. + Qed. + + Lemma go_leak_load: forall sz (x a ofs: Z) (addr: word) (initialL: RiscvMachineL) post (f : option LeakageEvent -> M unit), + valid_register a -> + map.get initialL.(getRegs) a = Some addr -> + mcomp_sat (f (Some (executeInstr (compile_load iset sz x a ofs) (leak_load iset sz addr)))) initialL post -> + mcomp_sat (Bind (leakage_of_instr Machine.getRegister (compile_load iset sz x a ofs)) f) initialL post. + Proof. + unfold leakage_of_instr, leakage_of_instr_I, leak_load, compile_load, Memory.bytes_per, Memory.bytes_per_word. + rewrite bitwidth_matches. + destruct width_cases as [E | E]; + (* note: "rewrite E" does not work because "width" also appears in the type of "word", + but we don't need to rewrite in the type of word, only in the type of the tuple, + which works if we do it before intro'ing it *) + (destruct (width =? 32)); + intros; destruct sz; + simulate''; + eassumption. + Qed. Arguments invalidateWrittenXAddrs: simpl never. @@ -485,6 +517,24 @@ Section FlatToRiscv1. simp; simulate''; simpl; simpl_word_exprs word_ok; try eassumption. Qed. + Lemma go_leak_store: forall sz (x a ofs: Z) (addr: word) (initialL: RiscvMachineL) post f, + valid_register a -> + map.get initialL.(getRegs) a = Some addr -> + mcomp_sat (f (Some (executeInstr (compile_store iset sz a x ofs) (leak_store iset sz addr)))) initialL post -> + mcomp_sat (Bind (leakage_of_instr Machine.getRegister (compile_store iset sz a x ofs)) f) initialL post. + Proof. + unfold leakage_of_instr, leakage_of_instr_I, leak_store, compile_store, Memory.bytes_per, Memory.bytes_per_word. + rewrite bitwidth_matches. + destruct width_cases as [E | E]; + (* note: "rewrite E" does not work because "width" also appears in the type of "word", + but we don't need to rewrite in the type of word, only in the type of the tuple, + which works if we do it before intro'ing it *) + (destruct (width =? 32)); + intros; destruct sz; + simulate''; + eassumption. + Qed. + Lemma run_compile_load: forall sz: Syntax.access_size, run_Load_spec iset (@Memory.bytes_per width sz) (compile_load iset sz) id. Proof using word_ok mem_ok PR BWM. @@ -537,6 +587,7 @@ Section FlatToRiscv1. getPc finalL = getNextPc initialL /\ getNextPc finalL = word.add (getPc finalL) (word.of_Z 4) /\ getMetrics finalL = addMetricInstructions 1 (addMetricLoads 2 (getMetrics initialL)) /\ + getTrace finalL = option_map (cons (executeInstr (compile_load iset Syntax.access_size.word rd rs ofs) (leak_load iset Syntax.access_size.word base))) (option_map (cons (fetchInstr initialL.(getPc))) initialL.(getTrace)) /\ valid_machine finalL). Proof using word_ok mem_ok PR BWM. intros. @@ -545,12 +596,15 @@ Section FlatToRiscv1. - eapply (run_compile_load Syntax.access_size.word); cycle -3; try eassumption. instantiate (2:=ltac:(destruct pf)); destruct pf; eassumption. - cbv beta. intros. simp. repeat split; try assumption. - etransitivity. 1: eassumption. - unfold id. - rewrite LittleEndian.combine_of_list, LittleEndianList.le_combine_split. - replace (BinInt.Z.of_nat (Memory.bytes_per Syntax.access_size.word) * 8) with width. - + rewrite word.wrap_unsigned. rewrite word.of_Z_unsigned. reflexivity. - + clear -BW. destruct width_cases as [E | E]; rewrite E; reflexivity. + + etransitivity. 1: eassumption. + unfold id. + rewrite LittleEndian.combine_of_list, LittleEndianList.le_combine_split. + replace (BinInt.Z.of_nat (Memory.bytes_per Syntax.access_size.word) * 8) with width. + * rewrite word.wrap_unsigned. rewrite word.of_Z_unsigned. reflexivity. + * clear -BW. destruct width_cases as [E | E]; rewrite E; reflexivity. + + rewrite H8p7. cbv [final_trace concrete_leakage_of_instr compile_load leak_load]. + destruct (getTrace initialL); [|repeat Tactics.destruct_one_match || reflexivity]. + destruct (bitwidth iset =? 32); simpl; rewrite Z.eqb_refl; reflexivity. - eapply LittleEndianList.length_le_split. Qed. @@ -578,6 +632,7 @@ Section FlatToRiscv1. getPc finalL = getNextPc initialL /\ getNextPc finalL = word.add (getPc finalL) (word.of_Z 4) /\ getMetrics finalL = addMetricInstructions 1 (addMetricStores 1 (addMetricLoads 1 (getMetrics initialL))) /\ + getTrace finalL = option_map (cons (executeInstr (compile_store iset Syntax.access_size.word rs1 rs2 ofs) (leak_store iset Syntax.access_size.word base))) (option_map (cons (fetchInstr initialL.(getPc))) initialL.(getTrace)) /\ valid_machine finalL). Proof using word_ok mem_ok PR BWM. intros. @@ -586,10 +641,14 @@ Section FlatToRiscv1. - eapply (run_compile_store Syntax.access_size.word); cycle -3; try eassumption. instantiate (2:=ltac:(destruct pf)); destruct pf; eassumption. - cbv beta. intros. simp. repeat split; try assumption. - unfold scalar, truncated_word, truncated_scalar, littleendian, ptsto_bytes in *. - rewrite HList.tuple.to_list_of_list. - rewrite LittleEndian.to_list_split in *. - eassumption. + + unfold scalar, truncated_word, truncated_scalar, littleendian, ptsto_bytes in *. + rewrite HList.tuple.to_list_of_list. + rewrite LittleEndian.to_list_split in *. + eassumption. + + rewrite H9p7. + cbv [final_trace concrete_leakage_of_instr compile_store leak_store]. + destruct (getTrace initialL); [|repeat Tactics.destruct_one_match || reflexivity]. + destruct (bitwidth iset =? 32); simpl; rewrite Z.eqb_refl; reflexivity. - eapply LittleEndianList.length_le_split. Qed. @@ -921,8 +980,12 @@ Ltac simulate'_step := (* lemmas introduced only in this file: *) | |- mcomp_sat (Monads.Bind (Execute.execute (compile_load _ _ _ _ _)) _) _ _ => eapply go_load; [ sidecondition.. | idtac ] + | |- mcomp_sat (Monads.Bind (leakage_of_instr _ (compile_load _ _ _ _ _)) _) _ _ => + eapply go_leak_load; [ sidecondition.. | idtac ] | |- mcomp_sat (Monads.Bind (Execute.execute (compile_store _ _ _ _ _)) _) _ _ => eapply go_store; [ sidecondition.. | idtac ] + | |- mcomp_sat (Monads.Bind (leakage_of_instr _ (compile_store _ _ _ _ _)) _) _ _ => + eapply go_leak_store; [ sidecondition.. | idtac ] (* simulate_step from GoFlatToRiscv: *) | |- _ => simulate_step | |- _ => simpl_modu4_0 diff --git a/compiler/src/compiler/FlatToRiscvDef.v b/compiler/src/compiler/FlatToRiscvDef.v index 649c6d2a0..a519bcfd7 100644 --- a/compiler/src/compiler/FlatToRiscvDef.v +++ b/compiler/src/compiler/FlatToRiscvDef.v @@ -1,6 +1,7 @@ Require Import coqutil.Macros.unique. Require Import coqutil.Decidable. Require Import compiler.FlatImp. +Require Import Coq.Wellfounded.Wellfounded. Require Import Coq.Lists.List. Import ListNotations. Require Import Coq.ZArith.ZArith. @@ -14,9 +15,12 @@ Require Import coqutil.Datatypes.ListSet. Require Import riscv.Utility.Encode. Require Import riscv.Utility.RegisterNames. Require Import bedrock2.Syntax. +Require Import bedrock2.LeakageSemantics. +Require Import compiler.FixEq. Require Import coqutil.Map.Interface. Require Import compiler.SeparationLogic. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import compiler.Registers. Require Import compiler.FlatImpConstraints. @@ -68,6 +72,7 @@ Section FlatToRiscv1. Context (iset: InstructionSet). + Context {width} {BW : Bitwidth width} {word: word.word width}. (* Part 2: compilation *) @@ -87,6 +92,21 @@ Section FlatToRiscv1. | access_size.word => if bitwidth iset =? 32 then Lw else Ld end. + Definition leak_Lbu x := ILeakage (Lbu_leakage x). + Definition leak_Lhu x := ILeakage (Lhu_leakage x). + Definition leak_Lw x := ILeakage (Lw_leakage x). + Definition leak_Lwu x := I64Leakage (Lwu_leakage x). + Definition leak_Ld x := I64Leakage (Ld_leakage x). + + Definition leak_load(sz: access_size) : + word -> InstructionLeakage := + match sz with + | access_size.one => leak_Lbu + | access_size.two => leak_Lhu + | access_size.four => if bitwidth iset =? 32 then leak_Lw else leak_Lwu + | access_size.word => if bitwidth iset =? 32 then leak_Lw else leak_Ld + end. + Definition compile_store(sz: access_size): Z -> Z -> Z -> Instruction := match sz with @@ -96,6 +116,20 @@ Section FlatToRiscv1. | access_size.word => if bitwidth iset =? 32 then Sw else Sd end. + Definition leak_Sb x := ILeakage (Sb_leakage x). + Definition leak_Sh x := ILeakage (Sh_leakage x). + Definition leak_Sw x := ILeakage (Sw_leakage x). + Definition leak_Sd x := I64Leakage (Sd_leakage x). + + Definition leak_store(sz: access_size) : + word -> InstructionLeakage := + match sz with + | access_size.one => leak_Sb + | access_size.two => leak_Sh + | access_size.four => leak_Sw + | access_size.word => if bitwidth iset =? 32 then leak_Sw else leak_Sd + end. + Definition compile_op_imm(rd: Z)(op: Syntax.bopname)(rs1: Z)(c2: Z): list Instruction := match op with | Syntax.bopname.add => [[Addi rd rs1 c2]] @@ -110,6 +144,30 @@ Section FlatToRiscv1. | _ => [InvalidInstruction (-1)] end. + Definition leak_Addi := ILeakage Addi_leakage. + Definition leak_Andi := ILeakage Andi_leakage. + Definition leak_Ori := ILeakage Ori_leakage. + Definition leak_Xori := ILeakage Xori_leakage. + Definition leak_Srli := ILeakage Srli_leakage. + Definition leak_Slli := ILeakage Slli_leakage. + Definition leak_Srai := ILeakage Srai_leakage. + Definition leak_Slti := ILeakage Slti_leakage. + Definition leak_Sltiu := ILeakage Sltiu_leakage. + + Definition leak_op_imm(op: Syntax.bopname) : list InstructionLeakage := + match op with + | Syntax.bopname.add => [leak_Addi] + | Syntax.bopname.and => [leak_Andi] + | Syntax.bopname.or => [leak_Ori] + | Syntax.bopname.xor => [leak_Xori] + | Syntax.bopname.sru => [leak_Srli] + | Syntax.bopname.slu => [leak_Slli] + | Syntax.bopname.srs => [leak_Srai] + | Syntax.bopname.lts => [leak_Slti] + | Syntax.bopname.ltu => [leak_Sltiu] + | _ => [] + end. + Definition compile_op_register(rd: Z)(op: Syntax.bopname)(rs1 rs2: Z): list Instruction := match op with | Syntax.bopname.add => [[Add rd rs1 rs2]] @@ -128,15 +186,61 @@ Section FlatToRiscv1. | Syntax.bopname.ltu => [[Sltu rd rs1 rs2]] | Syntax.bopname.eq => [[Sub rd rs1 rs2; Seqz rd rd]] end. + + Definition leak_Add := ILeakage Add_leakage. + Definition leak_Sub := ILeakage Sub_leakage. + Definition leak_Mul := MLeakage Mul_leakage. + Definition leak_Mulhu := MLeakage Mulhu_leakage. + Definition leak_Divu (num den : word) := MLeakage (Divu_leakage num den). + Definition leak_Remu (num den : word) := MLeakage (Remu_leakage num den). + Definition leak_And := ILeakage And_leakage. + Definition leak_Or := ILeakage Or_leakage. + Definition leak_Xor := ILeakage Xor_leakage. + Definition leak_Srl := ILeakage Srl_leakage. + Definition leak_Sll := ILeakage Sll_leakage. + Definition leak_Sra := ILeakage Sra_leakage. + Definition leak_Slt := ILeakage Slt_leakage. + Definition leak_Sltu := ILeakage Sltu_leakage. + Definition leak_Seqz := ILeakage Sltiu_leakage. + + Definition leak_op_register (op: Syntax.bopname) (x1 x2 : word) : list InstructionLeakage := + match op with + | Syntax.bopname.add => [ leak_Add ] + | Syntax.bopname.sub => [ leak_Sub ] + | Syntax.bopname.mul => [ leak_Mul ] + | Syntax.bopname.mulhuu => [ leak_Mulhu ] + | Syntax.bopname.divu => [ leak_Divu x1 x2 ] + | Syntax.bopname.remu => [ leak_Remu x1 x2 ] + | Syntax.bopname.and => [ leak_And ] + | Syntax.bopname.or => [ leak_Or ] + | Syntax.bopname.xor => [ leak_Xor ] + | Syntax.bopname.sru => [ leak_Srl ] + | Syntax.bopname.slu => [ leak_Sll ] + | Syntax.bopname.srs => [ leak_Sra ] + | Syntax.bopname.lts => [ leak_Slt ] + | Syntax.bopname.ltu => [ leak_Sltu ] + | Syntax.bopname.eq => [ leak_Sub ; leak_Seqz ] + end. + Definition compile_op(rd: Z)(op: Syntax.bopname)(op1 : Z)(op2: operand): list Instruction := match op2 with | Var v2 => compile_op_register rd op op1 v2 | Const c2 => compile_op_imm rd op op1 c2 end. + Definition leak_op (op : Syntax.bopname) (op2: @operand Z) (x1 x2 : word) : + list InstructionLeakage := + match op2 with + | Var _ => leak_op_register op x1 x2 + | Const c2 => leak_op_imm op + end. + Definition compile_lit_12bit(rd: Z)(v: Z): list Instruction := [[ Addi rd Register0 (signExtend 12 v) ]]. + Definition leak_lit_12bit : list InstructionLeakage := + [ leak_Addi ]. + (* On a 64bit machine, loading a constant -2^31 <= v < 2^31 is not always possible with a Lui followed by an Addi: If the constant is of the form 0x7ffffXXX, and XXX has its highest bit set, we would @@ -160,6 +264,11 @@ Section FlatToRiscv1. let hi := Z.lxor (signExtend 32 v) lo in [[ Lui rd hi ; Xori rd rd lo ]]. + Definition leak_Lui := ILeakage Lui_leakage. + + Definition leak_lit_32bit : list InstructionLeakage := + [ leak_Lui ; leak_Xori ]. + Definition compile_lit_64bit(rd: Z)(v: Z): list Instruction := let v0 := bitSlice v 0 11 in let v1 := bitSlice v 11 22 in @@ -173,6 +282,15 @@ Section FlatToRiscv1. Slli rd rd 11 ; Xori rd rd v0 ]]. + Definition leak_lit_64bit : list InstructionLeakage := + leak_lit_32bit ++ + [ leak_Slli ; + leak_Xori ; + leak_Slli ; + leak_Xori ; + leak_Slli ; + leak_Xori ]. + Definition compile_lit(rd: Z)(v: Z): list Instruction := if ((-2^11 <=? v)%Z && (v InstructionLeakage := + match cond with + | CondBinary op _ _ => + match op with + | BEq => leak_Bne + | BNe => leak_Beq + | BLt => leak_Bge + | BGe => leak_Blt + | BLtu => leak_Bgeu + | BGeu => leak_Bltu + end + | CondNez _ => + leak_Beq + end. + Local Notation bytes_per_word := (Memory.bytes_per_word (bitwidth iset)). Fixpoint save_regs(regs: list Z)(offset: Z): list Instruction := @@ -206,6 +354,13 @@ Section FlatToRiscv1. :: (save_regs regs (offset + bytes_per_word)) end. + Fixpoint leak_save_regs + (sp_val: word)(regs: list Z): list InstructionLeakage := + match regs with + | nil => nil + | r :: regs' => leak_store access_size.word sp_val :: leak_save_regs sp_val regs' + end. + Fixpoint load_regs(regs: list Z)(offset: Z): list Instruction := match regs with | nil => nil @@ -213,6 +368,13 @@ Section FlatToRiscv1. :: (load_regs regs (offset + bytes_per_word)) end. + Fixpoint leak_load_regs + (sp_val: word)(regs: list Z): list InstructionLeakage := + match regs with + | nil => nil + | r :: regs' => leak_load access_size.word sp_val :: leak_load_regs sp_val regs' + end. + (* number of words of stack allocation space needed within current frame *) Fixpoint stackalloc_words(s: stmt Z): Z := match s with @@ -242,12 +404,16 @@ Section FlatToRiscv1. (* All positions are relative to the beginning of the progam, so we get completely position independent code. *) + Context {mem: map.map word byte}. Context {env: map.map String.string (list Z * list Z * stmt Z)}. Context {pos_map: map.map String.string Z}. Context (compile_ext_call: pos_map -> Z -> Z -> stmt Z -> list Instruction). + Context (leak_ext_call: word(*program_base*) -> pos_map -> Z (*sp_val*) -> Z (*stackoffset*) -> stmt Z -> list word (*source-level leakage*) -> list LeakageEvent). Section WithEnv. Variable e: pos_map. + Variable e_env: env. + Variable program_base: word. (* mypos: position of the code relative to the positions in e stackoffset: $sp + stackoffset is the (last) highest used stack address (for SStackalloc) @@ -332,7 +498,7 @@ Section FlatToRiscv1. [[ Jalr zero ra 0 ]]. Definition add_compiled_function(state: list Instruction * pos_map)(fname: String.string) - (fimpl: list Z * list Z * stmt Z): list Instruction * pos_map := + (fimpl: list Z * list Z * stmt Z): list Instruction * pos_map := let '(old_insts, infomap) := state in let pos := 4 * Z.of_nat (length (old_insts)) in let new_insts := compile_function pos fimpl in @@ -342,6 +508,288 @@ Section FlatToRiscv1. Definition compile_funs: env -> list Instruction * pos_map := map.fold add_compiled_function (nil, map.empty). + + Definition leak_Jal := ILeakage Jal_leakage. + Definition leak_Jalr x := ILeakage (Jalr_leakage x). + + Definition tuple : Type := stmt Z * leakage * list LeakageEvent * Z * word * Z * (leakage -> list LeakageEvent -> list LeakageEvent * word). + + (*having to add in this extra information is inconvenient. + it'd be nice to prove that it's all redundant, and hence can be reconstructed from just the 'leakages' list. + This would probably be feasible if we could prove that only the program is ever executed, and the program is + never modified. But proving these things would probably be more effort than just including the redundant information.*) + Fixpoint leakage_events (abs_pos: word) (instrs : list Instruction) (leakages : list InstructionLeakage) : list LeakageEvent := + match instrs, leakages with + | instr :: instrs, leakage :: leakages => + fetchInstr abs_pos :: + executeInstr instr leakage :: + leakage_events (word.add abs_pos (word.of_Z 4)) instrs leakages + | _, _ => nil + end. + + Definition leakage_events_rel (mypos: Z) := + leakage_events (word.add program_base (word.of_Z mypos)). + + Definition project_tuple (tup : tuple) : nat * stmt Z := + let '(s, k, rk_so_far, mypos, sp_val, stackoffset, f) := tup in + (length k, s). + + Definition lt_tuple (x y : tuple) := + lt_tuple' (project_tuple x) (project_tuple y). + + Lemma lt_tuple_wf : well_founded lt_tuple. + Proof. + cbv [lt_tuple]. apply wf_inverse_image. apply lt_tuple'_wf. + Defined. + + Definition fun_leakage_helper + (mypos : Z) (sp_val ra_val : word) rets fbody := + let need_to_save := list_diff Z.eqb (modVars_as_list Z.eqb fbody) rets in + let scratchwords := stackalloc_words fbody in + let framesize := (bytes_per_word * + (Z.of_nat (length need_to_save) + 1 + scratchwords))%Z in + let sp_val' := word.add sp_val (word.of_Z (-framesize)) in + let beforeBodyInstrs := + [[ Addi sp sp (-framesize) ]] ++ + [[ compile_store access_size.word sp ra + (bytes_per_word * (Z.of_nat (length need_to_save) + scratchwords)) ]] ++ + save_regs need_to_save (bytes_per_word * scratchwords) in + let beforeBodyLeakage := + [ leak_Addi ] ++ (* Addi sp sp (-framesize) *) + [ leak_store access_size.word sp_val' ] ++ + leak_save_regs sp_val' need_to_save in + let afterBodyInstrs := + load_regs need_to_save (bytes_per_word * scratchwords) ++ + [[ compile_load access_size.word ra sp + (bytes_per_word * (Z.of_nat (length need_to_save) + scratchwords)) ]] ++ + [[ Addi sp sp framesize ]] ++ + [[ Jalr zero ra 0 ]] in + let afterBodyLeakage := + leak_load_regs sp_val' need_to_save ++ + [ leak_load access_size.word sp_val' ] ++ + [ leak_Addi ] ++ (* Addi sp sp framesize *) + [ leak_Jalr ra_val ] in + let body_pos := mypos + 4 * (2 + Z.of_nat (length need_to_save)) in + let body_stackoffset := (bytes_per_word * scratchwords)%Z in + let fbody_length := 4 * Z.of_nat (length (compile_stmt body_pos body_stackoffset fbody)) in + (beforeBodyInstrs, beforeBodyLeakage, afterBodyInstrs, afterBodyLeakage, body_pos, body_pos + fbody_length, sp_val', body_stackoffset). + + Definition stmt_leakage_body + (tup : tuple) + (stmt_leakage : forall othertup, lt_tuple othertup tup -> list LeakageEvent * word) + : list LeakageEvent * word. + refine ( + match tup as x return tup = x -> _ with + | (s, k, rk_so_far, mypos, sp_val, stackoffset, f) => + fun _ => + let leakage_events' := leakage_events_rel mypos (compile_stmt mypos stackoffset s) in + match s as x return s = x -> _ with + | SLoad sz x y o => + fun _ => + match k with + | leak_word addr :: k' => + f [leak_word addr] (rk_so_far ++ leakage_events' [leak_load sz (word.sub addr (word.of_Z o) (*this is silly to add it and then subtract it off again.... the alternative is to have Flattening trace transformation function not be the identity, though*))]) + | _ => (nil, word.of_Z 0) + end + | SStore sz x y o => + fun _ => + match k with + | leak_word addr :: k' => + f [leak_word addr] (rk_so_far ++ leakage_events' [leak_store sz (word.sub addr (word.of_Z o))]) + | _ => (nil, word.of_Z 0) + end + | SInlinetable sz x t i => + fun _ => + match k with + | leak_word i' :: k' => + f [leak_word i'] (rk_so_far ++ leakage_events_rel mypos [[ Jal x (4 + Z.of_nat (length (compile_byte_list t)) * 4)]] [leak_Jal] ++ leakage_events_rel (mypos + (4 + Z.of_nat (length (compile_byte_list t)) * 4)) [[Add x x i; compile_load sz x x 0 ]] [leak_Add; leak_load sz (word.add (word.add (word.add program_base (word.of_Z mypos)) (word.of_Z 4)) i')]) + | _ => (nil, word.of_Z 0) + end + | SStackalloc x n body => + fun _ => + match k as x return k = x -> _ with + | leak_unit :: k' => + fun _ => + stmt_leakage (body, k', rk_so_far ++ leakage_events_rel mypos [[Addi x sp (stackoffset-n)]] [ leak_Addi ], + mypos + 4, sp_val, stackoffset - n, fun skip => f (leak_unit :: skip)) _ + | _ => fun _ => (nil, word.add sp_val (word.of_Z (stackoffset - n))) + end eq_refl + | SLit _ v => + fun _ => + f [] (rk_so_far ++ leakage_events' (leak_lit v)) + | SOp _ op _ operand2 => + fun _ => + let newt_operands := + match op with + | Syntax.bopname.divu + | Syntax.bopname.remu => + match k with + | leak_word x1 :: leak_word x2 :: k' => Some ([leak_word x1; leak_word x2], x1, x2) + | _ => None + end + | Syntax.bopname.slu + | Syntax.bopname.sru + | Syntax.bopname.srs => + match k with + | leak_word x2 :: k' => Some ([leak_word x2], word.of_Z 0, x2) + | _ => None + end + | _ => Some ([], word.of_Z 0, word.of_Z 0) + end + in + match newt_operands with + | Some (newt, x1, x2) => + f newt (rk_so_far ++ leakage_events' (leak_op op operand2 x1 x2)) + | None => (nil, word.of_Z 0) + end + | SSet _ _ => + fun _ => + f [] (rk_so_far ++ leakage_events' [ leak_Add ]) + | SIf cond bThen bElse => + fun _ => + let thenLength := Z.of_nat (length (compile_stmt (mypos + 4) stackoffset bThen)) in + let elseLength := Z.of_nat (length (compile_stmt (mypos + 4 + 4 * thenLength + 4) stackoffset bElse)) in + match k as x return k = x -> _ with + | leak_bool b :: k' => + fun _ => + stmt_leakage (if b then bThen else bElse, + k', + rk_so_far ++ leakage_events_rel mypos [[compile_bcond_by_inverting cond ((thenLength + 2) * 4)]] [ leak_bcond_by_inverting cond (negb b) ], + if b then mypos + 4 else mypos + 4 + 4 * thenLength + 4, + sp_val, + stackoffset, + fun skip rk_so_far' => + f (leak_bool b :: skip) + (rk_so_far' ++ if b then leakage_events_rel (mypos + 4 + 4 * thenLength) [[Jal Register0 ((elseLength + 1) * 4)]] [leak_Jal] else [])) _ + | _ => fun _ => (nil, word.of_Z 0) + end eq_refl + | SLoop body1 cond body2 => + fun _ => + let body1Length := Z.of_nat (length (compile_stmt mypos stackoffset body1)) in + let body2Length := Z.of_nat (length (compile_stmt (mypos + (body1Length + 1) * 4) stackoffset body2)) in + stmt_leakage (body1, k, rk_so_far, mypos, sp_val, stackoffset, + fun skip rk_so_far' => + Let_In_pf_nd (List.skipn (length skip) k) + (fun k' _ => + match k' as x return k' = x -> _ with + | leak_bool true :: k'' => + fun _ => + stmt_leakage (body2, + k'', + rk_so_far' ++ leakage_events_rel (mypos + body1Length * 4) [[compile_bcond_by_inverting cond ((body2Length + 2) * 4)]] [ leak_bcond_by_inverting cond (negb true) ], + mypos + (body1Length + 1) * 4, sp_val, stackoffset, + fun skip' rk_so_far'' => + let k''' := List.skipn (length skip') k'' in + stmt_leakage (s, + k''', + rk_so_far'' ++ leakage_events_rel (mypos + (body1Length + 1 + body2Length) * 4) [[Jal Register0 (- (body1Length + 1 + body2Length) * 4)]] [ leak_Jal ], + mypos, sp_val, stackoffset, + fun skip'' => f (skip ++ leak_bool true :: skip' ++ skip'')) _) _ + | leak_bool false :: k'' => + fun _ => + f (skip ++ [leak_bool false]) (rk_so_far' ++ leakage_events_rel (mypos + body1Length * 4) [[compile_bcond_by_inverting cond ((body2Length + 2) * 4)]] [ leak_bcond_by_inverting cond (negb false) ]) + | _ => fun _ => (nil, word.of_Z 0) + end eq_refl)) _ + | SSeq s1 s2 => + fun _ => + let s1Length := Z.of_nat (length (compile_stmt mypos stackoffset s1)) in + stmt_leakage (s1, k, rk_so_far, mypos, sp_val, stackoffset, + fun skip rk_so_far' => + let k' := List.skipn (length skip) k in + stmt_leakage (s2, k', rk_so_far', mypos + 4 * s1Length, sp_val, stackoffset, + fun skip' => f (skip ++ skip')) _) _ + | SSkip => fun _ => f [] rk_so_far + | SCall resvars fname argvars => + fun _ => + match k as x return k = x -> _ with + | leak_unit :: k' => + fun _ => + match map.get e_env fname, map.get e fname with + | Some (params, rets, fbody), Some fpos => + let '(beforeBodyInstrs, beforeBodyLeakage, afterBodyInstrs, afterBodyLeakage, mypos', after_fun_pos, sp_val', stackoffset') := fun_leakage_helper fpos sp_val (word.add (word.add program_base (word.of_Z mypos)) (word.of_Z 4)) rets fbody in + stmt_leakage (fbody, + k', + rk_so_far ++ leakage_events_rel mypos [[ Jal ra (fpos - mypos) ]] [leak_Jal] ++ leakage_events_rel fpos beforeBodyInstrs beforeBodyLeakage, + mypos', sp_val', stackoffset', + fun skip rk_so_far' => + let k'' := List.skipn (length skip) k' in + f (leak_unit :: skip) (rk_so_far' ++ leakage_events_rel after_fun_pos afterBodyInstrs afterBodyLeakage)) _ + | _, _ => (nil, word.of_Z 0) + end + | _ => fun _ => (nil, word.of_Z 0) + end eq_refl + | SInteract _ _ _ => + fun _ => + match k with + | leak_list l :: k' => + f [leak_list l] (rk_so_far ++ leak_ext_call program_base e mypos stackoffset s l) + | _ => (nil, word.of_Z 0) + end + end eq_refl + end eq_refl). + Proof. (*copied verbatim from spilling*) + Unshelve. + all: intros. + all: cbv [lt_tuple project_tuple]. + all: subst. + all: repeat match goal with + | t := List.skipn ?n ?k |- _ => + let H := fresh "H" in + assert (H := List.skipn_length n k); subst t end. + all: try (right; constructor; constructor). + all: try (left; constructor; constructor). + - assert (H0 := skipn_length (length skip) k). left. + rewrite H. assert (H1:= @f_equal _ _ (@length _) _ _ e3). + simpl in H1. blia. + - assert (H := skipn_length (length skip) k). left. + assert (H1 := @f_equal _ _ (@length _) _ _ e3). simpl in H1. blia. + - destruct (length (List.skipn (length skip) k) =? length k)%nat eqn:E. + + apply Nat.eqb_eq in E. rewrite E. right. constructor. constructor. + + apply Nat.eqb_neq in E. left. blia. + Defined. + + Definition stmt_leakage + := Fix lt_tuple_wf _ stmt_leakage_body. + + Definition Equiv (x y : tuple) := + let '(x1, x2, x3, x4, x5, x6, fx) := x in + let '(y1, y2, y3, y4, y5, y6, fy) := y in + (x1, x2, x3, x4, x5, x6) = (y1, y2, y3, y4, y5, y6) /\ + forall k rk, + fx k rk = fy k rk. + + Lemma stmt_leakage_body_ext : + forall (x1 x2 : tuple) (f1 : forall y : tuple, lt_tuple y x1 -> list LeakageEvent * word) + (f2 : forall y : tuple, lt_tuple y x2 -> list LeakageEvent * word), + Equiv x1 x2 -> + (forall (y1 y2 : tuple) (p1 : lt_tuple y1 x1) (p2 : lt_tuple y2 x2), + Equiv y1 y2 -> f1 y1 p1 = f2 y2 p2) -> @stmt_leakage_body x1 f1 = @stmt_leakage_body x2 f2. + Proof. + intros. cbv [stmt_leakage_body]. cbv beta. + destruct x1 as [ [ [ [ [ [s_1 k_1] rk_so_far_1] mypos_1] sp_val_1] stackoffset_1] f_1]. + destruct x2 as [ [ [ [ [ [s_2 k_2] rk_so_far_2] mypos_2] sp_val_2] stackoffset_2] f_2]. + cbv [Equiv] in H. destruct H as [H1 H2]. injection H1. intros. subst. clear H1. + repeat (Tactics.destruct_one_match || rewrite H || apply H0 || cbv [Equiv] || intuition auto || match goal with | |- _ :: _ = _ :: _ => f_equal end || intuition auto(*why does putting this here make this work*)). + apply Let_In_pf_nd_ext. + repeat (Tactics.destruct_one_match || rewrite H || apply H0 || cbv [Equiv] || intuition auto || match goal with | |- _ :: _ = _ :: _ => f_equal end || intuition auto). + Qed. + + Lemma fix_step tup : + stmt_leakage tup = @stmt_leakage_body tup (fun y _ => stmt_leakage y). + Proof. + cbv [stmt_leakage]. + apply (@Fix_eq'_nondep _ _ lt_tuple_wf _ (stmt_leakage_body) Equiv eq). + { apply stmt_leakage_body_ext. } + { cbv [Equiv]. destruct tup as [ [ [ [ [ [x1 x2] x3] x4] x5] x6] fx]. intuition. } + Qed. + + Definition fun_leakage + (k : leakage) (rk_so_far : list LeakageEvent) (fpos : Z) (sp_val ra_val : word) + (rets : list Z) fbody + (f : leakage -> list LeakageEvent -> list LeakageEvent * word) := + let '(beforeBodyInstrs, beforeBodyLeakage, afterBodyInstrs, afterBodyLeakage, mypos', after_fun_pos, sp_val', stackoffset') := fun_leakage_helper fpos sp_val ra_val rets fbody in + stmt_leakage (fbody, k, rk_so_far ++ leakage_events_rel fpos beforeBodyInstrs beforeBodyLeakage, mypos', sp_val', stackoffset', + fun skip rk_so_far' => f skip (rk_so_far' ++ leakage_events_rel after_fun_pos afterBodyInstrs afterBodyLeakage)). End WithEnv. (* compiles all functions just to obtain their code size *) diff --git a/compiler/src/compiler/FlatToRiscvFunctions.v b/compiler/src/compiler/FlatToRiscvFunctions.v index 84aacc06e..b85883063 100644 --- a/compiler/src/compiler/FlatToRiscvFunctions.v +++ b/compiler/src/compiler/FlatToRiscvFunctions.v @@ -1,9 +1,11 @@ +Require Import bedrock2.LeakageSemantics. Require Import coqutil.Tactics.rdelta. Require Import coqutil.Tactics.rewr. Require Import coqutil.Datatypes.PropSet. Require Import Coq.Logic.FunctionalExtensionality. Require Import Coq.Logic.PropExtensionality. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Spec.Primitives. Require Import riscv.Platform.RiscvMachine. Require Import riscv.Platform.MetricRiscvMachine. @@ -53,9 +55,9 @@ Section Proofs. Context {env: map.map String.string (list Z * list Z * FlatImp.stmt Z)}. Context {M: Type -> Type}. Context {MM: Monads.Monad M}. - Context {RVM: Machine.RiscvProgram M word}. + Context {RVM: Machine.RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. - Context {ext_spec: Semantics.ExtSpec}. + Context {ext_spec: LeakageSemantics.ExtSpec}. Context {word_riscv_ok: RiscvWordProperties.word.riscv_ok word}. Context {locals_ok: map.ok locals}. Context {mem_ok: map.ok mem}. @@ -64,6 +66,7 @@ Section Proofs. Context {PR: MetricPrimitives.MetricPrimitives PRParams}. Context {BWM: bitwidth_iset width iset}. Context (compile_ext_call: pos_map -> Z -> Z -> stmt Z -> list Instruction). + Context (leak_ext_call: word -> pos_map -> Z -> Z -> stmt Z -> list word -> list LeakageEvent). Add Ring wring : (word.ring_theory (word := word)) (preprocess [autorewrite with rew_word_morphism], @@ -230,6 +233,10 @@ Section Proofs. (* `exists stack_trash frame_trash, ...` from goodMachine *) | |- exists _ _, _ = _ /\ _ = _ /\ (_ * _)%sep _ => eexists _, _; (split; [|split]); [..|wcancel_assumption]; blia + | |- exists _ _, _ => do 2 eexists; split; [align_trace|]; split; [reflexivity|]; intros; + rewrite fix_step; simpl; simpl_rev; repeat rewrite <- app_assoc; simpl; + cbn [leakage_events_rel leakage_events]; + try solve [repeat solve_word_eq word_ok || f_equal] | |- _ => solve [ solve_valid_machine word_ok ] | H:subset (footpr _) _ |- subset (footpr _) _ => eapply rearrange_footpr_subset; [ exact H | solve [ wwcancel ] ] @@ -385,7 +392,34 @@ Section Proofs. Hint Extern 3 (map.only_differ _ _ _) => eapply only_differ_trans_l; [eassumption|eauto with map_hints ..] : map_hints. *) + + Lemma leakage_events_app a i1 i2 l1 l2 : + length i1 = length l1 -> + leakage_events a (i1 ++ i2) (l1 ++ l2) = + leakage_events a i1 l1 ++ leakage_events (word.add a (word.of_Z (4 * Z.of_nat (length i1)))) i2 l2. + Proof. + intros. revert a l1 H. induction i1; intros ad l1 H. + - simpl. destruct l1; [|simpl in H; congruence]. + f_equal. solve_word_eq word_ok. + - simpl. destruct l1; [simpl in H; congruence|]. simpl. f_equal. + f_equal. simpl in H. injection H as H. rewrite IHi1 by assumption. f_equal. + f_equal. solve_word_eq word_ok. + Qed. + + Lemma length_load_regs i l x : + length (load_regs i l x) = length l. + Proof. + revert x. induction l. 1: reflexivity. + intros. simpl. rewrite IHl. reflexivity. + Qed. + Lemma length_leak_load_regs i x l : + length (leak_load_regs i x l) = length l. + Proof. + revert x. induction l. 1: reflexivity. + intros. simpl. rewrite IHl. reflexivity. + Qed. + Lemma compile_bcond_by_inverting_correct: forall cond (amt: Z) (initialL: RiscvMachineL) l b (Exec R Rexec: mem -> Prop), subset (footpr Exec) (of_list (initialL.(getXAddrs))) -> @@ -414,23 +448,23 @@ Section Proofs. (Platform.MetricLogging.addMetricLoads 1 (Platform.MetricLogging.addMetricInstructions 1 initialL.(getMetrics))) else (Platform.MetricLogging.addMetricJumps 1 (Platform.MetricLogging.addMetricLoads 1 (Platform.MetricLogging.addMetricInstructions 1 initialL.(getMetrics))))) /\ + finalL.(getTrace) = option_map (List.app (rev (leakage_events initialL.(getPc) [compile_bcond_by_inverting cond amt] [leak_bcond_by_inverting cond (negb b)]))) initialL.(getTrace) /\ valid_machine finalL). Proof. intros. get_run1_valid_for_free. inline_iff1. destruct_RiscvMachine initialL. - simulate'. + simulate'. simpl. destruct b. - (* don't branch *) - destruct cond; [destruct op | ]; - simpl in *; Simp.simp; repeat (simulate'; simpl_bools; simpl); try intuition congruence. + destruct cond; [destruct op | ]; destruct getTrace; + simpl in *; Simp.simp; repeat (simulate'; simpl_bools; simpl); intuition. - (* branch *) - destruct cond; [destruct op | ]; - simpl in *; Simp.simp; repeat (simulate'; simpl_bools; simpl); try intuition congruence. + destruct cond; [destruct op | ]; destruct getTrace; + simpl in *; Simp.simp; repeat (simulate'; simpl_bools; simpl); intuition. Qed. - - - Local Notation exec := (exec PostSpill isRegZ). + + Local Notation exec e pick_sp := (@exec _ _ _ _ _ _ _ _ PostSpill isRegZ pick_sp e). Definition cost_compile_spec mc := Platform.MetricLogging.addMetricInstructions 95 @@ -439,41 +473,50 @@ Section Proofs. (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) + (st0 : locals) (post outcome : leakage -> Semantics.trace -> mem -> locals -> MetricLog -> Prop) (argnames retnames : list Z) (body : stmt Z) (program_base : word) - (pos : Z) (ret_addr : word) (mach : RiscvMachineL) (e_impl : env) + (pos : Z) initialKL (ret_addr : word) (mach : RiscvMachineL) (e_impl : env) (e_pos : pos_map) (binds_count : nat) (insts : list Instruction) - (xframe : mem -> Prop) (t : list LogItem) (g : GhostConsts) + (xframe : mem -> Prop) pick_sp1 kH (t : list LogItem) (g : GhostConsts) cont (IH: forall (g0 : GhostConsts) (insts0 : list Instruction) (xframe0 : mem -> Prop) - (initialL : RiscvMachineL) (pos0 : Z), + (initialL : RiscvMachineL) (pos0 : Z) initialKL0 cont0, fits_stack (rem_framewords g0) (rem_stackwords g0) e_impl body -> compile_stmt iset compile_ext_call e_pos pos0 (bytes_per_word * rem_framewords g0) body = insts0 -> pos0 mod 4 = 0 -> getPc initialL = program_base + !pos0 -> + initialL.(getTrace) = Some initialKL0 -> iff1 (allx g0) ((xframe0 * program iset (program_base + !pos0) insts0)%sep * FlatToRiscvCommon.functions compile_ext_call program_base e_pos e_impl) -> goodMachine t m st0 g0 initialL -> + (forall k, pick_sp1 (k ++ kH) = snd (stmt_leakage iset compile_ext_call leak_ext_call e_pos e_impl_full program_base + (body, rev k, rev initialKL0, pos0, g0.(p_sp), (bytes_per_word * rem_framewords g0)%Z, cont0 k))) -> runsTo initialL (fun finalL => - exists finalTrace finalMH finalRegsH finalMetricsH, - outcome finalTrace finalMH finalRegsH finalMetricsH /\ + exists finalK finalTrace finalMH finalRegsH finalMetricsH, + outcome finalK finalTrace finalMH finalRegsH finalMetricsH /\ getPc finalL = getPc initialL + !(4 * #(Datatypes.length insts0)) /\ map.only_differ (getRegs initialL) (union (of_list (modVars_as_list Z.eqb body)) (singleton_set RegisterNames.ra)) (getRegs finalL) /\ (getMetrics finalL - getMetrics initialL <= lowerMetrics (finalMetricsH - mc))%metricsL /\ - goodMachine finalTrace finalMH finalRegsH g0 finalL)) - (HOutcome: forall (t' : Semantics.trace) (m' : mem) (mc' : MetricLog) (st1 : locals), - outcome t' m' st1 mc' -> + (exists kH'' finalKL, + finalK = kH'' ++ kH /\ + finalL.(getTrace) = Some finalKL /\ + forall k cont, + stmt_leakage iset compile_ext_call leak_ext_call e_pos e_impl_full program_base + (body, rev kH'' ++ k, rev initialKL0, pos0, g0.(p_sp), (bytes_per_word * rem_framewords g0)%Z, cont) = cont (rev kH'') (rev finalKL)) /\ + goodMachine finalTrace finalMH finalRegsH g0 finalL)) + (HOutcome: forall kH' t' m' mc' (st1 : locals), + outcome kH' t' m' st1 mc' -> exists (retvs : list word) (l' : locals), map.getmany_of_list st1 retnames = Some retvs /\ map.putmany_of_list_zip (List.firstn binds_count (reg_class.all reg_class.arg)) retvs l = Some l' /\ - post t' m' l' mc'), + post kH' t' m' l' mc'), (binds_count <= 8)%nat -> map.of_list_zip argnames argvs = Some st0 -> - exec e_impl_full body t m st0 mc outcome -> + exec pick_sp1 e_impl_full body kH t m st0 mc outcome -> map.getmany_of_list l (List.firstn (List.length argnames) (reg_class.all reg_class.arg)) = Some argvs -> map.extends e_impl_full e_impl -> @@ -488,13 +531,16 @@ Section Proofs. map.get (getRegs mach) RegisterNames.ra = Some ret_addr -> word.unsigned ret_addr mod 4 = 0 -> getPc mach = program_base + !pos -> + mach.(getTrace) = Some initialKL -> iff1 (allx g) ((xframe * program iset (program_base + !pos) insts)%sep * FlatToRiscvCommon.functions compile_ext_call program_base e_pos e_impl) -> goodMachine t m l g mach -> + (forall k, pick_sp1 (k ++ kH) = snd (fun_leakage iset compile_ext_call leak_ext_call e_pos e_impl_full program_base + (rev k) (rev initialKL) pos g.(p_sp) ret_addr retnames body (cont k))) -> runsToNonDet.runsTo (mcomp_sat (Run.run1 iset)) mach (fun finalL => - exists finalTrace finalMH finalRegsH finalMetricsH, - post finalTrace finalMH finalRegsH finalMetricsH /\ + exists finalK finalTrace finalMH finalRegsH finalMetricsH, + post finalK finalTrace finalMH finalRegsH finalMetricsH /\ getPc finalL = ret_addr /\ map.only_differ (getRegs mach) (union @@ -503,9 +549,16 @@ Section Proofs. (singleton_set RegisterNames.ra)) (getRegs finalL) /\ (getMetrics finalL - cost_compile_spec (getMetrics mach) <= lowerMetrics (finalMetricsH - mc))%metricsL /\ - goodMachine finalTrace finalMH finalRegsH g finalL). + (exists kH'' finalKL, + finalK = kH'' ++ kH /\ + finalL.(getTrace) = Some finalKL /\ + forall k cont, + fun_leakage iset compile_ext_call leak_ext_call e_pos e_impl_full program_base + (rev kH'' ++ k) (rev initialKL) pos g.(p_sp) ret_addr retnames body cont = + cont (rev kH'') (rev finalKL)) /\ + goodMachine finalTrace finalMH finalRegsH g finalL). Proof. - intros * IHexec OC BC OL Exb GetMany Ext GE FS C V Mo Mo' Gra RaM GPC A GM. + intros * IHexec OC BC OL Exb GetMany Ext GE FS C V Mo Mo' Gra RaM GPC GT A GM PSP. assert (valid_register RegisterNames.ra) by (cbv; auto). assert (valid_register RegisterNames.sp) by (cbv; auto). @@ -690,6 +743,7 @@ Section Proofs. replace y' with y; [exact H|] end. ring_simplify. reflexivity. } + { reflexivity. } { move OL at bottom. unfold map.extends. intros *. intro G. rewrite !map.get_put_dec. destruct_one_match. { @@ -745,6 +799,15 @@ Section Proofs. - blia. - wcancel_assumption. } + { simpl. intros. rewrite PSP. + cbv [fun_leakage fun_leakage_helper]. + simpl_rev. + rewrite BPW in *. repeat rewrite <- app_assoc. cbn [List.app]. + simpl. + f_equal. f_equal. f_equal. 2: reflexivity. f_equal. + remember (p_sp + _) as new_sp. eassert (Esp: new_sp = _). + 2: { rewrite Esp. reflexivity. } + subst. solve_word_eq word_ok. } } unfold goodMachine. @@ -757,7 +820,7 @@ Section Proofs. end. subst. match goal with - | H: outcome _ _ _ _ |- _ => rename H into HO + | H: outcome _ _ _ _ _ |- _ => rename H into HO end. match goal with | H: forall _, _ |- _ => @@ -931,7 +994,7 @@ Section Proofs. (* computed postcondition satisfies required postcondition: *) apply runsToDone. simpl_MetricRiscvMachine_get_set. - do 4 eexists. + do 5 eexists. match goal with | |- _ /\ _ /\ ?ODGoal /\ _ => assert ODGoal as OD end. @@ -1115,6 +1178,16 @@ Section Proofs. (* cost_compile_spec constraint: cost_compile_spec >= (...93...) i think? *) blia. + + do 2 eexists. split; [solve [align_trace]|]. simpl. cbv [final_trace]. simpl. + split; [reflexivity|]. intros. cbv [fun_leakage fun_leakage_helper]. + simpl_rev. rewrite BPW in *. repeat rewrite <- app_assoc in *. cbn [List.app] in *. + remember (p_sp - _) as new_sp. eassert (Esp: new_sp = _). + 2: { rewrite Esp in *. rewrite H2p7p2. f_equal. repeat f_equal. + cbv [leakage_events_rel]. rewrite leakage_events_app. + 2: { rewrite length_load_regs, length_leak_load_regs. reflexivity. } + rewrite length_load_regs. simpl. repeat solve_word_eq word_ok || f_equal. } + subst. solve_word_eq word_ok. + + rename l into lH, finalRegsH into lFH', finalRegsH' into lH', st0 into lFH, middle_regs into lL. @@ -1361,10 +1434,10 @@ Section Proofs. Lemma compile_stmt_correct: (forall resvars extcall argvars, - compiles_FlatToRiscv_correctly compile_ext_call + compiles_FlatToRiscv_correctly compile_ext_call leak_ext_call compile_ext_call (SInteract resvars extcall argvars)) -> (forall s, - compiles_FlatToRiscv_correctly compile_ext_call + compiles_FlatToRiscv_correctly compile_ext_call leak_ext_call (compile_stmt iset compile_ext_call) s). Proof. (* by induction on the FlatImp execution, symbolically executing through concrete RISC-V instructions, and using the IH for lists of abstract instructions (eg a then or else branch), @@ -1390,12 +1463,12 @@ Section Proofs. (postH := post) (g := {| allx := allx |}) (pos := pos) (extcall := action) (argvars := argvars) (resvars := resvars) (initialMH := m); simpl; - clear compile_ext_call_correct; cycle -1. + clear compile_ext_call_correct; cycle -2. { unfold goodMachine, valid_FlatImp_var in *. simpl. ssplit; eauto. } all: eauto using exec.interact, fits_stack_interact. + simpl. intros finalL A. destruct_RiscvMachine finalL. unfold goodMachine in *. simpl in *. destruct_products. subst. - do 4 eexists; ssplit; eauto. + do 5 eexists; ssplit; eauto. - idtac "Case compile_stmt_correct/SCall". (* We have one "map.get e fname" from exec, one from fits_stack, make them match *) @@ -1457,16 +1530,28 @@ Section Proofs. end. match goal with | |- ?G => replace G with - (runsToNonDet.runsTo (mcomp_sat (Run.run1 iset)) mach - (fun finalL : RiscvMachineL => exists - finalTrace finalMH finalRegsH finalMetricsH, - post finalTrace finalMH finalRegsH finalMetricsH /\ - getPc finalL = program_base + !pos + !(4 * #1) /\ - map.only_differ initialL_regs - (union (of_list (list_union Z.eqb binds [])) (singleton_set RegisterNames.ra)) - (getRegs finalL) /\ - (getMetrics finalL - initialL_metrics <= lowerMetrics (finalMetricsH - mc))%metricsL /\ - goodMachine finalTrace finalMH finalRegsH g finalL)) + (runsToNonDet.runsTo (mcomp_sat (Run.run1 iset)) mach + (fun finalL : RiscvMachineL => + exists + (finalK : leakage) (finalTrace : Semantics.trace) (finalMH : mem) + (finalRegsH : locals) (finalMetricsH : MetricLog), + post finalK finalTrace finalMH finalRegsH finalMetricsH /\ + getPc finalL = program_base + !pos + !(4 * #1) /\ + map.only_differ initialL_regs + (union (of_list (list_union Z.eqb binds [])) (singleton_set RegisterNames.ra)) + (getRegs finalL) /\ + (getMetrics finalL - initialL_metrics <= lowerMetrics (finalMetricsH - mc))%metricsL /\ + (exists (kH'' : list leakage_event) (finalKL : list LeakageEvent), + finalK = kH'' ++ k /\ + getTrace finalL = Some finalKL /\ + (forall (k0 : list leakage_event) + (cont0 : leakage -> list LeakageEvent -> list LeakageEvent * word), + stmt_leakage iset compile_ext_call leak_ext_call e_pos e_impl_full + program_base + (SCall binds fname args, rev kH'' ++ k0, rev initialKL, pos, p_sp, + (bytes_per_word * #(Datatypes.length unused_part_of_caller_frame))%Z, + cont0) = cont0 (rev kH'') (rev finalKL))) /\ + goodMachine finalTrace finalMH finalRegsH g finalL)) end. 2: { subst. reflexivity. } @@ -1502,7 +1587,7 @@ Section Proofs. } move C after A. match goal with - | H: exec _ body _ _ _ _ _ |- _ => rename H into Exb + | H: exec _ _ body _ _ _ _ _ _ |- _ => rename H into Exb end. move Exb before C. assert (Ext: map.extends e_impl_full e_impl). { @@ -1564,7 +1649,7 @@ Section Proofs. let T := type of IHexec in let T' := open_constr:( (forall (g : GhostConsts) (e_impl : env) (e_pos : pos_map) (program_base : word) (insts : list Instruction) (xframe : mem -> Prop) - (initialL : RiscvMachineL) (pos : Z), + (initialL : RiscvMachineL) (pos : Z) initialKL cont, map.extends e_impl_full e_impl -> good_e_impl e_impl e_pos -> fits_stack (rem_framewords g) (rem_stackwords g) e_impl body -> @@ -1575,22 +1660,28 @@ Section Proofs. pos mod 4 = 0 -> word.unsigned program_base mod 4 = 0 -> getPc initialL = program_base + !pos -> + getTrace initialL = Some initialKL -> iff1 (FlatToRiscvCommon.allx g) ((xframe * program iset (program_base + !pos) insts)%sep * functions program_base e_pos e_impl) -> goodMachine mid_log m st0 g initialL -> + (forall k0 : list leakage_event, + pick_sp1 (k0 ++ leak_unit :: k) = + snd + (stmt_leakage iset compile_ext_call leak_ext_call e_pos e_impl_full program_base + (body, rev k0, rev initialKL, pos, FlatToRiscvCommon.p_sp g, + (bytes_per_word * rem_framewords g)%Z, cont k0))) -> runsTo initialL (fun finalL : RiscvMachineL => exists - (finalTrace : Semantics.trace) (finalMH : mem) (finalRegsH : locals) + finalK (finalTrace : Semantics.trace) (finalMH : mem) (finalRegsH : locals) (finalMetricsH : MetricLog), - outcome finalTrace finalMH finalRegsH finalMetricsH /\ + outcome finalK finalTrace finalMH finalRegsH finalMetricsH /\ getPc finalL = getPc initialL + !(4 * #(Datatypes.length insts)) /\ map.only_differ (getRegs initialL) (union (of_list (modVars_as_list Z.eqb body)) (singleton_set RegisterNames.ra)) (getRegs finalL) /\ - (* (getMetrics finalL - initialL_metrics <= lowerMetrics (finalMetricsH - mc))%metricsL /\ *) - _ /\ + _ /\ _ /\ goodMachine finalTrace finalMH finalRegsH g finalL)) ) in replace T with T' in IHexec. 2: { @@ -1629,21 +1720,37 @@ Section Proofs. | H: (binds_count <= 8)%nat |- _ => rename H into BC end. move BC after OC. + (*We need to remember the definition of mach, or at least that getTrace mach = _ :: getTrace...*) + eassert (Hmach : mach.(getTrace) = _). + { subst mach. simpl. cbv [final_trace]. simpl. reflexivity. } repeat match goal with | x := _ |- _ => clearbody x end. - revert IHexec OC BC OL Exb GetMany Ext GE FS C V Mo Mo' Gra RaM GPC A GM. - eapply compile_function_body_correct. - } + eapply compile_function_body_correct; try eassumption. + 1: eapply OC. + intros. rewrite associate_one_left. rewrite H16. + cbv [fun_leakage]. rewrite fix_step. simpl_rev. cbn [stmt_leakage_body]. + rewrite H. rewrite GetPos. repeat rewrite <- app_assoc. f_equal. + subst g. cbv [FlatToRiscvCommon.p_sp]. rewrite Heqret_addr. + repeat Tactics.destruct_one_match. + instantiate (1 := fun k0 a b => cont (k0 ++ [leak_unit]) (leak_unit :: a) b). + repeat rewrite <- app_assoc. reflexivity. } subst mach. simpl_MetricRiscvMachine_get_set. - intros. fwd. eexists. eexists. eexists. eexists. + intros. fwd. do 5 eexists. split; [ eapply H0p0 | ]. split; eauto 8 with map_hints. split; eauto 8 with map_hints. - split; eauto 8 with map_hints. - (* cost_compile_spec constraint: cost_compile_spec + (1,1,1,0) <= cost_call *) - unfold cost_compile_spec, cost_call in *. - MetricsToRiscv.solve_MetricLog. + split. + { (* cost_compile_spec constraint: cost_compile_spec + (1,1,1,0) <= cost_call *) + unfold cost_compile_spec, cost_call in *. + MetricsToRiscv.solve_MetricLog. } + split; [|eassumption]. + do 2 eexists. split; [align_trace|]. split; [eassumption|]. + intros. simpl_rev. erewrite <- (H0p4p2 _ (fun x => cont0 (leak_unit :: x))). + rewrite fix_step. cbv [stmt_leakage_body fun_leakage]. rewrite H. rewrite GetPos. + repeat rewrite <- app_assoc. f_equal. subst g. cbv [FlatToRiscvCommon.p_sp]. + rewrite Heqret_addr. repeat Tactics.destruct_one_match. + repeat rewrite <- app_assoc. reflexivity. - idtac "Case compile_stmt_correct/SLoad". progress unfold Memory.load, Memory.load_Z in *. fwd. @@ -1654,13 +1761,12 @@ 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. unfold Memory.store, Memory.store_Z in *. - change Memory.store_bytes with (Platform.Memory.store_bytes(word:=word)) in *. + change store_bytes with (Platform.Memory.store_bytes(word:=word)) in H1. match goal with | H: Platform.Memory.store_bytes _ _ _ _ = _ |- _ => unshelve epose proof (store_bytes_frame H _) as P; cycle 2 @@ -1758,13 +1864,17 @@ Section Proofs. Z.to_euclidean_division_equations. blia. } + assert (sp_val :pick_sp1 k = p_sp + !(bytes_per_word * #(Datatypes.length remaining_frame))). + { specialize (H14 nil). simpl in H14. rewrite H14. + rewrite fix_step. simpl. simpl_addrs. reflexivity. } + rewrite sp_val in *. clear sp_val. + eapply runsTo_trans; simpl_MetricRiscvMachine_get_set. + eapply IHexec with (g := {| p_sp := p_sp; rem_framewords := Z.of_nat (List.length remaining_frame); |}) (program_base := program_base) (e_impl := e_impl) (pos := (pos + 4)%Z) - (a := p_sp + !bytes_per_word * !#(Datatypes.length remaining_frame)) (mStack := mStack) (mCombined := map.putmany mSmall mStack); simpl_MetricRiscvMachine_get_set; @@ -1772,16 +1882,20 @@ Section Proofs. rewrite ?@length_save_regs, ?@length_load_regs in *; simpl_word_exprs word_ok; ssplit; - cycle -5. + cycle -6. { reflexivity. } { eassumption. } { exists stack_trash, remaining_frame. ssplit. 1,2: reflexivity. wcancel_assumption. } { reflexivity. } { assumption. } + { intros. rewrite associate_one_left. rewrite H14. + simpl_rev. rewrite fix_step. simpl. simpl_addrs. + instantiate (2 := rev _). rewrite rev_involutive. reflexivity. } { match goal with | |- ?G => let t := type of Ab in replace G with t; [exact Ab|f_equal] end. + 1: solve_word_eq word_ok. rewrite @coqutil.Datatypes.List.length_flat_map with (n := Z.to_nat bytes_per_word). - simpl_addrs. rewrite !Z2Nat.id by blia. rewrite <- BPW. rewrite <- Z_div_exact_2; blia. - clear. intros. eapply LittleEndianList.length_le_split. @@ -1799,6 +1913,8 @@ Section Proofs. { safe_sidecond. } { safe_sidecond. } { safe_sidecond. } + { simpl. cbv [final_trace leakage_events_rel leakage_events]. + simpl. simpl_rev. simpl_addrs. reflexivity. } { etransitivity. 1: eassumption. wwcancel. } { match goal with | |- map.extends (map.put _ ?k ?v1) (map.put _ ?k ?v2) => replace v1 with v2 @@ -1835,6 +1951,8 @@ Section Proofs. - simpl. reflexivity. } eauto with map_hints. + * simpl_rev. cbv [leakage_events_rel leakage_events] in H7p6p2. simpl_addrs. + rewrite H7p6p2. reflexivity. * edestruct hl_mem_to_ll_mem with (mL := middle_mem) (mTraded := mStack') as (returned_bytes & L & Q). 1, 2: eassumption. @@ -1949,6 +2067,10 @@ Section Proofs. after_IH. all: try safe_sidecond. all: try safe_sidecond. + { simpl. subst. reflexivity. } + { intros. rewrite associate_one_left, H13. rewrite fix_step. + simpl_rev. simpl. cbn [leakage_events_rel leakage_events]. + repeat rewrite <- app_assoc. reflexivity. } * (* jump over else-branch *) simpl. intros. destruct_RiscvMachine middle. fwd. subst. @@ -1956,7 +2078,9 @@ Section Proofs. { eapply run_Jal0; try safe_sidecond. solve_divisibleBy4. } simpl_MetricRiscvMachine_get_set. - intros. destruct_RiscvMachine mid. fwd. run1done. finishcost. + intros. destruct_RiscvMachine mid. fwd. run1done. 1: finishcost. + intros. simpl. repeat rewrite <- app_assoc in *. rewrite H4p9p2. + repeat solve_word_eq word_ok || f_equal. - idtac "Case compile_stmt_correct/SIf/Else". (* execute branch instruction, which will jump over then-branch *) @@ -1979,9 +2103,16 @@ Section Proofs. all: after_IH. all: try safe_sidecond. all: try safe_sidecond. + { simpl. subst. reflexivity. } + { intros. rewrite associate_one_left, H13. rewrite fix_step. + simpl_rev. simpl. cbn [leakage_events_rel leakage_events]. + repeat rewrite <- app_assoc. reflexivity. } * (* 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. finishcost. + simpl. intros. destruct_RiscvMachine middle. fwd. subst. run1done. + 1: finishcost. + intros. simpl. repeat rewrite <- app_assoc in *. rewrite H4p9p2. + rewrite app_nil_r. reflexivity. - idtac "Case compile_stmt_correct/SLoop". match goal with @@ -2000,15 +2131,17 @@ Section Proofs. | H: iff1 allx ?RHS |- _ => match RHS with | context[compile_stmt _ _ _ ?POS _ body1] => - eapply IH1 with (g := {| allx := allx |}) (pos := POS) + eapply IH1 with (g := {| allx := allx |}) (pos := POS) (cont := fun _ => _) end end. all: after_IH. all: try safe_sidecond. all: try safe_sidecond. + 1: reflexivity. + intros. rewrite H18. rewrite fix_step. simpl. reflexivity. + simpl in *. simpl. intros. destruct_RiscvMachine middle. match goal with - | H: exists _ _ _ _, _ |- _ => destruct H as [ tH' [ mH' [ lH' [ mcH' H ] ] ] ] + | H: exists _ _ _ _ _, _ |- _ => destruct H as [ kH' [ tH' [ mH' [ lH' [ mcH' H ] ] ] ] ] end. fwd. destruct (eval_bcond lH' cond) as [condB|] eqn: E. @@ -2031,13 +2164,20 @@ Section Proofs. | H: iff1 allx ?RHS |- _ => match RHS with | context[compile_stmt _ _ _ ?POS _ body2] => - eapply IH2 with (g := {| allx := allx |}) (pos := POS) + eapply IH2 with (g := {| allx := allx |}) (pos := POS) (cont := fun _ => _) end end. all: after_IH. 1: eassumption. all: try safe_sidecond. - all: try safe_sidecond. } + all: try safe_sidecond. + 1: reflexivity. + intros. repeat rewrite associate_one_left, app_assoc. rewrite H18. + rewrite fix_step. simpl. simpl_rev. rewrite H3p4p2. + rewrite List.skipn_app_r by reflexivity. cbn [FixEq.Let_In_pf_nd]. + simpl_rev. repeat rewrite <- app_assoc. + cbn [leakage_events_rel leakage_events]. + repeat solve_word_eq word_ok || f_equal. reflexivity. } simpl in *. intros. destruct_RiscvMachine middle. fwd. (* jump back to beginning of loop: *) eapply runsToStep. @@ -2046,14 +2186,34 @@ Section Proofs. intros. destruct_RiscvMachine mid. fwd. eapply runsTo_trans. { (* 3rd application of IH: run the whole loop again *) - eapply IH12 with (g := {| allx := allx |}) (pos := pos). + eapply IH12 with (g := {| allx := allx |}) (pos := pos) (cont := fun _ => _). 1: eassumption. all: after_IH. all: try safe_sidecond. all: try safe_sidecond. - } + 1: reflexivity. + intros. rewrite associate_one_left, 2 app_assoc. rewrite H18. + rewrite fix_step. simpl. simpl_rev. rewrite H3p4p2. + rewrite List.skipn_app_r by reflexivity. cbn [FixEq.Let_In_pf_nd]. + simpl_rev. repeat rewrite <- app_assoc. + cbn [leakage_events_rel leakage_events]. + eassert (rev finalKL ++ _ = _) as ->. 2: rewrite H3p12p2. + { repeat rewrite <- app_assoc. simpl. repeat solve_word_eq word_ok || f_equal. } + rewrite List.skipn_app_r by reflexivity. + repeat solve_word_eq word_ok || reflexivity || f_equal. } (* at end of loop, just prove that computed post satisfies required post *) - simpl. intros. destruct_RiscvMachine middle. fwd. run1done. finishcost. + simpl. intros. destruct_RiscvMachine middle. fwd. run1done. 1: finishcost. + intros. rewrite H3p4p2. rewrite List.skipn_app_r by reflexivity. + cbn [FixEq.Let_In_pf_nd]. + simpl_rev. repeat rewrite <- app_assoc. + cbn [leakage_events_rel leakage_events]. + eassert (rev finalKL ++ _ = _) as ->. 2: rewrite H3p12p2. + { repeat rewrite <- app_assoc. simpl. repeat solve_word_eq word_ok || f_equal. } + rewrite List.skipn_app_r by reflexivity. + repeat solve_word_eq word_ok || reflexivity || f_equal. + eassert (rev finalKL0 ++ _ = _) as ->. 2: rewrite H3p22p2. + { repeat rewrite <- app_assoc. simpl. repeat solve_word_eq word_ok || f_equal. } + reflexivity. * (* false: done, jump over body2 *) eapply runsToStep. { @@ -2062,7 +2222,9 @@ Section Proofs. try safe_sidecond. } simpl_MetricRiscvMachine_get_set. - intros. destruct_RiscvMachine mid. fwd. run1done. finishcost. + intros. destruct_RiscvMachine mid. fwd. run1done. 1: finishcost. + rewrite H3p4p2. rewrite List.skipn_app_r by reflexivity. + cbn [FixEq.Let_In_pf_nd]. repeat solve_word_eq word_ok || f_equal. - idtac "Case compile_stmt_correct/SSeq". on hyp[(FlatImpConstraints.uses_standard_arg_regs s1); runsTo] @@ -2070,10 +2232,12 @@ Section Proofs. on hyp[(FlatImpConstraints.uses_standard_arg_regs s2); runsTo] do (fun H => rename H into IH2). eapply runsTo_trans. - + eapply IH1 with (g := {| allx := allx; |}) (pos := pos). + + eapply IH1 with (g := {| allx := allx; |}) (pos := pos) (cont := fun _ => _). all: after_IH. all: try safe_sidecond. all: try safe_sidecond. + 1: reflexivity. + intros. rewrite H14. rewrite fix_step. reflexivity. + simpl. intros. destruct_RiscvMachine middle. fwd. eapply runsTo_trans. * match goal with @@ -2087,7 +2251,13 @@ Section Proofs. 1: eassumption. all: try safe_sidecond. all: try safe_sidecond. + 1: reflexivity. + intros. rewrite app_assoc. rewrite H14. rewrite fix_step. simpl. + simpl_rev. rewrite H1p4p2. rewrite List.skipn_app_r by reflexivity. + reflexivity. * simpl. intros. destruct_RiscvMachine middle. fwd. run1done. + rewrite H1p4p2. rewrite List.skipn_app_r by reflexivity. + rewrite H1p12p2. reflexivity. - idtac "Case compile_stmt_correct/SSkip". run1done. @@ -2095,4 +2265,3 @@ Section Proofs. End Proofs. - diff --git a/compiler/src/compiler/FlatToRiscvLiterals.v b/compiler/src/compiler/FlatToRiscvLiterals.v index cb19be2b0..05c2eb1af 100644 --- a/compiler/src/compiler/FlatToRiscvLiterals.v +++ b/compiler/src/compiler/FlatToRiscvLiterals.v @@ -22,9 +22,9 @@ Section FlatToRiscvLiterals. Context {mem: map.map word byte}. Context {M: Type -> Type}. Context {MM: Monads.Monad M}. - Context {RVM: Machine.RiscvProgram M word}. + Context {RVM: Machine.RiscvProgramWithLeakage M word}. Context {PRParams: Primitives.PrimitivesParams M MetricRiscvMachine}. - Context {ext_spec: Semantics.ExtSpec}. + Context {ext_spec: LeakageSemantics.ExtSpec}. Context {word_riscv_ok: RiscvWordProperties.word.riscv_ok word}. Context {locals_ok: map.ok locals}. Context {mem_ok: map.ok mem}. @@ -87,7 +87,8 @@ Section FlatToRiscvLiterals. runsTo (withRegs (map.put initialL.(getRegs) x (word.of_Z v)) (withPc (add initialL.(getPc) d) (withNextPc (add initialL.(getNextPc) d) - (withMetrics (updateMetricsForLiteral v initialL.(getMetrics)) initialL)))) + (withMetrics (updateMetricsForLiteral v initialL.(getMetrics)) + (withLeakageEvents (Some (rev (leakage_events initialL.(getPc) insts (leak_lit iset v)))) initialL))))) post -> runsTo initialL post. Proof. @@ -95,17 +96,16 @@ Section FlatToRiscvLiterals. simpl in *. destruct_RiscvMachine initialL. subst d insts initialL_npc. - unfold compile_lit, updateMetricsForLiteral in *. - (* TODO once we're on 8.16, it should be possible to replace "F, P, N" by "*" - https://github.com/coq/coq/pull/15426 *) - rewrite bitwidth_matches in F, P, N. + unfold compile_lit, leakage_events, leak_lit, updateMetricsForLiteral in *. + rewrite bitwidth_matches in *. destruct_one_match_hyp; [|destruct_one_match_hyp]. - - unfold compile_lit_12bit in *. + - unfold compile_lit_12bit, leak_lit_12bit in *. run1det. simpl_word_exprs word_ok. match_apply_runsTo. erewrite signExtend_nop; eauto; try blia. - - unfold compile_lit_32bit in *. + destruct_one_match; reflexivity. + - unfold compile_lit_32bit, leak_lit_32bit in *. simpl in P. run1det. run1det. match_apply_runsTo. @@ -145,7 +145,8 @@ Section FlatToRiscvLiterals. } + solve_word_eq word_ok. + solve_word_eq word_ok. - - unfold compile_lit_64bit, compile_lit_32bit in *. + + simpl. destruct_one_match; reflexivity. + - unfold compile_lit_64bit, leak_lit_64bit, compile_lit_32bit, compile_lit_32bit in *. remember (signExtend 12 (signExtend 32 (bitSlice v 32 64))) as mid. remember (signExtend 32 (signExtend 32 (bitSlice v 32 64))) as hi. protect_equalities. @@ -188,6 +189,7 @@ Section FlatToRiscvLiterals. all: Btauto.btauto. + solve_word_eq word_ok. + solve_word_eq word_ok. + + destruct_one_match; reflexivity. Qed. Lemma compile_lit_correct_full: forall (initialL: RiscvMachineL) post x v R Rexec, @@ -202,7 +204,8 @@ Section FlatToRiscvLiterals. runsTo (withRegs (map.put initialL.(getRegs) x (word.of_Z v)) (withPc (add initialL.(getPc) d) (withNextPc (add initialL.(getNextPc) d) - (withMetrics (updateMetricsForLiteral v initialL.(getMetrics)) initialL)))) + (withMetrics (updateMetricsForLiteral v initialL.(getMetrics)) + (withLeakageEvents (Some (rev (leakage_events (initialL.(getPc)) insts (leak_lit iset v)))) initialL))))) post -> runsTo initialL post. Proof. (* by case distinction on literal size and symbolic execution through the instructions diff --git a/compiler/src/compiler/FlattenExpr.v b/compiler/src/compiler/FlattenExpr.v index 516c483e3..9d8767b6f 100644 --- a/compiler/src/compiler/FlattenExpr.v +++ b/compiler/src/compiler/FlattenExpr.v @@ -8,7 +8,7 @@ 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 bedrock2.LeakageSemantics bedrock2.MetricLeakageSemantics. Require Import coqutil.Macros.unique. Require Import Coq.Bool.Bool. Require Import coqutil.Datatypes.PropSet. @@ -28,6 +28,7 @@ Section FlattenExpr1. {mem: map.map word Byte.byte} {FlatImp_env: map.map string (list string * list string * FlatImp.stmt string)} {ext_spec: ExtSpec} + {pick_sp: PickSp} {NGstate: Type} {NG: NameGen String.string NGstate} {locals_ok: map.ok locals} @@ -346,13 +347,13 @@ Section FlattenExpr1. 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', - mid t' m' l' mc' -> + Lemma seq_with_modVars: forall env k t m (l: locals) mc s1 s2 mid post, + exec env s1 k t m l mc mid -> + (forall k' t' m' l' mc', + mid k' t' m' l' mc' -> map.only_differ l (FlatImp.modVars s1) l' -> - exec env s2 t' m' l' mc' post) -> - exec env (FlatImp.SSeq s1 s2) t m l mc post. + exec env s2 k' t' m' l' mc' post) -> + exec env (FlatImp.SSeq s1 s2) k t m l mc post. Proof. intros *. intros E1 E2. eapply @FlatImp.exec.seq. - eapply FlatImp.exec.intersect. @@ -368,15 +369,16 @@ Section FlattenExpr1. Goal True. idtac "FlattenExpr: Entering slow lemmas section". Abort. - Lemma flattenExpr_correct_aux : forall e fenv oResVar ngs1 ngs2 resVar s initialH initialL initialM initialMcH initialMcL finalMcH res t, + Lemma flattenExpr_correct_aux : forall e fenv oResVar ngs1 ngs2 resVar s initialH initialK finalKH initialL initialM initialMcH initialMcL finalMcH res t, flattenExpr ngs1 oResVar e = (s, resVar, ngs2) -> map.extends initialL initialH -> 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) -> - 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). + eval_expr initialM initialH e initialK initialMcH = Some (res, finalKH, finalMcH) -> + exec fenv s initialK t initialM initialL initialMcL + (fun finalKL t' finalM finalL finalMcL => + finalKL = finalKH /\ t' = t /\ finalM = initialM /\ map.get finalL resVar = Some res /\ + (finalMcL - initialMcL <= finalMcH - initialMcH)%metricsH). Proof. induction e; intros *; intros F Ex U D Ev; simpl in *; simp. @@ -392,7 +394,8 @@ Section FlattenExpr1. 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. cost_hammer. + eapply @FlatImp.exec.load; t_safe; rewrite ?word.add_0_r; try eassumption. + 1: reflexivity. cost_hammer. - (* expr.inlinetable *) repeat match goal with @@ -502,33 +505,33 @@ Section FlattenExpr1. 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, + Lemma flattenExpr_correct_with_modVars : forall e fenv oResVar ngs1 ngs2 resVar s initialK finalKH t m lH lL initialMcH initialMcL finalMcH res, flattenExpr ngs1 oResVar e = (s, resVar, ngs2) -> map.extends lL lH -> 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) -> - exec fenv s t m lL initialMcL (fun t' m' lL' finalMcL => + eval_expr m lH e initialK initialMcH = Some (res, finalKH, finalMcH) -> + exec fenv s initialK t m lL initialMcL (fun finalKL 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). + finalKL = finalKH /\ t' = t /\ m' = m /\ map.get lL' resVar = Some res /\ + (finalMcL - initialMcL <= finalMcH - initialMcH)%metricsH). Proof. intros *. intros F Ex U D Ev. - epose proof (flattenExpr_correct_aux _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ F Ex U D Ev) as P. + epose proof (flattenExpr_correct_aux _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ F Ex U D Ev) as P. eapply FlatImp.exec.intersect; cycle 1. - exact P. - rapply FlatImp.modVarsSound. exact P. Qed. - Lemma flattenExprs_correct: forall es fenv ngs1 ngs2 resVars s t m lH lL initialMcH initialMcL finalMcH resVals, + Lemma flattenExprs_correct: forall es fenv ngs1 ngs2 resVars s initialK finalKH t m lH lL initialMcH initialMcL finalMcH resVals, flattenExprs ngs1 es = (s, resVars, ngs2) -> map.extends lL lH -> map.undef_on lH (allFreshVars ngs1) -> disjoint (ExprImp.allVars_exprs es) (allFreshVars ngs1) -> - eval_call_args m lH es initialMcH = Some (resVals, finalMcH) -> + eval_call_args m lH es initialK initialMcH = Some (resVals, finalKH, finalMcH) -> (* List.option_all (List.map (eval_expr m lH) es) = Some resVals -> *) - exec fenv s t m lL initialMcL (fun t' m' lL' finalMcL => - t' = t /\ m' = m /\ + exec fenv s initialK t m lL initialMcL (fun finalKL t' m' lL' finalMcL => + finalKL = finalKH /\ t' = t /\ m' = m /\ map.getmany_of_list lL' resVars = Some resVals /\ map.only_differ lL (FlatImp.modVars s) lL' /\ (finalMcL - initialMcL <= finalMcH - initialMcH)%metricsH). @@ -594,14 +597,14 @@ Section FlattenExpr1. | intros; simpl in *; simp; repeat rewrite_match; t_safe ]. Lemma flattenBooleanExpr_correct_aux : - forall e fenv ngs1 ngs2 resCond (s: FlatImp.stmt string) (initialH initialL: locals) initialM t initialMcH initialMcL finalMcH res, + forall e fenv ngs1 ngs2 resCond (s: FlatImp.stmt string) (initialH initialL: locals) initialK finalKH initialM t initialMcH initialMcL finalMcH res, flattenExprAsBoolExpr ngs1 e = (s, resCond, ngs2) -> map.extends initialL initialH -> map.undef_on initialH (allFreshVars ngs1) -> disjoint (ExprImp.allVars_expr e) (allFreshVars ngs1) -> - eval_expr initialM initialH e initialMcH = Some (res, finalMcH) -> - exec fenv s t initialM initialL initialMcL (fun t' finalM finalL finalMcL => - t' = t /\ finalM = initialM /\ + eval_expr initialM initialH e initialK initialMcH = Some (res, finalKH, finalMcH) -> + exec fenv s initialK t initialM initialL initialMcL (fun finalKL t' finalM finalL finalMcL => + finalKL = finalKH /\ t' = t /\ finalM = initialM /\ FlatImp.eval_bcond finalL resCond = Some (negb (word.eqb res (word.of_Z 0))) /\ (finalMcL - initialMcL <= finalMcH - initialMcH)%metricsH). Proof. @@ -631,14 +634,14 @@ Section FlattenExpr1. Goal True. idtac "FlattenExpr: flattenBooleanExpr_correct_aux done". Abort. Lemma flattenBooleanExpr_correct_with_modVars: - forall e fenv ngs1 ngs2 resCond (s: FlatImp.stmt string) (initialH initialL: locals) initialM t initialMcH initialMcL finalMcH res, + forall e fenv ngs1 ngs2 resCond (s: FlatImp.stmt string) (initialH initialL: locals) initialK finalKH initialM t initialMcH initialMcL finalMcH res, flattenExprAsBoolExpr ngs1 e = (s, resCond, ngs2) -> map.extends initialL initialH -> map.undef_on initialH (allFreshVars ngs1) -> disjoint (ExprImp.allVars_expr e) (allFreshVars ngs1) -> - eval_expr initialM initialH e initialMcH = Some (res, finalMcH) -> - exec fenv s t initialM initialL initialMcL (fun t' finalM finalL finalMcL => - (t' = t /\ finalM = initialM /\ + eval_expr initialM initialH e initialK initialMcH = Some (res, finalKH, finalMcH) -> + exec fenv s initialK t initialM initialL initialMcL (fun finalKL t' finalM finalL finalMcL => + (finalKL = finalKH /\ t' = t /\ finalM = initialM /\ FlatImp.eval_bcond finalL resCond = Some (negb (word.eqb res (word.of_Z 0))) /\ (finalMcL - initialMcL <= finalMcH - initialMcH)%metricsH) /\ map.only_differ initialL (FlatImp.modVars s) finalL (* <-- added *)). @@ -686,16 +689,16 @@ Section FlattenExpr1. Lemma flattenStmt_correct_aux: forall eH eL, flatten_functions eH = Success eL -> - forall eH0 sH t m mcH lH post, - MetricSemantics.exec eH0 sH t m lH mcH post -> + forall eH0 sH k t m mcH lH post, + MetricLeakageSemantics.exec eH0 sH k t m lH mcH post -> eH0 = eH -> forall ngs ngs' sL lL mcL, flattenStmt ngs sH = (sL, ngs') -> map.extends lL lH -> map.undef_on lH (allFreshVars ngs) -> disjoint (ExprImp.allVars_cmd sH) (allFreshVars ngs) -> - 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 *) + exec eL sL k t m lL mcL (fun k' t' m' lL' mcL' => exists lH' mcH', + post k' 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 ExprImp.modVarsSound). In the previous proof, which was by induction @@ -748,7 +751,7 @@ Section FlattenExpr1. eapply @FlatImp.exec.stackalloc. 1: eassumption. intros. rename H2 into IHexec. eapply @FlatImp.exec.weaken. - { eapply IHexec; try reflexivity; try eassumption; maps. } + { subst a. 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. cost_hammer. } @@ -831,8 +834,8 @@ Section FlattenExpr1. postcondition offered by IHexec - the metric differences are now between the start and end of the loop rather than after the expression execution and the end of the loop *) - eapply @FlatImp.exec.loop with (mid2 := (fun t' m' lL' mcL' => exists lH' mcH', - mid t' m' lH' mcH' /\ + eapply @FlatImp.exec.loop with (mid2 := (fun k' t' m' lL' mcL' => exists lH' mcH', + mid k' t' m' lH' mcH' /\ map.extends lL' lH' /\ map.only_differ l (ExprImp.modVars c) lH' /\ (mcL' - mcL <= mcH' - mc)%metricsH)); @@ -870,7 +873,8 @@ Section FlattenExpr1. + simpl in *. intros. simp. rename l into lH, l' into lL'. rename l0 into argValNames. unfold flatten_functions in H. - pose proof (map.try_map_values_fw _ _ _ H _ _ H0). + Fail pose proof (@map.try_map_values_fw _ _ _ _ _ _ _ _ _ _ _ _ H _ _ H0). + pose proof (@map.try_map_values_fw _ _ _ _ _ _ _ Semantics.env_ok _ _ _ _ H _ _ H0). simp. unfold flatten_function, ExprImp2FlatImp in *. destruct v2 as [[params' rets'] fbody']. @@ -886,7 +890,7 @@ Section FlattenExpr1. -- clear -locals_ok. maps. -- unfold map.undef_on, map.agree_on. intros. rewrite map.get_empty. - destr (map.get lf k); [exfalso|reflexivity]. + destr (map.get lf k0); [exfalso|reflexivity]. eapply freshNameGenState_spec. 2: eassumption. pose proof H2 as G. unfold map.of_list_zip in G. @@ -930,12 +934,12 @@ Section FlattenExpr1. Qed. Goal True. idtac "FlattenExpr: flattenStmt_correct_aux done". Abort. - Lemma flattenStmt_correct: forall eH eL sH sL lL t m mc post, + Lemma flattenStmt_correct: forall eH eL sH sL lL k t m mc post, flatten_functions eH = Success eL -> ExprImp2FlatImp sH = sL -> - MetricSemantics.exec eH sH t m map.empty mc post -> - exec eL sL t m lL mc (fun t' m' lL' mcL' => exists lH' mcH', - post t' m' lH' mcH' /\ + MetricLeakageSemantics.exec eH sH k t m map.empty mc post -> + exec eL sL k t m lL mc (fun k' t' m' lL' mcL' => exists lH' mcH', + post k' t' m' lH' mcH' /\ map.extends lL' lH' /\ (mcL' - mc <= mcH' - mc)%metricsH). Proof. diff --git a/compiler/src/compiler/ForeverSafe.v b/compiler/src/compiler/ForeverSafe.v index 7bf0e073f..51ccb73ee 100644 --- a/compiler/src/compiler/ForeverSafe.v +++ b/compiler/src/compiler/ForeverSafe.v @@ -26,7 +26,7 @@ Section ForeverSafe. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {PR: MetricPrimitives PRParams}. Variable iset: Decode.InstructionSet. diff --git a/compiler/src/compiler/GoFlatToRiscv.v b/compiler/src/compiler/GoFlatToRiscv.v index 92cfd4cb5..4763e593c 100644 --- a/compiler/src/compiler/GoFlatToRiscv.v +++ b/compiler/src/compiler/GoFlatToRiscv.v @@ -22,6 +22,7 @@ Require Export coqutil.Word.SimplWordExpr. Require Import compiler.DivisibleBy4. Require Import bedrock2.ptsto_bytes. Require Import bedrock2.Scalars. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Utility.Encode. Require Import riscv.Proofs.EncodeBound. Require Import riscv.Proofs.DecodeEncode. @@ -42,7 +43,7 @@ Section Go. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {PR: MetricPrimitives PRParams}. @@ -236,6 +237,11 @@ Section Go. mcomp_sat (Bind (Machine.storeDouble kind addr v) f) initialL post. Proof. t spec_storeDouble. Qed. + Lemma go_leakEvent : forall (initialL: RiscvMachineL) (f: unit -> M unit) (e : option LeakageEvent) post, + mcomp_sat (f tt) (withLeakageEvent e initialL) post -> + mcomp_sat (Bind (leakEvent e) f) initialL post. + Proof. t spec_leakEvent. Qed. + Lemma go_getPC: forall (initialL: RiscvMachineL) (f: word -> M unit) post, mcomp_sat (f initialL.(getPc)) initialL post -> mcomp_sat (Bind getPC f) initialL post. @@ -427,13 +433,25 @@ Section Go. | _ => True end. + Lemma getXAddrs_withLeakageEvent e x : + getXAddrs (withLeakageEvent e x) = getXAddrs x. + Proof. destruct x, getMachine. reflexivity. Qed. + + Lemma getMem_withLeakageEvent e x : + getMem (withLeakageEvent e x) = getMem x. + Proof. destruct x, getMachine. reflexivity. Qed. + Lemma go_fetch_inst{initialL: RiscvMachineL} {inst pc0 R Rexec} (post: RiscvMachineL -> Prop): pc0 = initialL.(getPc) -> subset (footpr (program iset pc0 [inst] * Rexec)%sep) (of_list initialL.(getXAddrs)) -> (program iset pc0 [inst] * Rexec * R)%sep initialL.(getMem) -> not_InvalidInstruction inst -> - mcomp_sat (Bind (execute inst) (fun _ => endCycleNormal)) - (updateMetrics (addMetricLoads 1) initialL) post -> + mcomp_sat (Bind (leakage_of_instr getRegister inst) + (fun e => Bind (leakEvent e) + (fun _ => Bind (execute inst) + (fun _ => endCycleNormal)))) + (updateMetrics (addMetricLoads 1) + (withLeakageEvent (Some (fetchInstr (getPc initialL))) initialL)) post -> mcomp_sat (run1 iset) initialL post. Proof. intros. subst. @@ -446,10 +464,11 @@ Section Go. assert ((T * R * Rexec * P1 * P2)%sep m) as A by ecancel_assumption; clear H end. do 2 (apply sep_emp_r in A; destruct A as [A ?]). + apply go_leakEvent. eapply go_loadWord_Fetch. - - eapply ptsto_instr_subset_to_isXAddr4. + - rewrite getXAddrs_withLeakageEvent. eapply ptsto_instr_subset_to_isXAddr4. eapply shrink_footpr_subset. 1: eassumption. simpl. ecancel. - - unfold Memory.loadWord. + - rewrite getMem_withLeakageEvent. unfold Memory.loadWord. unfold truncated_scalar, littleendian, Memory.bytes_per in A. eapply load_bytes_of_sep with (n:=(length (LittleEndianList.le_split 4 (encode inst)))). (* TODO here it would be useful if seplog unfolded Memory.bytes_per for me, @@ -722,6 +741,7 @@ Ltac simpl_MetricRiscvMachine_get_set := getMem getXAddrs getLog + getTrace withRegs withPc withNextPc @@ -730,6 +750,8 @@ Ltac simpl_MetricRiscvMachine_get_set := withLog withLogItem withLogItems + withLeakageEvent + withLeakageEvents RiscvMachine.withRegs RiscvMachine.withPc RiscvMachine.withNextPc @@ -738,6 +760,8 @@ Ltac simpl_MetricRiscvMachine_get_set := RiscvMachine.withLog RiscvMachine.withLogItem RiscvMachine.withLogItems + RiscvMachine.withLeakageEvent + RiscvMachine.withLeakageEvents ]. Ltac simpl_MetricRiscvMachine_mem := @@ -838,6 +862,7 @@ Ltac simulate_step := | eapply go_loadDouble ; [sidecondition..|] | eapply go_storeDouble ; [sidecondition..|] *) + | refine (go_leakEvent _ _ _ _ _); [sidecondition..|] | refine (go_getPC _ _ _ _); [sidecondition..|] | refine (go_setPC _ _ _ _ _); [sidecondition..|] | refine (go_endCycleNormal _ _ _); [sidecondition..|] diff --git a/compiler/src/compiler/LowerPipeline.v b/compiler/src/compiler/LowerPipeline.v index 37ed158ff..4bfc2dd1d 100644 --- a/compiler/src/compiler/LowerPipeline.v +++ b/compiler/src/compiler/LowerPipeline.v @@ -1,8 +1,11 @@ + +Require Import bedrock2.LeakageSemantics. Require Import Coq.Logic.FunctionalExtensionality. Require Import coqutil.Map.Interface. Require Import coqutil.Map.MapEauto. Require Import coqutil.Tactics.Tactics. Require Import coqutil.Tactics.fwd. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Spec.Decode. Require Import riscv.Spec.Primitives. Require Import riscv.Platform.RiscvMachine. @@ -353,11 +356,12 @@ Section LowerPipeline. Context {M: Type -> Type}. Context {MM: Monads.Monad M}. - Context {RVM: Machine.RiscvProgram M word}. + Context {RVM: Machine.RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {PR: MetricPrimitives.MetricPrimitives PRParams}. - Context {ext_spec: Semantics.ExtSpec}. + Context {ext_spec: LeakageSemantics.ExtSpec}. Context {word_riscv_ok: RiscvWordProperties.word.riscv_ok word}. + Context (leak_ext_call: word -> pos_map -> Z -> Z -> stmt Z -> list word -> list LeakageEvent). Definition machine_ok{BWM: bitwidth_iset width iset} (p_functions: word)(stack_start stack_pastend: word) @@ -384,17 +388,19 @@ Section LowerPipeline. = Some vals. Hypothesis compile_ext_call_correct: forall resvars extcall argvars, - compiles_FlatToRiscv_correctly compile_ext_call compile_ext_call + compiles_FlatToRiscv_correctly compile_ext_call leak_ext_call compile_ext_call (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)(mc: MetricLog) - (post: Semantics.trace -> mem -> list word -> MetricLog -> Prop): Prop := + (p_funcs stack_pastend ret_addr : word)(f_name: string)(kL: list LeakageEvent) + (t: Semantics.trace)(mH: mem)(argvals: list word)(mc: MetricLog) + (post: list LeakageEvent -> 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 /\ - forall p_funcs stack_start stack_pastend ret_addr Rdata Rexec (initial: MetricRiscvMachine), + forall stack_start Rdata Rexec (initial: MetricRiscvMachine), map.get initial.(getRegs) RegisterNames.ra = Some ret_addr -> + initial.(getTrace) = Some kL -> initial.(getLog) = t -> raiseMetrics (cost_compile_spec initial.(getMetrics)) = mc -> word.unsigned ret_addr mod 4 = 0 -> @@ -403,11 +409,12 @@ Section LowerPipeline. word.unsigned (word.sub stack_pastend stack_start) mod bytes_per_word = 0 -> initial.(getPc) = word.add p_funcs (word.of_Z f_rel_pos) -> machine_ok p_funcs stack_start stack_pastend instrs mH Rdata Rexec initial -> - runsTo initial (fun final => exists mH' retvals, + runsTo initial (fun final => exists kL' mH' retvals, arg_regs_contain final.(getRegs) retvals /\ - post final.(getLog) mH' retvals (raiseMetrics final.(getMetrics)) /\ + post kL' final.(getLog) mH' retvals (raiseMetrics final.(getMetrics)) /\ map.only_differ initial.(getRegs) reg_class.caller_saved final.(getRegs) /\ final.(getPc) = ret_addr /\ + final.(getTrace) = Some kL' /\ machine_ok p_funcs stack_start stack_pastend instrs mH' Rdata Rexec final). Definition same_finfo_and_length: @@ -483,19 +490,34 @@ 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 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 mcH (fun t' m' l' mc' => - exists retvals, map.getmany_of_list l' retnames = Some retvals /\ - post t' m' retvals mc')) -> + forall fname kH kL t m argvals mcH post argnames retnames fbody l, + forall (p_funcs stack_pastend ret_addr : word), + let '(instrs, finfo, req_stack_size) := p2 in + let f_rel_pos := + match map.get finfo fname with + | Some f_rel_pos => f_rel_pos + | None => 0 + end in + (map.get p1 fname = Some (argnames, retnames, fbody) /\ + map.of_list_zip argnames argvals = Some l /\ + FlatImp.exec (pick_sp := fun k => snd (fun_leakage iset compile_ext_call leak_ext_call finfo p1 p_funcs + (skipn (length kH) (rev k)) (rev kL) f_rel_pos stack_pastend ret_addr retnames fbody (fun _ rk => (rk, word.of_Z 0)))) + PostSpill isRegZ p1 fbody kH t m l mcH + (fun kH' t' m' l' mc' => + exists retvals, map.getmany_of_list l' retnames = Some retvals /\ + post kH' t' m' retvals mc')) -> 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'). + riscv_call p2 p_funcs stack_pastend ret_addr fname kL t m argvals mcL + (fun kL' t m a mcL' => + exists mcH' kH' kH'', + metricsLeq (mcL' - mcL) (mcH' - mcH) /\ + kH' = kH'' ++ kH /\ + fst (fun_leakage iset compile_ext_call leak_ext_call finfo p1 p_funcs + (rev kH'') (rev kL) f_rel_pos stack_pastend ret_addr retnames fbody (fun _ rk => (rk, word.of_Z 0))) = rev kL' /\ + post kH' t m a mcH'). Proof. unfold riscv_call. - intros. destruct p2 as ((finstrs & finfo) & req_stack_size). + intros p1 p2. destruct p2 as ((finstrs & finfo) & req_stack_size). intros. match goal with | H: riscvPhase _ = _ |- _ => pose proof H as RP; unfold riscvPhase in H end. @@ -504,7 +526,7 @@ Section LowerPipeline. rewrite E0 in P. specialize P with (1 := H1p0). cbn in P. pose proof (compile_funs_finfo_idemp _ _ _ E0) as Q. subst r. fwd. - eexists. split. 1: eassumption. + eexists. split. 1: reflexivity. intros. assert (word.unsigned p_funcs mod 4 = 0). { unfold machine_ok in *. fwd. @@ -530,9 +552,9 @@ Section LowerPipeline. (pos := pos2) (program_base := p_funcs) (l := l) - (post := fun t' m' l' mc' => + (post := fun kH' t' m' l' mc' => (exists retvals, - map.getmany_of_list l' retnames = Some retvals /\ post t' m' retvals mc')). + map.getmany_of_list l' retnames = Some retvals /\ post kH' t' m' retvals mc')). eapply Q with (g := {| rem_stackwords := word.unsigned (word.sub stack_pastend stack_start) / bytes_per_word; @@ -559,6 +581,8 @@ Section LowerPipeline. { eassumption. } { eassumption. } { eassumption. } + { eassumption. } + { intros. rewrite <- H18. forget (k ++ kH) as kk. reflexivity. } + cbv beta. intros. fwd. rewrite <- Vp1. assert (HLR: Datatypes.length retnames = Datatypes.length retvals). { @@ -603,6 +627,7 @@ Section LowerPipeline. + eassumption. + assumption. + unfold machine_ok in *. fwd. assumption. + + rewrite H1. reflexivity. + simpl_g_get. reflexivity. + unfold machine_ok in *. subst. fwd. unfold goodMachine; simpl_g_get. ssplit. @@ -674,11 +699,14 @@ Section LowerPipeline. apply Hlength_stack_trash_words. } { reflexivity. } { assumption. } + + simpl. intros. simpl_rev. + rewrite List.skipn_app_r by (rewrite rev_length; reflexivity). + reflexivity. - cbv beta. unfold goodMachine. simpl_g_get. unfold machine_ok in *. intros. fwd. assert (0 < bytes_per_word). { (* TODO: deduplicate *) unfold bytes_per_word; simpl; destruct width_cases as [EE | EE]; rewrite EE; cbv; trivial. } - eexists _, _. ssplit. + eexists _, _, _. ssplit. + eapply map.getmany_of_list_extends. 1: eassumption. match goal with | H: map.getmany_of_list finalRegsH _ = Some _ |- @@ -691,13 +719,11 @@ Section LowerPipeline. symmetry. eapply map.getmany_of_list_length. exact GM. - + eexists. split; [|eassumption]. - subst. - apply raise_metrics_ineq in H10p3. - unfold_MetricLog. - cbn in H10p3. - cbn. - solve_MetricLog. + + do 3 eexists. ssplit. 4: eassumption. + * subst. apply raise_metrics_ineq in H11p3. + unfold_MetricLog. cbn in H11p3. cbn. solve_MetricLog. + * align_trace. + * rewrite rev_involutive. reflexivity. + eapply only_differ_subset. 1: eassumption. rewrite ListSet.of_list_list_union. rewrite ?singleton_set_eq_of_list. @@ -720,6 +746,9 @@ Section LowerPipeline. * unfold reg_class.get. subst k. cbn. exact I. * contradiction. + reflexivity. + + rewrite H11p4p1. + specialize (H11p4p2 nil). rewrite app_nil_r in H11p4p2. rewrite H11p4p2. + simpl. rewrite rev_involutive. reflexivity. + cbv [mem_available]. repeat rewrite ?(iff1ToEq (sep_ex1_r _ _)), ?(iff1ToEq (sep_ex1_l _ _)). exists (List.flat_map (fun x => HList.tuple.to_list (LittleEndian.split (Z.to_nat bytes_per_word) (word.unsigned x))) stack_trash). diff --git a/compiler/src/compiler/MMIO.v b/compiler/src/compiler/MMIO.v index df348f136..6fc67e893 100644 --- a/compiler/src/compiler/MMIO.v +++ b/compiler/src/compiler/MMIO.v @@ -5,6 +5,7 @@ Require Import compiler.util.Common. Require Import bedrock2.Semantics. Require Import riscv.Utility.Monads. Require Import compiler.FlatImp. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Spec.Decode. Require Import coqutil.sanity. Require Import riscv.Utility.MkMachineWidth. @@ -117,12 +118,36 @@ Section MMIO1. morphism (word.ring_morph (word := word)), constants [word_cst]). - Definition compile_ext_call(_: funname_env Z)(_ _: Z)(s: stmt Z) := - match s with - | SInteract resvars action argvars => compile_interact resvars action argvars - | _ => [] - end. + Definition leak_interact(abs_pos: word)(results: list Z) a (args: list Z) (leakage: list word): + list LeakageEvent := + match leakage with + | [addr'] => + if String.eqb "MMIOWRITE" a then + match results, args with + | [], [addr; val] => + leakage_events abs_pos [[ Sw addr val 0 ]] [ ILeakage (Sw_leakage addr') ] + | _, _ => [] (* invalid, excluded by ext_spec *) + end + else + match results, args with + | [res], [addr] => leakage_events abs_pos [[ Lw res addr 0 ]] [ ILeakage (Lw_leakage addr') ] + | _, _ => [] (* invalid, excluded by ext_spec *) + end + | _ => [] (*invalid*) + end. + Definition compile_ext_call(_: funname_env Z)(_ _: Z)(s: stmt Z) := + match s with + | SInteract resvars action argvars => compile_interact resvars action argvars + | _ => [] + end. + + Definition leak_ext_call(program_base: word)(_: funname_env Z)(pos _: Z)(s: stmt Z)(l: list word) := + match s with + | SInteract resvars action argvars => leak_interact (word.add program_base (word.of_Z pos)) resvars action argvars l + | _ => [] + end. + Section CompilationTest. Definition magicMMIOAddrLit: Z := 0x10024000. Variable addr: Z. @@ -246,7 +271,7 @@ Section MMIO1. Time Qed. Lemma compile_ext_call_correct: forall resvars extcall argvars, - FlatToRiscvCommon.compiles_FlatToRiscv_correctly compile_ext_call compile_ext_call + FlatToRiscvCommon.compiles_FlatToRiscv_correctly compile_ext_call leak_ext_call compile_ext_call (FlatImp.SInteract resvars extcall argvars). Proof. unfold FlatToRiscvCommon.compiles_FlatToRiscv_correctly. simpl. intros. @@ -257,17 +282,17 @@ Section MMIO1. destruct_RiscvMachine initialL. unfold FlatToRiscvCommon.goodMachine in *. match goal with - | H: forall _ _, outcome _ _ -> _ |- _ => specialize H with (mReceive := map.empty) + | H: forall _ _ _, outcome _ _ _ -> _ |- _ => specialize H with (mReceive := map.empty) end. destruct (String.eqb "MMIOWRITE" action) eqn: E; cbn [getRegs getPc getNextPc getMem getLog getMachine getMetrics getXAddrs] in *. + (* MMOutput *) progress simpl in *|-. match goal with - | H: FE310CSemantics.ext_spec _ _ _ _ _ |- _ => rename H into Ex + | H: FE310CSemantics.leakage_ext_spec _ _ _ _ _ |- _ => rename H into Ex end. unfold compile_interact in *. - cbv [FE310CSemantics.ext_spec] in Ex. + cbv [FE310CSemantics.leakage_ext_spec] in Ex. rewrite E in *. destruct Ex as (?&?&?&(?&?&?)&?). subst mGive argvals. repeat match goal with @@ -300,7 +325,7 @@ Section MMIO1. | H: map.split _ _ map.empty |- _ => rewrite map.split_empty_r in H; subst end. match goal with - | HO: outcome _ _, H: _ |- _ => specialize (H _ HO); rename H into HP + | HO: outcome _ _ _, H: _ |- _ => specialize (H _ _ HO); rename H into HP end. destruct g. FlatToRiscvCommon.simpl_g_get. simp. @@ -337,7 +362,7 @@ Section MMIO1. } repeat fwd. - unfold getReg. + unfold RiscvMachine.withLeakageEvent, getRegs, getReg. destr ((0 rename H into Ex + | H: FE310CSemantics.leakage_ext_spec _ _ _ _ _ |- _ => rename H into Ex end. unfold compile_interact in *. - cbv [FE310CSemantics.ext_spec] in Ex. + cbv [FE310CSemantics.leakage_ext_spec] in Ex. simpl in *|-. rewrite E in *. @@ -480,7 +508,7 @@ Section MMIO1. repeat fwd. - unfold getReg. + unfold getReg, getRegs, RiscvMachine.withLeakageEvent. destr ((0 _, OC: forall _, outcome _ _ |- _ => - epose proof (A (cons _ nil) (OC _)) as P; clear A + | A: forall _ _, outcome _ _ _ -> _, OC: forall _, outcome _ _ _ |- _ => + epose proof (A (cons _ nil) (cons _ nil) (OC _)) as P; clear A end. cbn in P. simp. @@ -513,7 +541,7 @@ Section MMIO1. specialize (Pp1 mKeep). rewrite map.split_empty_r in Pp1. specialize (Pp1 eq_refl). unfold setReg. destr ((0 Prop; - Call(p: Program)(funcname: string) - (t: trace)(m: mem)(argvals: list word)(mc: MetricLog) - (post: trace -> mem -> list word -> MetricLog -> Prop): Prop; + Leakage: Type; + Call(p: Program)(pick_sp: list Leakage -> word)(funcname: string) + (k: list Leakage)(t: trace)(m: mem)(argvals: list word)(mc: MetricLog) + (post: list Leakage -> trace -> mem -> list word -> MetricLog -> Prop): Prop; }. Record phase_correct{L1 L2: Lang} @@ -80,12 +83,17 @@ Section WithWordAndMem. phase_preserves_post: forall p1 p2, L1.(Valid) p1 -> compile p1 = Success p2 -> - 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' - ); + forall fname k1 k2 pick_sp2, + exists pick_sp1 k2'', + forall t m argvals mc1 post, + L1.(Call) p1 pick_sp1 fname k1 t m argvals mc1 post -> + forall mc2, + L2.(Call) p2 pick_sp2 fname k2 t m argvals mc2 + (fun k2' t m a mc2' => + exists mc1' k1', + post k1' t m a mc1' /\ + k2' = k2'' k1' /\ + metricsLeq (mc2' - mc2) (mc1' - mc1)); }. Arguments phase_correct : clear implicits. @@ -110,38 +118,45 @@ Section WithWordAndMem. Proof. unfold compose_phases. intros [V12 C12] [V23 C23]. - split; intros; fwd; [eauto|]. - erewrite f_equal; [eauto|]. - post_ext. - split; [|destruct 1 as (?&?&?&?&?)]; eauto with metric_arith. + split; intros; fwd; eauto. + specialize (V12 p1 a E H). + specialize (C23 a p2 V12 H0 fname nil k2 pick_sp2). + destruct C23 as [pick_sp_mid [f23 C23] ]. + specialize (C12 p1 a H E fname k1 nil pick_sp_mid). + destruct C12 as [pick_sp1 [f12 C12] ]. + exists pick_sp1. exists (fun k3' => f23 (f12 k3')). + intros. eapply C12 in H1. eapply C23 in H1. + eassert ((fun _ _ _ _ _ => exists _, _) = _) as ->. 2: exact H1. + post_ext. split; intros; fwd; eauto 10 with metric_arith. + Unshelve. all: exact EmptyMetricLog. Qed. Section WithMoreParams. Context {Zlocals: map.map Z word} {string_keyed_map: forall T: Type, map.map string T} (* abstract T for better reusability *) - {ext_spec: ExtSpec} + {ext_spec: LeakageSemantics.ExtSpec} {word_ok : word.ok word} {mem_ok: map.ok mem} {string_keyed_map_ok: forall T, map.ok (string_keyed_map T)} {Zlocals_ok: map.ok Zlocals}. - Context {ext_spec_ok : Semantics.ext_spec.ok ext_spec}. + Context {ext_spec_ok : LeakageSemantics.ext_spec.ok ext_spec}. (* for riscv phase: *) Context {Registers: map.map Z word}. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {word_riscv_ok: RiscvWordProperties.word.riscv_ok word}. Context {Registers_ok: map.ok Registers}. Context {PR: MetricPrimitives PRParams}. Context {iset: InstructionSet}. Context {BWM: bitwidth_iset width iset}. - Context (compile_ext_call : string_keyed_map Z -> Z -> Z -> FlatImp.stmt Z -> - list Instruction). + Context (compile_ext_call : string_keyed_map Z -> Z -> Z -> FlatImp.stmt Z -> list Instruction). + Context (leak_ext_call : word -> string_keyed_map Z -> Z -> Z -> FlatImp.stmt Z -> list word -> list LeakageEvent). Context (compile_ext_call_correct: forall resvars extcall argvars, - compiles_FlatToRiscv_correctly compile_ext_call compile_ext_call + compiles_FlatToRiscv_correctly compile_ext_call leak_ext_call compile_ext_call (FlatImp.SInteract resvars extcall argvars)). Context (compile_ext_call_length_ignores_positions: forall stackoffset posmap1 posmap2 c pos1 pos2, @@ -153,51 +168,39 @@ Section WithWordAndMem. Definition is12BitImmediate(x: Z) : bool := andb (Z.leb (-2048) x) (Z.ltb x 2048). - Definition locals_based_call_spec{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 (cost_spill_spec mc) (fun t' m' l' mc' => - exists retvals, map.getmany_of_list l' retnames = Some 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 := + Definition locals_based_call_spec {Var Cmd: Type}{locals: map.map Var word} + {string_keyed_map': forall T: Type, map.map string T} + (Exec: PickSp -> string_keyed_map' (list Var * list Var * Cmd)%type -> + Cmd -> leakage -> trace -> mem -> locals -> MetricLog -> + (leakage -> trace -> mem -> locals -> MetricLog -> Prop) -> Prop) + (spilled : bool)(e: string_keyed_map' (list Var * list Var * Cmd)%type) + (pick_sp: PickSp) (f: string) + (k: leakage)(t: trace)(m: mem)(argvals: list word)(mc: MetricLog) + (post: leakage -> 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'). - + map.of_list_zip argnames argvals = Some l /\ + Exec pick_sp e fbody k t m l (if spilled then mc else cost_spill_spec mc) + (fun k' t' m' l' mc' => + exists retvals, map.getmany_of_list l' retnames = Some retvals /\ + post k' t' m' retvals mc'). + Definition ParamsNoDup{Var: Type}: (list Var * list Var * FlatImp.stmt Var) -> Prop := fun '(argnames, retnames, body) => NoDup argnames /\ NoDup retnames. Definition SrcLang: Lang := {| Program := Semantics.env; Valid := map.forall_values ExprImp.valid_fun; - Call := locals_based_call_spec MetricSemantics.exec; + Call := locals_based_call_spec (fun pick_sp e => @MetricLeakageSemantics.exec _ _ _ _ _ _ e pick_sp) false; |}. + (* | *) (* | FlattenExpr *) (* V *) 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 PreSpill isRegStr); + Call := locals_based_call_spec (fun pick_sp e => @FlatImp.exec _ _ _ _ _ _ _ _ PreSpill isRegStr e pick_sp) false; |}. (* | *) @@ -216,27 +219,29 @@ 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 PreSpill isRegZ); + Call := locals_based_call_spec (fun pick_sp e => @FlatImp.exec _ _ _ _ _ _ _ _ PreSpill isRegZ e pick_sp) false; |}. + (* | *) (* | Spilling *) (* V *) 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_spilled (FlatImp.exec PostSpill isRegZ); + Call := locals_based_call_spec (fun e pick_sp => @FlatImp.exec _ _ _ _ _ _ _ _ PostSpill isRegZ pick_sp e) true; |}. + (* | *) (* | FlatToRiscv *) (* V *) - Definition RiscvLang: Lang := {| + Definition RiscvLang (p_funcs stack_pastend ret_addr : word) : Lang := {| Program := list Instruction * (* <- code of all functions concatenated *) string_keyed_map Z * (* <- position (offset) for each function *) Z; (* <- required stack space in XLEN-bit words *) (* bounds in instructions are checked by `ptsto_instr` *) Valid '(insts, finfo, req_stack_size) := True; - Call := riscv_call; + Call := (fun p _ => riscv_call p p_funcs stack_pastend ret_addr); |}. Lemma flatten_functions_NoDup: forall funs funs', @@ -252,6 +257,18 @@ Section WithWordAndMem. eapply H. eassumption. Qed. + (*TODO: I could've used this in some other compiler proof files too..*) + Definition remove_n_r {X : Type} (n : nat) (l : list X) := + rev (skipn n (rev l)). + + Lemma remove_n_r_spec {X : Type} (l1 l2 : list X) : + remove_n_r (length l2) (l1 ++ l2) = l1. + Proof. + cbv [remove_n_r]. rewrite rev_app_distr. rewrite List.skipn_app_r. + - apply rev_involutive. + - rewrite rev_length. reflexivity. + Qed. + Lemma flattening_correct: phase_correct SrcLang FlatWithStrVars flatten_functions. Proof. unfold SrcLang, FlatWithStrVars. @@ -262,7 +279,9 @@ Section WithWordAndMem. unfold valid_fun in *. intros. specialize H0 with (1 := H2). simpl in H0. eapply H0. } - unfold locals_based_call_spec. intros. fwd. + unfold locals_based_call_spec. intros. + exists (fun k => pick_sp2 (remove_n_r (length k1) k ++ k2)). + exists (fun k => remove_n_r (length k1) k ++ k2). intros. fwd. pose proof H0 as GF. unfold flatten_functions in GF. eapply map.try_map_values_fw in GF. 2: eassumption. @@ -272,7 +291,12 @@ Section WithWordAndMem. eapply FlatImp.exec.weaken. - eapply flattenStmt_correct_aux. + eassumption. - + eauto. + + eapply MetricLeakageSemantics.exec.exec_ext. + 1: eapply MetricLeakageSemantics.exec.exec_to_other_trace. + 1: eapply H1p2. + intros. simpl. simpl_rev. rewrite List.skipn_app_r. + 2: rewrite rev_length; reflexivity. + rewrite remove_n_r_spec. rewrite rev_involutive. reflexivity. + reflexivity. + match goal with | |- ?p = _ => rewrite (surjective_pairing p) @@ -293,7 +317,8 @@ Section WithWordAndMem. + eapply @freshNameGenState_disjoint_fbody. - simpl. intros. fwd. eexists. split. + eauto using map.getmany_of_list_extends. - + eexists. split; [|eassumption]. unfold cost_spill_spec in *; solve_MetricLog. + + do 2 eexists. split; [eassumption|]. rewrite remove_n_r_spec. + split; [reflexivity|]. unfold cost_spill_spec in *; solve_MetricLog. Qed. Lemma useimmediate_functions_NoDup: forall funs funs', @@ -323,7 +348,9 @@ Section WithWordAndMem. simpl in H0. assumption. } - unfold locals_based_call_spec. intros. fwd. + unfold locals_based_call_spec. intros. + exists (fun k => pick_sp2 ((rev (skipn (length k1) (rev k)) ++ k2))). + exists (fun k => remove_n_r (length k1) k ++ k2). intros. fwd. pose proof H0 as GI. unfold useimmediate_functions in GI. eapply map.try_map_values_fw in GI. 2: eassumption. @@ -332,12 +359,20 @@ Section WithWordAndMem. intros. eapply exec.weaken. - eapply useImmediate_correct_aux; eauto. - - simpl. destruct 1 as (?&?&?&?&?). - repeat (eexists; split; try eassumption). + eapply FlatImp.exec.exec_ext. + 1: eapply FlatImp.exec.exec_to_other_trace. + 1: eassumption. + intros. simpl. simpl_rev. rewrite List.skipn_app_r. + 2: rewrite rev_length; reflexivity. + simpl_rev. rewrite List.skipn_app_r. + 2: rewrite rev_length; reflexivity. + rewrite rev_involutive. reflexivity. + - simpl. intros. fwd. eexists. intuition eauto. + do 2 eexists. intuition eauto. + { rewrite remove_n_r_spec. reflexivity. } unfold cost_spill_spec in *; solve_MetricLog. Qed. - Lemma dce_functions_NoDup: forall funs funs', (forall f argnames retnames body, map.get funs f = Some (argnames, retnames, body) -> NoDup argnames /\ NoDup retnames) -> @@ -361,8 +396,17 @@ Section WithWordAndMem. intros. specialize H0 with (1 := H2). simpl in H0. assumption. } - - unfold locals_based_call_spec. intros. fwd. + unfold locals_based_call_spec. intros. + assert (E: (exists x, map.get p1 fname = Some x) -> map.get p1 fname = Some (match (map.get p1 fname) with | Some finfo => finfo | None => (nil, nil, SSkip) end)). + { intros. destruct H1 as [x H1]. destruct (map.get p1 fname); congruence. } + destruct (match map.get p1 fname with + | Some finfo => finfo + | None => ([], [], SSkip) + end) as [ [argnames0 retnames0] fbody0 ]. + eexists. eexists. intros. fwd. + assert (H1p0': map.get p1 fname = Some (argnames0, retnames0, fbody0)). + { eapply E. eexists. apply H1p0. } + rewrite H1p0 in H1p0'. inversion H1p0'. subst. clear H1p0'. pose proof H0 as GI. unfold dce_functions in GI. eapply map.try_map_values_fw in GI. 2: eassumption. @@ -370,13 +414,26 @@ Section WithWordAndMem. eexists _, _, _, _. split. 1: eassumption. split. 1: eassumption. intros. eapply @exec.weaken. - - eapply dce_correct_aux; eauto. - eapply MapEauto.agree_on_refl. + - eapply exec.exec_ext. 1: eapply dce_correct_aux; eauto. + { eapply MapEauto.agree_on_refl. } + 2: { intros. simpl. instantiate (1 := fun x => pick_sp2 (rev x)). + simpl. rewrite rev_involutive. reflexivity. } + intros. remember (k ++ k1) as kk eqn:Hkk. + rewrite <- (remove_n_r_spec k k1). forget (k ++ k1) as kk0. + subst. instantiate (2 := fun _ => _). simpl. reflexivity. - unfold compile_post. intros. fwd. exists retvals. split. + erewrite MapEauto.agree_on_getmany; [ eauto | eapply MapEauto.agree_on_comm; [ eassumption ] ]. - + eexists; split; eauto. unfold cost_spill_spec in *; solve_MetricLog. + + do 3 eexists; ssplit; eauto. 2: unfold cost_spill_spec in *; solve_MetricLog. + instantiate (1 := fun _ => _). simpl. + specialize (H1p8 nil (fun _ => nil)). simpl in H1p8. + do 2 rewrite app_nil_r in H1p8. + rewrite <- (rev_involutive kL''). rewrite <- H1p8. + remember (kH'' ++ k1) as kk eqn:Hkk. + rewrite <- (remove_n_r_spec kH'' k1). + rewrite <- Hkk. reflexivity. + Unshelve. all: auto. Qed. Lemma regalloc_functions_NoDup: forall funs funs', @@ -403,7 +460,8 @@ Section WithWordAndMem. eapply regalloc_functions_NoDup; eassumption. } unfold locals_based_call_spec. - intros. fwd. + intros. exists (fun k => pick_sp2 (remove_n_r (length k1) k ++ k2)). + exists (fun kH'' => remove_n_r (length k1) kH'' ++ k2). intros. fwd. pose proof H0 as GR. unfold regalloc_functions in GR. fwd. rename E into GR. @@ -426,13 +484,20 @@ Section WithWordAndMem. unfold map.of_list_zip in *. eapply FlatImp.exec.weaken. - eapply checker_correct; eauto. - eapply states_compat_precond. - 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. + + eapply FlatImp.exec.exec_ext. + 1: eapply FlatImp.exec.exec_to_other_trace. + 1: eassumption. + intros. simpl. rewrite remove_n_r_spec. simpl_rev. rewrite List.skipn_app_r. + 2: rewrite rev_length; reflexivity. + rewrite rev_involutive. reflexivity. + + eapply states_compat_precond. + 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. + eauto using states_compat_getmany. - + eexists. split; [|eassumption]. unfold cost_spill_spec in *; solve_MetricLog. + + do 2 eexists. rewrite remove_n_r_spec. intuition eauto. + unfold cost_spill_spec in *; solve_MetricLog. Qed. Ltac debool := @@ -486,7 +551,12 @@ Section WithWordAndMem. Proof. unfold FlatWithZVars, FlatWithRegs. split; cbn. 1: exact spilling_preserves_valid. - unfold locals_based_call_spec, locals_based_call_spec_spilled. intros. fwd. + unfold locals_based_call_spec. intros. + set (finfo := (match (map.get p1 fname) with | Some finfo => finfo | None => (nil, nil, SSkip) end)). + exists (fun k => snd (fun_leakage p1 pick_sp2 finfo (skipn (length k1) (rev k)) (rev k2))). + exists (fun k1' => rev (fst (fun_leakage p1 pick_sp2 finfo (rev (remove_n_r (length k1) k1')) (rev k2)))). + intros. fwd. + subst finfo. rewrite H1p0 in *. pose proof H0 as GL. unfold spill_functions in GL. eapply map.try_map_values_fw in GL. 2: eassumption. @@ -505,16 +575,48 @@ Section WithWordAndMem. 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). + unfold call_spec. intros * E. rewrite E in *. fwd. exact H1p2. + - cbv beta. intros. fwd. eexists. intuition eauto. + do 2 eexists. intuition eauto. rewrite remove_n_r_spec. + rewrite H2p1p2. rewrite rev_involutive. reflexivity. + Qed. + + Lemma riscv_call_weaken : + forall p p_funcs stack_pastend ret_addr funcname + k t m argvals mc post1, + riscv_call p p_funcs stack_pastend ret_addr funcname k t m argvals mc post1 -> + forall post2, + (forall k' t' m' retvals mc', + post1 k' t' m' retvals mc' -> post2 k' t' m' retvals mc') -> + riscv_call p p_funcs stack_pastend ret_addr funcname k t m argvals mc post2. + Proof. + intros. cbv [riscv_call] in *. destruct p. destruct p. + destruct H as [f_rel_pos H]. exists f_rel_pos. intuition eauto. + cbv [runsTo]. eapply runsToNonDet.runsTo_weaken. 1: eapply H2. + all: try eassumption. simpl. intros. fwd. do 3 eexists. + intuition eauto. Qed. - Lemma riscv_phase_correct: phase_correct FlatWithRegs RiscvLang (riscvPhase compile_ext_call). + Lemma riscv_phase_correct p_funcs stack_pastend ret_addr : phase_correct FlatWithRegs (RiscvLang p_funcs stack_pastend ret_addr) (riscvPhase compile_ext_call). Proof. unfold FlatWithRegs, RiscvLang. split; cbn. - intros p1 ((? & finfo) & ?). intros. exact I. - - eapply flat_to_riscv_correct; eassumption. + - intros. set (finfo := match (map.get p1 fname) with | Some finfo => finfo | None => (nil, nil, SSkip) end). + destruct finfo as [ [argnames0 retnames0] fbody0 ] eqn:E. + destruct p2 as [ [instrs finfo_] req_stack_size] eqn:E'. + eexists. eexists. intros. + cbv [locals_based_call_spec] in H1. fwd. subst finfo. rewrite H1p0 in E. + inversion E. subst. pose proof flat_to_riscv_correct as H'. + repeat specialize (H' ltac:(assumption)). + specialize (H' (instrs, finfo_, req_stack_size)). + do 2 specialize (H' ltac:(assumption)). cbv beta match in H'. idtac. + eapply riscv_call_weaken. 1: eapply H'; eauto. + simpl. intros. fwd. do 2 eexists. split; [eassumption|]. split; [|eassumption]. + rewrite <- (rev_involutive k'). rewrite <- H1p5. + remember (kH'' ++ k1) as k''. replace kH'' with (remove_n_r (length k1) k''). + 2: { subst. apply remove_n_r_spec. } + instantiate (1 := fun _ => _). simpl. reflexivity. Qed. Definition composed_compile: @@ -527,7 +629,7 @@ Section WithWordAndMem. (compose_phases spill_functions (riscvPhase compile_ext_call)))))). - Lemma composed_compiler_correct: phase_correct SrcLang RiscvLang composed_compile. + Lemma composed_compiler_correct p_funcs stack_pastend ret_addr : phase_correct SrcLang (RiscvLang p_funcs stack_pastend ret_addr) composed_compile. Proof. unfold composed_compile. exact (compose_phases_correct flattening_correct @@ -535,7 +637,7 @@ Section WithWordAndMem. (compose_phases_correct dce_correct (compose_phases_correct regalloc_correct (compose_phases_correct spilling_correct - riscv_phase_correct))))). + (riscv_phase_correct _ _ _)))))). Qed. Definition compile(funs: list (string * (list string * list string * cmd))): @@ -581,27 +683,29 @@ Section WithWordAndMem. (* output of compilation: *) (instrs: list Instruction) (finfo: list (string * Z)) (req_stack_size: Z) (* function we choose to call: *) - (fname: string) + (fname: string) (p_funcs stack_hi ret_addr : word), (* high-level initial state & post on final state: *) - (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) -> + forall kH kL, + exists f pick_sp, + forall t mH argvals mc (post: leakage -> trace -> mem -> list word -> MetricLog -> Prop), (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 /\ - MetricSemantics.exec (map.of_list functions) fbody t mH l + MetricLeakageSemantics.exec (pick_sp := pick_sp) (map.of_list functions) fbody kH t mH l (cost_spill_spec mc) - (fun t' m' l' mc' => exists retvals: list word, + (fun kH' t' m' l' mc' => exists retvals: list word, map.getmany_of_list l' retnames = Some retvals /\ - post t' m' retvals mc')) -> + post kH' 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: *) (initial: MetricRiscvMachine) (* ghost vars that help describe the low-level machine: *) - (stack_lo stack_hi: word) (ret_addr p_funcs: word) (Rdata Rexec: mem -> Prop), + (stack_lo : word) (Rdata Rexec: mem -> Prop), 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 -> initial.(getPc) = word.add p_funcs (word.of_Z f_rel_pos) -> @@ -609,21 +713,23 @@ Section WithWordAndMem. word.unsigned ret_addr mod 4 = 0 -> arg_regs_contain initial.(getRegs) argvals -> initial.(getLog) = t -> + initial.(getTrace) = Some kL -> 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, + exists kH' mH' retvals, arg_regs_contain (getRegs final) 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 /\ + (exists mcH' : MetricLog, + ((raiseMetrics final.(getMetrics)) - mcL <= mcH' - mc)%metricsH /\ + post kH' (getLog final) mH' retvals mcH') /\ + final.(getTrace) = Some (f kH') /\ + 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. - pose proof (phase_preserves_post composed_compiler_correct) as C. - unfold Call, SrcLang, RiscvLang, locals_based_call_spec, riscv_call in C. + pose proof (phase_preserves_post (composed_compiler_correct p_funcs stack_hi ret_addr)) as C. + unfold Call, Leakage, SrcLang, RiscvLang, locals_based_call_spec, riscv_call in C. eapply valid_src_funs_correct in H. specialize C with (1 := H). assert (composed_compile (map.of_list functions) = @@ -632,10 +738,14 @@ Section WithWordAndMem. rewrite map.of_list_tuples. reflexivity. } specialize C with (1 := H0'). - specialize C with (1 := H1). - specialize (C mcL). - cbv iota in C. - fwd. eauto 10. + specialize (C fname kH kL (fun _ => word.of_Z 0)). + destruct C as [pick_spH [L C] ]. + eexists. intros. eexists. intros. + specialize C with (1 := H1). specialize (C mcL). + destruct C as [f_rel_pos C]. exists f_rel_pos. fwd. + intuition eauto. + eapply runsToNonDet.runsTo_weaken. + 1: eapply Cp1; eauto. cbv beta. intros. fwd. do 3 eexists. intuition eauto. Qed. Definition instrencode(p: list Instruction): list byte := @@ -655,19 +765,19 @@ Section WithWordAndMem. (* output of compilation: *) (instrs: list Instruction) (finfo: list (string * Z)) (req_stack_size: Z) (* function we choose to call: *) - (fname: string) (f_rel_pos: Z) + (fname: string) (p_funcs stack_hi ret_addr : word) (f_rel_pos: Z), + valid_src_funs fs = true -> + compile fs = Success (instrs, finfo, req_stack_size) -> + forall kH kL, + exists f pick_sp, forall (* high-level initial state & post on final state: *) - (t: trace) (mH: mem) (argvals: list word) (mc: MetricLog) (post: trace -> mem -> list word -> MetricLog -> Prop) + (t: trace) (mH: mem) (argvals: list word) (mcH : MetricLog) (mcL : MetricLog) (post: leakage -> 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) + (stack_lo : word) (Rdata Rexec: mem -> Prop) (* low-level machine on which we're going to run the compiled program: *) (initial: MetricRiscvMachine), - valid_src_funs fs = true -> NoDup (map fst fs) -> - compile fs = Success (instrs, finfo, req_stack_size) -> - MetricWeakestPrecondition.call (map.of_list fs) fname t mH argvals - (cost_spill_spec mc) post -> - forall mcL, + MetricLeakageSemantics.call (pick_sp := pick_sp) (map.of_list fs) fname kH t mH argvals (cost_spill_spec mcH) post -> 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 -> @@ -676,27 +786,30 @@ Section WithWordAndMem. word.unsigned ret_addr mod 4 = 0 -> arg_regs_contain initial.(getRegs) argvals -> initial.(getLog) = t -> + initial.(getTrace) = Some kL -> 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, - arg_regs_contain (getRegs final) 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). + exists kH' mH' retvals, + arg_regs_contain (getRegs final) retvals /\ + (exists mcH' : MetricLog, + ((raiseMetrics final.(getMetrics)) - mcL <= mcH' - mcH)%metricsH /\ + post kH' (getLog final) mH' retvals mcH') /\ + final.(getTrace) = Some (f kH') /\ + 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 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 MetricSemantics.call in WP. fwd. - do 5 eexists. 1: eassumption. split. 1: eassumption. - intros. assumption. + destruct (compiler_correct fs instrs finfo req_stack_size fname p_funcs stack_hi ret_addr H H0 kH kL) as + [f [pick_sp compiler_correct'] ]. + exists f, pick_sp. intros. + let H := hyp (MetricLeakageSemantics.call (pick_sp := pick_sp)) in rename H into WP. + unfold MetricLeakageSemantics.call in WP. fwd. + edestruct compiler_correct' as (f_rel_pos' & G & C). + { do 4 eexists. eauto. } + rewrite H3 in G. inversion G; subst; clear G. + apply C; auto. Qed. End WithMoreParams. diff --git a/compiler/src/compiler/RegAlloc.v b/compiler/src/compiler/RegAlloc.v index 20ceb8882..fcd67a3c3 100644 --- a/compiler/src/compiler/RegAlloc.v +++ b/compiler/src/compiler/RegAlloc.v @@ -1,4 +1,5 @@ Require Import compiler.util.Common. +Require Import bedrock2.LeakageSemantics. Require Import compiler.FlatImp. Require Import coqutil.Tactics.simpl_rewrite. Require Import Coq.Lists.List. Import ListNotations. @@ -916,7 +917,8 @@ Section CheckerCorrect. Context {impLocalsOk: map.ok impLocals}. Context {srcEnv: map.map String.string (list srcvar * list srcvar * stmt)} {srcEnvOk: map.ok srcEnv}. Context {impEnv: map.map String.string (list impvar * list impvar * stmt')} {impEnvOk: map.ok impEnv}. - Context {ext_spec: Semantics.ExtSpec}. + Context {ext_spec: LeakageSemantics.ExtSpec}. + Context {pick_sp: PickSp}. Definition states_compat(st: srcLocals)(corresp: list (srcvar * impvar))(st': impLocals) := forall (x: srcvar) (x': impvar), @@ -1192,16 +1194,16 @@ Section CheckerCorrect. 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, + Lemma checker_correct: forall (e: srcEnv) (e': impEnv) s k t m lH mcH post, check_funcs e e' = Success tt -> - exec PreSpill isRegStr e s t m lH mcH post -> + exec PreSpill isRegStr e s k 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 PreSpill isRegZ e' s' t m lL mcL (fun t' m' lL' mcL' => + exec PreSpill isRegZ e' s' k t m lL mcL (fun k' t' m' lL' mcL' => exists lH' mcH', states_compat lH' corresp' lL' /\ (mcL' - mcL <= mcH' - mcH)%metricsH /\ - post t' m' lH' mcH'). + post k' t' m' lH' mcH'). Proof. induction 2; intros; match goal with @@ -1300,13 +1302,13 @@ 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' : MetricLog) => + (mid2 := (fun k'0 t'0 m'0 lL' mcL' => exists (lH' : srcLocals) (mcH' : MetricLog), states_compat lH' a1 lL' /\ (exists mcHmid mcLmid, mcLmid - mcL <= mcHmid - mc /\ mcL' - mcLmid <= mcH' - mcHmid)%metricsH /\ - mid2 t'0 m'0 lH' mcH')). + mid2 k'0 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. + cbv beta. intros. fwd. eexists. eexists. (* exists (exec.cost_SLoop_false isRegStr cond mcH'). *) @@ -1341,7 +1343,7 @@ Section CheckerCorrect. eapply exec.seq. + eapply IH1. 1: eassumption. eapply states_compat_precond. assumption. - + cbv beta. intros t' m' l' mcblah ?. fwd. + + cbv beta. intros k' t' m' l' mcblah ?. fwd. 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. diff --git a/compiler/src/compiler/RiscvEventLoop.v b/compiler/src/compiler/RiscvEventLoop.v index cd82e5bf3..41550dd11 100644 --- a/compiler/src/compiler/RiscvEventLoop.v +++ b/compiler/src/compiler/RiscvEventLoop.v @@ -44,7 +44,7 @@ Section EventLoop. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {PR: MetricPrimitives PRParams}. @@ -67,11 +67,12 @@ Section EventLoop. goodReadyState done m -> m.(getPc) = if done then pc_end else pc_start. Hypothesis goodReadyState_preserved_by_jump_back: - forall (state: RiscvMachineL) newMetrics, + forall (state: RiscvMachineL) newMetrics newLeakage, goodReadyState true state -> let state' := (withPc pc_start (withNextPc (word.add pc_start (word.of_Z 4)) - (withMetrics newMetrics state))) in + (withMetrics newMetrics + (withLeakageEvents (Some newLeakage) state)))) in valid_machine state' -> goodReadyState false state'. @@ -135,12 +136,15 @@ Section EventLoop. | |- ?G => let T := type of H in replace G with T; [exact H|] end. repeat f_equal. - all: solve_word_eq word_ok. + all: try solve_word_eq word_ok. + destruct getTrace; simpl. + { instantiate (1 := [_;_]). reflexivity. } + reflexivity. + simpl. match goal with | H: valid_machine ?m1 |- valid_machine ?m2 => replace m2 with m1; [exact H|] end. - f_equal. f_equal; solve_word_eq word_ok. + f_equal. f_equal; try solve_word_eq word_ok. destruct getTrace; reflexivity. - intros state [C1 C2]. apply goodReadyState_checks_PC in C1. apply goodReadyState_checks_PC in C2. diff --git a/compiler/src/compiler/RunInstruction.v b/compiler/src/compiler/RunInstruction.v index ed2def57c..afc7b1be9 100644 --- a/compiler/src/compiler/RunInstruction.v +++ b/compiler/src/compiler/RunInstruction.v @@ -31,6 +31,8 @@ Require Export coqutil.Word.SimplWordExpr. Require Import coqutil.Tactics.Simp. Require Import compiler.DivisibleBy4. Require Import compiler.ZLemmas. +Require Import riscv.Spec.LeakageOfInstr. +Require Import coqutil.Datatypes.Option. Import Utility. Notation Register0 := 0%Z (only parsing). @@ -104,7 +106,7 @@ Section Run. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {PR: MetricPrimitives PRParams}. @@ -124,6 +126,13 @@ Section Run. Context (iset: InstructionSet). + Notation w0 := (word.of_Z 0). + Definition final_trace (r1 r2 : Z) (v1 v2 : word) (i : Instruction) (pc : word) + (initial_trace : option (list LeakageEvent)) := + option_map2 cons + (concrete_leakage_of_instr (fun r => if r =? r1 then v1 else v2) i) + (option_map (cons (fetchInstr pc)) initial_trace). + Definition run_Jalr0_spec := forall (rs1: Z) (oimm12: Z) (initialL: RiscvMachineL) (Exec R Rexec: mem -> Prop) (dest: word), @@ -148,6 +157,7 @@ Section Run. finalL.(getPc) = word.add dest (word.of_Z oimm12) /\ finalL.(getNextPc) = word.add finalL.(getPc) (word.of_Z 4) /\ finalL.(getMetrics) = addMetricInstructions 1 (addMetricJumps 1 (addMetricLoads 1 initialL.(getMetrics))) /\ + finalL.(getTrace) = final_trace rs1 0 dest w0 (Jalr RegisterNames.zero rs1 oimm12) initialL.(getPc) initialL.(getTrace) /\ valid_machine finalL). Definition run_Jal_spec := @@ -167,6 +177,7 @@ Section Run. finalL.(getPc) = word.add initialL.(getPc) (word.of_Z jimm20) /\ finalL.(getNextPc) = word.add finalL.(getPc) (word.of_Z 4) /\ finalL.(getMetrics) = addMetricInstructions 1 (addMetricJumps 1 (addMetricLoads 1 initialL.(getMetrics))) /\ + finalL.(getTrace) = final_trace 0 0 w0 w0 (Jal rd jimm20) initialL.(getPc) initialL.(getTrace) /\ valid_machine finalL). Definition run_Jal0_spec := @@ -184,6 +195,7 @@ Section Run. finalL.(getPc) = word.add initialL.(getPc) (word.of_Z jimm20) /\ finalL.(getNextPc) = word.add finalL.(getPc) (word.of_Z 4) /\ finalL.(getMetrics) = addMetricInstructions 1 (addMetricJumps 1 (addMetricLoads 1 initialL.(getMetrics))) /\ + finalL.(getTrace) = final_trace 0 0 w0 w0 (Jal Register0 jimm20) initialL.(getPc) initialL.(getTrace) /\ valid_machine finalL). Definition run_ImmReg_spec(Op: Z -> Z -> Z -> Instruction) @@ -205,6 +217,7 @@ Section Run. finalL.(getPc) = initialL.(getNextPc) /\ finalL.(getNextPc) = word.add finalL.(getPc) (word.of_Z 4) /\ finalL.(getMetrics) = addMetricInstructions 1 (addMetricLoads 1 initialL.(getMetrics)) /\ + finalL.(getTrace) = final_trace rs 0 rs_val w0 (Op rd rs imm) initialL.(getPc) initialL.(getTrace) /\ valid_machine finalL). Definition run_Load_spec(n: nat)(L: Z -> Z -> Z -> Instruction) @@ -230,6 +243,7 @@ Section Run. finalL.(getPc) = initialL.(getNextPc) /\ finalL.(getNextPc) = word.add finalL.(getPc) (word.of_Z 4) /\ finalL.(getMetrics) = addMetricInstructions 1 (addMetricLoads 2 initialL.(getMetrics)) /\ + finalL.(getTrace) = final_trace rs 0 base w0 (L rd rs ofs) initialL.(getPc) initialL.(getTrace) /\ valid_machine finalL). Definition run_Store_spec(n: nat)(S: Z -> Z -> Z -> Instruction): Prop := @@ -255,6 +269,7 @@ Section Run. finalL.(getPc) = initialL.(getNextPc) /\ finalL.(getNextPc) = word.add finalL.(getPc) (word.of_Z 4) /\ finalL.(getMetrics) = addMetricInstructions 1 (addMetricStores 1 (addMetricLoads 1 initialL.(getMetrics))) /\ + finalL.(getTrace) = final_trace rs1 rs2 base v_new (S rs1 rs2 ofs) initialL.(getPc) initialL.(getTrace) /\ valid_machine finalL). Ltac inline_iff1 := @@ -268,13 +283,15 @@ Section Run. end; simpl in *; subst; simulate'; + cbv [final_trace]; simpl; + repeat rewrite Z.eqb_refl; repeat match goal with - | |- _ /\ _ => split - | |- _ => solve [auto] - | |- _ => ecancel_assumption - | |- _ => solve_MetricLog - end. + | |- _ /\ _ => split + | |- _ => solve [auto] + | |- _ => ecancel_assumption + | |- _ => solve_MetricLog + end. Ltac t := repeat intro; inline_iff1; get_run1_valid_for_free; t0. diff --git a/compiler/src/compiler/Spilling.v b/compiler/src/compiler/Spilling.v index 14c656809..0ea499735 100644 --- a/compiler/src/compiler/Spilling.v +++ b/compiler/src/compiler/Spilling.v @@ -1,12 +1,15 @@ Require Import compiler.util.Common. +Require Import bedrock2.LeakageSemantics. Require Import bedrock2.Map.SeparationLogic. Require Import compiler.FlatImp. +Require Import Coq.Init.Wf Wellfounded. Require Import Coq.Lists.List. Import ListNotations. Require Import Coq.Logic.PropExtensionality. Require Import Coq.Logic.FunctionalExtensionality. Require Import riscv.Utility.Utility. Require Import coqutil.Map.MapEauto. Require Import coqutil.Tactics.Simp. +Require Import compiler.FixEq. Require Import compiler.Registers. Require Import compiler.SeparationLogic. Require Import compiler.SpillingMapGoals. @@ -21,8 +24,8 @@ Open Scope Z_scope. Section Spilling. Notation stmt := (stmt Z). - Notation execpre := (exec PreSpill isRegZ). - Notation execpost := (exec PostSpill isRegZ). + Notation execpre pick_sp e := (@exec _ _ _ _ _ _ _ _ PreSpill isRegZ e pick_sp). + Notation execpost pick_sp e := (@exec _ _ _ _ _ _ _ _ PostSpill isRegZ e pick_sp). Definition zero := 0. Definition ra := 1. @@ -59,12 +62,24 @@ Section Spilling. | None => SSkip end. + Definition leak_load_iarg_reg(fpval: word) (r: Z): leakage := + match stack_loc r with + | Some o => [leak_word (word.add fpval (word.of_Z o))] + | None => nil + end. + Definition save_ires_reg(r: Z): stmt := match stack_loc r with | Some o => SStore Syntax.access_size.word fp (spill_tmp 1) o | None => SSkip end. + Definition leak_save_ires_reg(fpval: word) (r: Z) : leakage := + match stack_loc r with + | Some o => [leak_word (word.add fpval (word.of_Z o))] + | None => nil + end. + Notation "s1 ;; s2" := (SSeq s1 s2) (right associativity, at level 60). (* reg must be <32, var might be >= 32 *) @@ -74,11 +89,23 @@ Section Spilling. | None => SSet reg var end. + Definition leak_set_reg_to_var(fpval: word) (var: Z) : leakage := + match stack_loc var with + | Some o => [leak_word (word.add fpval (word.of_Z o))] + | None => nil + end. + Fixpoint set_reg_range_to_vars(range_start: Z)(srcs: list Z): stmt := match srcs with | nil => SSkip | x :: xs => set_reg_range_to_vars (range_start+1) xs;; set_reg_to_var range_start x end. + + Fixpoint leak_set_reg_range_to_vars(fpval: word)(srcs: list Z) : leakage := + match srcs with + | nil => nil + | x :: xs => leak_set_reg_range_to_vars fpval xs ++ leak_set_reg_to_var fpval x + end. (* var might be >=32, reg must be < 32 *) Definition set_var_to_reg(var reg: Z): stmt := @@ -87,6 +114,12 @@ Section Spilling. | None => SSet var reg end. + Definition leak_set_var_to_reg(fpval: word) (var: Z) : leakage := + match stack_loc var with + | Some o => [leak_word (word.add fpval (word.of_Z o))] + | None => nil + end. + Fixpoint set_vars_to_reg_range(dests: list Z)(range_start: Z): stmt := match dests with | nil => SSkip @@ -100,18 +133,32 @@ Section Spilling. (do_first;; set_var_to_reg x range_start) xs (range_start+1) end. + Fixpoint leak_set_vars_to_reg_range(fpval: word) (dests: list Z) : leakage := + match dests with + | nil => nil + | x :: xs => leak_set_var_to_reg fpval x ++ leak_set_vars_to_reg_range fpval xs + end. + Definition prepare_bcond(c: bcond Z): stmt := match c with | CondBinary _ x y => load_iarg_reg 1 x;; load_iarg_reg 2 y | CondNez x => load_iarg_reg 1 x end. + Definition leak_prepare_bcond(fpval: word) (c: bcond Z) : leakage := + match c with + | CondBinary _ x y => leak_load_iarg_reg fpval x ++ (leak_load_iarg_reg fpval y) + | CondNez x => leak_load_iarg_reg fpval x + end. + Definition spill_bcond(c: bcond Z): bcond Z := match c with | CondBinary op x y => CondBinary op (iarg_reg 1 x) (iarg_reg 2 y) | CondNez x => CondNez (iarg_reg 1 x) end. + Definition leak_spill_bcond : leakage := nil. + Fixpoint spill_stmt(s: stmt): stmt := match s with | SLoad sz x y o => @@ -161,6 +208,237 @@ Section Spilling. set_vars_to_reg_range resvars a0 end. + Definition tuple : Type := stmt * leakage * leakage * word * (leakage -> leakage -> leakage * word). + + Definition project_tuple (tup : tuple) : nat * stmt := + let '(s, k, sk_so_far, fpval, f) := tup in (length k, s). + Definition lt_tuple (x y : tuple) := + lt_tuple' (project_tuple x) (project_tuple y). + + Lemma lt_tuple_wf : well_founded lt_tuple. + Proof. + cbv [lt_tuple]. apply wf_inverse_image. apply lt_tuple'_wf. + Defined. + + Definition stmt_leakage_body + {env: map.map String.string (list Z * list Z * stmt)} + (e: env) + (pick_sp : leakage -> word) + (tup : stmt * leakage * leakage * word * (leakage (*skip*) -> leakage (*sk_so_far*) -> leakage * word)) + (stmt_leakage : forall othertup, lt_tuple othertup tup -> leakage * word) + : leakage * word. + refine ( + match tup as x return tup = x -> _ with + | (s, k, sk_so_far, fpval, f) => + fun _ => + match s as x return s = x -> _ with + | SLoad sz x y o => + fun _ => + match k with + | leak_word addr :: k' => + f [leak_word addr] (sk_so_far ++ leak_load_iarg_reg fpval y ++ [leak_word addr] ++ leak_save_ires_reg fpval x) + | _ => (nil, word.of_Z 0) + end + | SStore sz x y o => + fun _ => + match k with + | leak_word addr :: k' => + f [leak_word addr] (sk_so_far ++ leak_load_iarg_reg fpval x ++ leak_load_iarg_reg fpval y ++ [leak_word addr]) + | _ => (nil, word.of_Z 0) + end + | SInlinetable _ x _ i => + fun _ => + match k with + | leak_word i' :: k' => + f [leak_word i'] (sk_so_far ++ leak_load_iarg_reg fpval i ++ [leak_word i'] ++ leak_save_ires_reg fpval x) + | _ => (nil, word.of_Z 0) + end + | SStackalloc x z body => + fun _ => + match k as x return k = x -> _ with + | _ :: k' => + fun _ => + stmt_leakage (body, k', sk_so_far ++ leak_unit :: leak_save_ires_reg fpval x, fpval, fun skip => f (leak_unit :: skip)) _ + | nil => + fun _ => + (nil, pick_sp (rev sk_so_far)) + end eq_refl + | SLit x _ => + fun _ => + f [] (sk_so_far ++ leak_save_ires_reg fpval x) + | SOp x op y oz => + fun _ => + let newt_a' := + match op with + | Syntax.bopname.divu + | Syntax.bopname.remu => + match k with + | leak_word x1 :: leak_word x2 :: k' => Some ([leak_word x1; leak_word x2], k') + | _ => None + end + | Syntax.bopname.slu + | Syntax.bopname.sru + | Syntax.bopname.srs => + match k with + | leak_word x2 :: k' => Some ([leak_word x2], k') + | _ => None + end + | _ => Some ([], k) + end + in + match newt_a' with + | Some (newt, a') => + f newt + (sk_so_far ++ + leak_load_iarg_reg fpval y ++ + match oz with + | Var z => leak_load_iarg_reg fpval z + | Const _ => [] + end + ++ newt + ++ leak_save_ires_reg fpval x) + | None => (nil, word.of_Z 0) + end + | SSet x y => + fun _ => + f [] (sk_so_far ++ leak_load_iarg_reg fpval y ++ leak_save_ires_reg fpval x) + | SIf c thn els => + fun _ => + match k as x return k = x -> _ with + | leak_bool b :: k' => + fun _ => + stmt_leakage (if b then thn else els, + k', + sk_so_far ++ leak_prepare_bcond fpval c ++ leak_spill_bcond ++ [leak_bool b], + fpval, + (fun skip => f (leak_bool b :: skip))) _ + | _ => fun _ => (nil, word.of_Z 0) + end eq_refl + | SLoop s1 c s2 => + fun _ => + stmt_leakage (s1, k, sk_so_far, fpval, + (fun skip sk_so_far' => + Let_In_pf_nd (List.skipn (length skip) k) + (fun k' _ => + match k' as x return k' = x -> _ with + | leak_bool true :: k'' => + fun _ => + stmt_leakage (s2, k'', sk_so_far' ++ leak_prepare_bcond fpval c ++ leak_spill_bcond ++ [leak_bool true], fpval, + (fun skip' sk_so_far'' => + let k''' := List.skipn (length skip') k'' in + stmt_leakage (s, k''', sk_so_far'', fpval, + (fun skip'' => f (skip ++ leak_bool true :: skip' ++ skip''))) _)) _ + | leak_bool false :: k'' => + fun _ => + f (skip ++ [leak_bool false]) (sk_so_far' ++ leak_prepare_bcond fpval c ++ leak_spill_bcond ++ [leak_bool false]) + | _ => fun _ => (nil, word.of_Z 0) + end eq_refl))) _ + | SSeq s1 s2 => + fun _ => + stmt_leakage (s1, k, sk_so_far, fpval, + (fun skip sk_so_far' => + let k' := List.skipn (length skip) k in + stmt_leakage (s2, k', sk_so_far', fpval, (fun skip' => f (skip ++ skip'))) _)) _ + | SSkip => fun _ => f [] sk_so_far + | SCall resvars fname argvars => + fun _ => + match k as x return k = x -> _ with + | leak_unit :: k' => + fun _ => + match @map.get _ _ env e fname with + | Some (params, rets, fbody) => + let sk_before_salloc := sk_so_far ++ leak_set_reg_range_to_vars fpval argvars ++ [leak_unit] in + let fpval' := pick_sp (rev sk_before_salloc) in + stmt_leakage (fbody, + k', + sk_before_salloc ++ leak_unit :: leak_set_vars_to_reg_range fpval' params, + fpval', + (fun skip sk_so_far' => + let k'' := List.skipn (length skip) k' in + f (leak_unit :: skip) (sk_so_far' ++ leak_set_reg_range_to_vars fpval' rets ++ leak_set_vars_to_reg_range fpval resvars))) _ + | None => (nil, word.of_Z 0) + end + | _ => fun _ => (nil, word.of_Z 0) + end eq_refl + | SInteract resvars _ argvars => + fun _ => + match k with + | leak_list l :: k' => + f [leak_list l] (sk_so_far ++ leak_set_reg_range_to_vars fpval argvars ++ [leak_list l] ++ leak_set_vars_to_reg_range fpval resvars) + | _ => (nil, word.of_Z 0) + end + end eq_refl + end%nat eq_refl). + Proof. + Unshelve. + all: intros. + all: cbv [lt_tuple project_tuple]. + all: subst. + all: repeat match goal with + | t := List.skipn ?n ?k |- _ => + let H := fresh "H" in + assert (H := List.skipn_length n k); subst t end. + all: try (right; constructor; constructor). + all: try (left; constructor; constructor). + - assert (H0 := skipn_length (length skip) k). left. + rewrite H. assert (H1:= @f_equal _ _ (@length _) _ _ e3). + simpl in H1. blia. + - assert (H := skipn_length (length skip) k). left. + assert (H1 := @f_equal _ _ (@length _) _ _ e3). simpl in H1. blia. + - destruct (length (List.skipn (length skip) k) =? length k)%nat eqn:E. + + apply Nat.eqb_eq in E. rewrite E. right. constructor. constructor. + + apply Nat.eqb_neq in E. left. blia. + Defined. + + Definition stmt_leakage + {env: map.map String.string (list Z * list Z * stmt)} e pick_sp + := Fix lt_tuple_wf _ (stmt_leakage_body e pick_sp). + + Definition Equiv (x y : tuple) := + let '(x1, x2, x3, x4, fx) := x in + let '(y1, y2, y3, y4, fy) := y in + (x1, x2, x3, x4) = (y1, y2, y3, y4) /\ + forall k sk, + fx k sk = fy k sk. + + Lemma stmt_leakage_body_ext {env: map.map String.string (list Z * list Z * stmt)} e pick_sp : + forall (x1 x2 : tuple) + (f1 : forall y : tuple, lt_tuple y x1 -> leakage * word) + (f2 : forall y : tuple, lt_tuple y x2 -> leakage * word), + Equiv x1 x2 -> + (forall (y1 y2 : tuple) (p1 : lt_tuple y1 x1) (p2 : lt_tuple y2 x2), + Equiv y1 y2 -> f1 y1 p1 = f2 y2 p2) -> + stmt_leakage_body e pick_sp x1 f1 = + stmt_leakage_body e pick_sp x2 f2. + Proof. + intros. cbv [stmt_leakage_body]. cbv beta. + destruct x1 as [ [ [ [s_1 k_1] sk_so_far_1] fpval_1] f_1]. + destruct x2 as [ [ [ [s_2 k_2] sk_so_far_2] fpval_2] f_2]. + cbv [Equiv] in H. destruct H as [H1 H2]. injection H1. intros. subst. clear H1. + repeat (Tactics.destruct_one_match || rewrite H || apply H0 || cbv [Equiv] || intuition auto || match goal with | |- _ :: _ = _ :: _ => f_equal end || intuition auto(*why does putting this here make this work*)). + apply Let_In_pf_nd_ext. + repeat (Tactics.destruct_one_match || rewrite H || apply H0 || cbv [Equiv] || intuition auto || match goal with | |- _ :: _ = _ :: _ => f_equal end || intuition auto). + Qed. + + Lemma sfix_step {env: map.map String.string (list Z * list Z * stmt)} e pick_sp tup : + stmt_leakage e pick_sp tup = stmt_leakage_body e pick_sp tup (fun y _ => stmt_leakage e pick_sp y). + Proof. + cbv [stmt_leakage]. + apply (@Fix_eq'_nondep _ _ lt_tuple_wf _ (stmt_leakage_body e pick_sp) Equiv eq). + { apply stmt_leakage_body_ext. } + { cbv [Equiv]. destruct tup as [ [ [x1 x2] x3] fx]. intuition. } + Qed. + + Lemma stmt_leakage_ext {env: map.map String.string (list Z * list Z * stmt)} e pick_sp x1 x2 : + Equiv x1 x2 -> + stmt_leakage e pick_sp x1 = stmt_leakage e pick_sp x2. + Proof. + revert x2. induction (lt_tuple_wf x1). intros. do 2 rewrite sfix_step. + apply stmt_leakage_body_ext. + - assumption. + - intros. apply H0; assumption. + Qed. + Definition max_var_bcond(c: bcond Z): Z := match c with | CondBinary _ x y => Z.max x y @@ -277,6 +555,15 @@ Section Spilling. else error:("Spilling got input program with invalid var names (please report as a bug)"). + Definition fun_leakage {env : map.map string (list Z * list Z * stmt)} (e : env) (pick_sp : leakage -> word) (f : list Z * list Z * stmt) (k : leakage) (sk_so_far : leakage) : leakage * word := + let '(argnames, resnames, body) := f in + let fpval := pick_sp (rev sk_so_far) in + stmt_leakage e pick_sp (body, + k, + sk_so_far ++ leak_unit :: leak_set_vars_to_reg_range fpval argnames, + fpval, + (fun skip sk_so_far' => (sk_so_far' ++ leak_set_reg_range_to_vars fpval resnames, word.of_Z 0))). + Lemma firstn_min_absorb_length_r{A: Type}: forall (l: list A) n, List.firstn (Nat.min n (length l)) l = List.firstn n l. Proof. @@ -311,7 +598,7 @@ Section Spilling. Context {locals: map.map Z word}. Context {localsOk: map.ok locals}. Context {env: map.map String.string (list Z * list Z * stmt)} {env_ok: map.ok env}. - Context {ext_spec: Semantics.ExtSpec}. + Context {ext_spec: LeakageSemantics.ExtSpec}. Definition spill_functions: env -> result env := map.try_map_values spill_fun. @@ -406,7 +693,7 @@ Section Spilling. nth_error stackwords (Z.to_nat (r - 32)) = Some v) /\ length stackwords = Z.to_nat (maxvar - 31). - Implicit Types post : Semantics.trace -> mem -> locals -> MetricLog -> Prop. + Implicit Types post : leakage -> Semantics.trace -> mem -> locals -> MetricLog -> Prop. Lemma put_arg_reg: forall l r v fpval lRegs, (eq lRegs * arg_regs * ptsto fp fpval)%sep l -> @@ -454,18 +741,17 @@ Section Spilling. unfold spill_tmp. eapply put_arg_reg; eassumption. Qed. - Lemma load_iarg_reg_correct(i: Z): forall r e2 t1 t2 m1 m2 l1 l2 mc2 fpval post frame maxvar v, + Lemma load_iarg_reg_correct {pick_sp: PickSp} (i: Z): forall r e2 k2 t1 t2 m1 m2 l1 l2 mc2 fpval post frame maxvar v, 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 -> (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 (mkMetricLog 1 0 2 0 + mc2)%metricsH)) -> - execpost e2 (load_iarg_reg i r) t2 m2 l2 mc2 post. + post (rev (leak_load_iarg_reg fpval r) ++ k2) t2 m2 (map.put l2 (iarg_reg i r) v) (if isRegZ r then mc2 else (mkMetricLog 1 0 2 0 + mc2)%metricsH)) -> + execpost pick_sp e2 (load_iarg_reg i r) k2 t2 m2 l2 mc2 post. Proof. intros. - unfold load_iarg_reg, stack_loc, iarg_reg, related in *. fwd. + unfold leak_load_iarg_reg, 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). @@ -505,16 +791,16 @@ 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, *) + + (* Lemma load_iarg_reg_correct' {pick_sp: PickSp} (i: Z): forall r e2 k1 k2 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'). *) + (* post k1 t1 m1 l1 mc1 -> *) + (* execpost e2 (load_iarg_reg i r) k2 t2 m2 l2 mc2 *) + (* (fun k2' t2' m2' l2' mc2' => exists k1' t1' m1' l1' mc1', *) + (* related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ post k1' t1' m1' l1' mc1'). *) (* Proof. *) (* intros. *) (* unfold load_iarg_reg, stack_loc, iarg_reg, related in *. fwd. *) @@ -556,18 +842,19 @@ Section Spilling. when the new postcondition is used as a "mid1" in exec.loop, and body1 is a seq in which this lemma was used, t2, m2, l2, mc2 are introduced after the evar "?mid1" is created (i.e. after exec.loop is applied), so they are not in the scope of "?mid1". *) - Lemma load_iarg_reg_correct''(i: Z): forall r e2 t1 t2 m1 m2 l1 l2 mc2 frame maxvar v fpval, + Lemma load_iarg_reg_correct'' {pick_sp: PickSp} (i: Z): forall r e2 k2 t1 t2 m1 m2 l1 l2 mc2 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 -> - execpost e2 (load_iarg_reg i r) t2 m2 l2 mc2 (fun t2' m2' l2' mc2' => + execpost pick_sp e2 (load_iarg_reg i r) k2 t2 m2 l2 mc2 (fun k2' t2' m2' l2' mc2' => + k2' = rev (leak_load_iarg_reg fpval r) ++ k2 /\ 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 mkMetricLog 1 0 2 0 + mc2))%metricsH). Proof. intros. - unfold load_iarg_reg, stack_loc, iarg_reg, related in *. fwd. + unfold load_iarg_reg, leak_load_iarg_reg, stack_loc, iarg_reg, related in *. fwd. destr (32 <=? r). - eapply exec.load. + eapply get_sep. ecancel_assumption. @@ -608,19 +895,20 @@ Section Spilling. `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 mc1' mc2 mc2' x v maxvar frame post fpval, - post t1 m1 (map.put l1 x v) mc1' -> + Lemma save_ires_reg_correct {pick_sp: PickSp} : forall e k1 k2 t1 t2 m1 m2 l1 l2 mc1 mc1' mc2 mc2' x v maxvar frame post fpval, + post k1 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 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'', + execpost pick_sp e (save_ires_reg x) k2 t2 m2 (map.put l2 (ires_reg x) v) mc2' + (fun k2' t2' m2' l2' mc2'' => exists k1' t1' m1' l1' mc1'', related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ - post t1' m1' l1' mc1'' /\ + post k1' t1' m1' l1' mc1'' /\ + k2' = rev (leak_save_ires_reg fpval x) ++ k2 /\ (mc2'' - mc2 <= mc1'' - mc1)%metricsH). Proof. intros. - unfold save_ires_reg, stack_loc, ires_reg, related in *. fwd. + unfold leak_save_ires_reg, save_ires_reg, stack_loc, ires_reg, related in *. fwd. destr (32 <=? x). - edestruct store_to_word_array as (m' & stackwords' & St & S' & Ni & Nj & L). 1: ecancel_assumption. @@ -662,7 +950,7 @@ Section Spilling. - eapply Nj. 1: blia. eauto. } 1: { unfold spill_tmp. eapply put_tmp; eauto. } - 1: blia. + 1: blia. 1: reflexivity. unfold cost_store. unfold spill_tmp; cbn. destr (isRegZ x); solve_MetricLog. } @@ -704,24 +992,25 @@ Section Spilling. repeat destruct_one_match_hyp; subst; fwd; try congruence; try blia. specialize H0p8 with (1 := H1). blia. } - all: try eassumption. + all: try eassumption. 1: reflexivity. 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. *) - Lemma save_ires_reg_correct'': forall e t1 t2 m1 m2 l1 l2 mc2 x v maxvar frame post fpval, + Lemma save_ires_reg_correct'' {pick_sp: PickSp} : forall e k2 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', related maxvar frame fpval t1 m1 (map.put l1 x v) t2' m2' l2' -> - 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. + post (rev (leak_save_ires_reg fpval x) ++ k2) t2' m2' l2' (if isRegZ x then mc2 else mkMetricLog 1 1 1 0 + mc2)%metricsH) -> + execpost pick_sp e (save_ires_reg x) k2 t2 m2 (map.put l2 (ires_reg x) v) mc2 post. Proof. intros. - unfold save_ires_reg, stack_loc, ires_reg, related in *. fwd. + unfold leak_save_ires_reg, save_ires_reg, stack_loc, ires_reg, related in *. fwd. destr (32 <=? x). - edestruct store_to_word_array as (m' & stackwords' & St & S' & Ni & Nj & L). 1: ecancel_assumption. @@ -873,8 +1162,8 @@ Section Spilling. do 2 rewrite MetricArith.add_assoc; rewrite (MetricArith.add_comm n); reflexivity. 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, + Lemma set_vars_to_reg_range_correct {pick_sp: PickSp} : + forall args start argvs e k2 t1 t2 m1 m2 l1 l1' l2 mc2 maxvar frame post fpval, related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> map.putmany_of_list_zip args argvs l1 = Some l1' -> map.getmany_of_list l2 (List.unfoldn (Z.add 1) (List.length args) start) = Some argvs -> @@ -882,11 +1171,12 @@ Section Spilling. a0 <= start -> start + Z.of_nat (List.length args) <= a7 + 1 -> Forall (fun x => fp < x <= maxvar /\ (x < a0 \/ a7 < x)) args -> - (forall m2' l2' mc2', + (forall k2' 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') -> - execpost e (set_vars_to_reg_range args start) t2 m2 l2 mc2 post. + mc2' = cost_set_vars_to_reg_range args start mc2 -> + k2' = rev (leak_set_vars_to_reg_range fpval args) ++ k2 -> + post k2' t2 m2' l2' mc2') -> + execpost pick_sp e (set_vars_to_reg_range args start) k2 t2 m2 l2 mc2 post. Proof. induction args; intros. - simpl. eapply exec.skip. fwd. eauto. @@ -894,6 +1184,7 @@ Section Spilling. unfold related in H. fwd. eapply exec.seq_cps. rewrite (Z.add_comm 1 start) in *. + simpl in H6. cbv [leak_set_var_to_reg stack_loc] in H6. destr (32 <=? a). + edestruct store_to_word_array with (i := a - 32). 1: ecancel_assumption. 1: blia. @@ -923,11 +1214,12 @@ Section Spilling. assumption. } { blia. } * intros. apply H6; auto. - cbn in *. destr (isRegZ start); try blia; destr (isRegZ a); try blia. - 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. + -- cbn in *. destr (isRegZ start); try blia; destr (isRegZ a); try blia. + 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. + -- subst. rewrite rev_app_distr. rewrite <- app_assoc. reflexivity. + eapply exec.set. { eassumption. } eapply IHargs; try eassumption; try blia. 2: { @@ -968,26 +1260,25 @@ Section Spilling. 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, + Lemma set_reg_range_to_vars_correct {pick_sp: PickSp} : + forall args argvs start e k2 t1 t2 m1 m2 l1 l2 mc2 maxvar frame post fpval, related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> (List.length args <= 8)%nat -> a0 <= start -> start + Z.of_nat (List.length args) <= a7 + 1 -> Forall (fun x => fp < x <= maxvar /\ (x < a0 \/ a7 < x)) args -> map.getmany_of_list l1 args = Some argvs -> - (forall l2' mc2', + (forall k2' 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') -> - execpost e (set_reg_range_to_vars start args) t2 m2 l2 mc2 post. + mc2' = cost_set_reg_range_to_vars start args mc2 -> + k2' = rev (leak_set_reg_range_to_vars fpval args) ++ k2 -> + post k2' t2 m2 l2' mc2') -> + execpost pick_sp e (set_reg_range_to_vars start args) k2 t2 m2 l2 mc2 post. Proof. induction args; intros. - - simpl. eapply exec.skip. eapply H5. 1: eassumption. - + simpl. destruct argvs. 1: reflexivity. discriminate. - + trivial. - - simpl. unfold set_reg_to_var, stack_loc. + - simpl. eapply exec.skip. eapply H5; eauto. + - simpl. unfold set_reg_to_var, leak_set_reg_to_var, stack_loc. destruct argvs as [|v vs]. { unfold map.getmany_of_list in H4. cbn in H4. simp. destr (List.option_all (map (map.get l1) args)); discriminate. @@ -995,6 +1286,7 @@ Section Spilling. eapply map.invert_getmany_of_list_cons in H4. destruct H4 as [G GM]. cbn [List.length] in *. simp. + cbn [leak_set_reg_range_to_vars] in H5. cbv [leak_set_reg_to_var stack_loc] in H5. destr (32 <=? a). + eapply exec.seq_cps. eapply IHargs; try eassumption; try blia. @@ -1025,6 +1317,7 @@ Section Spilling. rewrite H6; unfold cost_load, isRegZ; cbn. rewrite (proj2 (Z.leb_le start 31)) by assumption. reflexivity. + -- subst. rewrite rev_app_distr. rewrite <- app_assoc. reflexivity. + eapply exec.seq_cps. eapply IHargs; try eassumption; try blia. intros. @@ -1056,6 +1349,7 @@ Section Spilling. rewrite (proj2 (Z.leb_le a 31)) by assumption. rewrite (proj2 (Z.leb_le start 31)) by assumption. reflexivity. + -- subst. rewrite app_nil_r. reflexivity. Qed. Lemma cost_set_reg_range_to_vars_bound : forall args start mc len, @@ -1273,39 +1567,41 @@ Section Spilling. Qed. 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), - execpre e1 s1 t1 m1 l1 mc1 post -> + forall pick_sp1 k1 t1 m1 l1 mc1 post, + execpre pick_sp1 e1 s1 k1 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), + forall pick_sp2 k2 t2 m2 l2 mc2 fpval f, related maxvar frame fpval t1 m1 l1 t2 m2 l2 -> - execpost e2 (spill_stmt s1) t2 m2 l2 mc2 - (fun (t2' : Semantics.trace) (m2' : mem) (l2' : locals) (mc2' : MetricLog) => - exists t1' m1' l1' mc1', + (forall k, pick_sp1 (k ++ k1) = snd (stmt_leakage e1 pick_sp2 (s1, rev k, rev k2, fpval, f k))) -> + execpost pick_sp2 e2 (spill_stmt s1) k2 t2 m2 l2 mc2 + (fun k2' t2' m2' l2' mc2' => + exists k1' t1' m1' l1' mc1' k1'', related maxvar frame fpval t1' m1' l1' t2' m2' l2' /\ - post t1' m1' l1' mc1' /\ - (mc2' - mc2 <= mc1' - mc1)%metricsH). + post k1' t1' m1' l1' mc1' /\ + (mc2' - mc2 <= mc1' - mc1)%metricsH /\ + k1' = k1'' ++ k1 /\ + forall k f, stmt_leakage e1 pick_sp2 (s1, rev k1'' ++ k, rev k2, fpval, f) = f (rev k1'') (rev k2')). (* 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 := + pick_sp k t m argvals mc + (post : leakage -> 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_spill_spec mc) (fun t' m' l' mc' => + execpre pick_sp e fbody k t m l (cost_spill_spec mc) (fun k' t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ - post t' m' retvals mc'). + post k' 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 := + pick_sp k t m argvals mc + (post : leakage -> 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' => + execpost pick_sp e fbody k t m l mc (fun k' t' m' l' mc' => exists retvals, map.getmany_of_list l' retnames = Some retvals /\ - post t' m' retvals mc'). + post k' t' m' retvals mc'). (* In exec.call, there are many maps of locals involved: @@ -1334,13 +1630,18 @@ Section Spilling. 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, + Lemma spill_fun_correct_aux: forall pick_sp2 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 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'). + forall argvals kH kL t m mcH mcL (post : leakage -> Semantics.trace -> mem -> list word -> MetricLog -> Prop), + call_spec e1 (argnames1, retnames1, body1) + (fun k' => snd (fun_leakage e1 pick_sp2 (argnames1, retnames1, body1) (skipn (length kH) (rev k')) (rev kL))) kH t m argvals mcH post -> + call_spec_spilled e2 (argnames2, retnames2, body2) pick_sp2 kL t m argvals mcL + (fun kL' t' m' l' mcL' => + exists kH'' mcH', + post (kH'' ++ kH) t' m' l' mcH' /\ + metricsLeq (mcL' - mcL) (mcH' - mcH) /\ + fst (fun_leakage e1 pick_sp2 (argnames1, retnames1, body1) (rev kH'') (rev kL)) = rev kL'). Proof. unfold call_spec, spilling_correct_for. intros * Sp IHexec * Ex lFL3 OL2. unfold spill_fun in Sp. fwd. @@ -1411,24 +1712,27 @@ Section Spilling. 2: eapply Forall_le_max. cbv beta. subst maxvar'. clear. blia. } - intros mL4 lFL4 mcL4 R Hcost. + intros kL4 mL4 lFL4 mcL4 R Hcost HkL4. eapply exec.seq_cps. eapply exec.weaken. { - eapply IHexec. 1: apply Ex. 2: exact R. - unfold valid_vars_src. - eapply Forall_vars_stmt_impl. - 2: eapply max_var_sound. - 2: eapply forallb_vars_stmt_correct. - 3: eassumption. - 2: { - unfold is_valid_src_var. - intros *. - rewrite ?Bool.andb_true_iff, ?Bool.orb_true_iff, ?Z.ltb_lt. reflexivity. - } - cbv beta. subst maxvar'. blia. - } - cbv beta. intros tL5 mL5 lFL5 mcL5 (tH5 & mH5 & lFH5 & mcH5 & R5 & OC). - fwd. + eapply IHexec with (f := fun _ => _). 1: apply Ex. 2: exact R. + { unfold valid_vars_src. + eapply Forall_vars_stmt_impl. + 2: eapply max_var_sound. + 2: eapply forallb_vars_stmt_correct. + 3: eassumption. + 2: { + unfold is_valid_src_var. + intros *. + rewrite ?Bool.andb_true_iff, ?Bool.orb_true_iff, ?Z.ltb_lt. reflexivity. + } + cbv beta. subst maxvar'. blia. } + intros. subst a kL4. simpl. simpl_rev. repeat rewrite <- app_assoc. + rewrite List.skipn_app_r. + 2: { rewrite rev_length. reflexivity. } + reflexivity. } + cbv beta. intros kL5 tL5 mL5 lFL5 mcL5 (kH5 & tH5 & mH5 & lFH5 & mcH5 & kH5'' & R5 & OC & Ek & Emc & CT). + subst. fwd. eapply set_reg_range_to_vars_correct. { eassumption. } { blia. } @@ -1447,7 +1751,7 @@ Section Spilling. subst maxvar'. clear. blia. } { eassumption. } rename R into R0. - intros lFL6 mcL6 R GM HCost. + intros kFL6 lFL6 mcL6 R GM HCost HkFL6. (* 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. @@ -1461,7 +1765,7 @@ Section Spilling. | |- exists _, _ => eexists | |- _ /\ _ => split end. - 5: eassumption. + 4: eassumption. 2: { unfold map.split. eauto. } @@ -1477,11 +1781,11 @@ Section Spilling. } blia. } { eassumption. } - { - add_bounds. + { add_bounds. unfold cost_stackalloc, cost_spill_spec in *. (* TODO XXX *) - destruct (isRegZ fp); solve_MetricLog. - } + destruct (isRegZ fp); solve_MetricLog. } + subst a. simpl. simpl_rev. repeat rewrite <- app_assoc in * || simpl in *. + specialize (CT nil). rewrite app_nil_r in CT. rewrite CT. reflexivity. Qed. @@ -1509,31 +1813,19 @@ Section Spilling. Ltac irs := cost_unfold; repeat isReg_helper; cost_solve. Ltac sirs := scost_unfold; repeat isReg_helper; scost_solve. + Ltac after_save_ires_reg_correct'' := + intros; do 6 eexists; ssplit; [eassumption | eassumption | solve [irs] || solve [sirs] || idtac | align_trace | intros; rewrite sfix_step; simpl; simpl_rev; repeat rewrite <- app_assoc; try reflexivity]. + - Lemma spilling_correct (e1 e2 : env) (Ev : spill_functions e1 = Success e2) - (s1 : stmt) - (t1 : Semantics.trace) - (m1 : mem) - (l1 : locals) - (mc1 : MetricLog) - (post : Semantics.trace -> mem -> locals -> MetricLog -> Prop): - 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 -> - 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' /\ - post t1' m1' l1' mc1' /\ - (mc2' - mc2 <= mc1' - mc1)%metricsH). + Lemma spilling_correct (e1 e2 : env) (Ev : spill_functions e1 = Success e2) (s1 : stmt) : + spilling_correct_for e1 e2 s1. Proof. + cbv [spilling_correct_for]. 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 CSet. clear l2 H4. + intros *. intros R GM ? ?. subst. clear l2 H4. unfold related in R. fwd. spec (subst_split (ok := mem_ok) m) as A. 1: eassumption. 1: ecancel_assumption. @@ -1571,18 +1863,18 @@ Section Spilling. { reflexivity. } { unfold a0, a7. blia. } { eassumption. } - { intros. do 4 eexists. split. 1: eassumption. split. - { eapply H2p1. + { intros. do 6 eexists. split. 1: eassumption. ssplit. + - 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. } - cbn in *. subst. - add_bounds. - cost_solve. + eapply map.shrink_disjoint_l; eassumption. + - cbn in *. subst. add_bounds. cost_solve. (* cost_SInteract constraint: prespill - postspill >= (...32...) i think? *) - } + - align_trace. + - intros. subst. rewrite sfix_step. simpl. simpl_rev. + repeat rewrite <- app_assoc. reflexivity. } (* related for set_vars_to_reg_range_correct: *) unfold related. eexists _, _, _. ssplit. @@ -1628,7 +1920,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 kCL2 lCL2 ? ? ? ? ?. subst. 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. } @@ -1697,23 +1989,23 @@ Section Spilling. 2: eapply Forall_le_max. cbv beta. subst maxvar'. clear. blia. } - intros mL4 lFL4 mcL4 R CSet. + intros kL4 mL4 lFL4 mcL4 R ? ?. subst. eapply exec.seq_cps. eapply exec.weaken. { - eapply IHexec. 2: exact R. - unfold valid_vars_src. - eapply Forall_vars_stmt_impl. - 2: eapply max_var_sound. - 2: eapply forallb_vars_stmt_correct. - 3: eassumption. - 2: { - unfold is_valid_src_var. - intros *. - rewrite ?Bool.andb_true_iff, ?Bool.orb_true_iff, ?Z.ltb_lt. reflexivity. - } - cbv beta. subst maxvar'. blia. - } - cbv beta. intros tL5 mL5 lFL5 mcL5 (tH5 & mH5 & lFH5 & mcH5 & R5 & OC & Hmetrics). + eapply IHexec with (f := fun _ => _). 2: exact R. + - unfold valid_vars_src. + eapply Forall_vars_stmt_impl. + 2: eapply max_var_sound. + 2: eapply forallb_vars_stmt_correct. + 3: eassumption. + 2: { unfold is_valid_src_var. + intros *. + rewrite ?Bool.andb_true_iff, ?Bool.orb_true_iff, ?Z.ltb_lt. reflexivity. } + cbv beta. subst maxvar'. blia. + - intros. rewrite associate_one_left. rewrite H6. rewrite sfix_step. + simpl_rev. simpl. rewrite H. simpl_rev. + repeat rewrite <- app_assoc. reflexivity. } + cbv beta. intros kL5 tL5 mL5 lFL5 mcL5 (kH5 & tH5 & mH5 & lFH5 & mcH5 & kH5'' & R5 & OC & Hmetrics & Hk5'' & CT). match goal with | H: context[outcome], A: context[outcome] |- _ => specialize H with (1 := A); move H at bottom; rename H into Q @@ -1737,7 +2029,7 @@ Section Spilling. subst maxvar'. clear. blia. } { eassumption. } rename R into R0. - intros lFL6 mcL6 R GM ?. + intros kFL6 lFL6 mcL6 R GM ? ?. subst. (* 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 *) @@ -1796,14 +2088,13 @@ Section Spilling. { reflexivity. } { unfold a0, a7. blia. } { eassumption. } - { intros m22 l22 mc22 R22. do 4 eexists. split. 1: eassumption. - split; try eassumption. - subst. - move Hmetrics at bottom. - add_bounds. - cost_solve. - (* cost_SCall constraint: prespill - postspill >= (...66...) i think? *) - } + { intros k22 m22 l22 mc22 R22 ? ?. subst. do 6 eexists. ssplit; try eassumption. + - move Hmetrics at bottom. add_bounds. cost_solve. + (* cost_SCall constraint: prespill - postspill >= (...66...) i think? *) + - align_trace. + - intros. rewrite sfix_step. simpl. simpl_rev. rewrite H. Search fbody. + repeat rewrite <- app_assoc in *. simpl in *. rewrite CT. + reflexivity. } - (* exec.load *) eapply exec.seq_cps. @@ -1818,15 +2109,11 @@ Section Spilling. 1: ecancel_assumption. unfold map.split in Sp. subst. fwd. unfold Memory.load, Memory.load_Z, Memory.load_bytes. erewrite map.getmany_of_tuple_in_disjoint_putmany; eauto. } - eapply save_ires_reg_correct. - + eassumption. - + eassumption. - + blia. - + irs. + eapply save_ires_reg_correct''; eauto. after_save_ires_reg_correct''. - (* 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. - pose proof H5 as A. unfold related in A. fwd. + pose proof H6 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. @@ -1836,7 +2123,7 @@ 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: split. 2: eassumption. + do 6 eexists. ssplit. 2: eassumption. + unfold related. repeat match goal with | |- exists _, _ => eexists @@ -1846,6 +2133,9 @@ Section Spilling. spec store_bytes_sep_hi2lo as A. 1: eassumption. all: ecancel_assumption. + irs. + + align_trace. + + intros. rewrite sfix_step. simpl. simpl_rev. repeat rewrite <- app_assoc. + reflexivity. - (* exec.inlinetable *) eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). intros. eapply exec.seq_cps. @@ -1853,21 +2143,23 @@ Section Spilling. { unfold ires_reg, iarg_reg, spill_tmp, fp, a0, a7 in *. destr (32 <=? x); destr (32 <=? i); try blia. } { rewrite map.get_put_same. reflexivity. } { eassumption. } - eapply save_ires_reg_correct. - + eassumption. - + eassumption. - + blia. - + irs. + eapply save_ires_reg_correct''; eauto. after_save_ires_reg_correct''. - (* exec.stackalloc *) rename H1 into IH. eapply exec.stackalloc. 1: assumption. intros. eapply exec.seq_cps. edestruct grow_related_mem as (mCombined1 & ? & ?). 1,2: eassumption. - eapply save_ires_reg_correct''. 1: eassumption. 1: blia. - intros. - eapply exec.weaken. { - eapply IH; eassumption. } + eapply save_ires_reg_correct''; eauto. + assert (Ha : a = pick_sp1 k). + { subst a. specialize (H4 nil). simpl in H4. rewrite H4. + rewrite sfix_step. simpl. simpl_rev. reflexivity. } + rewrite Ha in *. clear Ha a. intros. + eapply exec.weaken. + { eapply IH with (f := fun _ => _); eauto. + intros. rewrite associate_one_left. rewrite H4. + rewrite sfix_step. simpl. simpl_rev. repeat rewrite <- app_assoc. + reflexivity. } cbv beta. intros. fwd. edestruct shrink_related_mem as (mSmall2 & ? & ?). 1,2: eassumption. repeat match goal with @@ -1875,43 +2167,34 @@ Section Spilling. | |- _ /\ _ => split end. 1,4,3,2: eassumption. - irs. + + irs. + + align_trace. + + intros. rewrite sfix_step. simpl. simpl_rev. + repeat rewrite <- app_assoc in *. rewrite H8p4. reflexivity. - (* exec.lit *) eapply exec.seq_cps. eapply exec.lit. - eapply save_ires_reg_correct. - + eassumption. - + eassumption. - + blia. - + irs. + eapply save_ires_reg_correct''; eauto. after_save_ires_reg_correct''. - (* 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. - { - eapply exec.seq_cps. eapply exec.seq_cps. eapply load_iarg_reg_correct; (blia || eassumption || idtac). - clear H2. intros. - eapply exec.op. + { eapply exec.seq_cps. eapply exec.seq_cps. + eapply load_iarg_reg_correct; (blia || eassumption || idtac). + clear H2. intros. 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). - sirs. } - } - { - eapply exec.seq_cps. eapply exec.op. + { eapply save_ires_reg_correct''; eauto. after_save_ires_reg_correct''. + destruct op; reflexivity. } } + { 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). - sirs. } - } + { eapply save_ires_reg_correct''; eauto. after_save_ires_reg_correct''. + destruct op; reflexivity. } } - (* exec.set *) 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. - + irs. + eapply save_ires_reg_correct''; eauto. after_save_ires_reg_correct''. - (* exec.if_true *) unfold prepare_bcond. destr cond; cbn [ForallVars_bcond eval_bcond spill_bcond] in *; fwd. + eapply exec.seq_assoc. @@ -1921,17 +2204,29 @@ Section Spilling. cbn. erewrite get_iarg_reg_1 by eauto with zarith. rewrite map.get_put_same. congruence. } eapply exec.weaken. - * eapply IHexec; eassumption. - * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. - sirs. + * eapply IHexec with (f := fun _ => _); try eassumption. + intros. rewrite associate_one_left, H3, sfix_step. + simpl. simpl_rev. repeat rewrite <- app_assoc. reflexivity. + * cbv beta; intros; fwd. eexists _, t1', m1', l1', mc1', _. + ssplit; try eassumption. + -- sirs. + -- align_trace. + -- intros. rewrite sfix_step. simpl. simpl_rev. + repeat rewrite <- app_assoc in *. rewrite H5p4. reflexivity. + 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 IHexec; eassumption. - * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. - sirs. + * eapply IHexec with (f := fun _ => _); try eassumption. + intros. rewrite associate_one_left, H3, sfix_step. + simpl. simpl_rev. repeat rewrite <- app_assoc. reflexivity. + * cbv beta; intros; fwd. eexists _, t1', m1', l1', mc1', _. + ssplit; try eassumption. + -- sirs. + -- align_trace. + -- intros. rewrite sfix_step. simpl. simpl_rev. + repeat rewrite <- app_assoc in *. rewrite H4p4. reflexivity. - (* exec.if_false *) unfold prepare_bcond. destr cond; cbn [ForallVars_bcond eval_bcond spill_bcond] in *; fwd. + eapply exec.seq_assoc. @@ -1941,28 +2236,40 @@ Section Spilling. cbn. erewrite get_iarg_reg_1 by eauto with zarith. rewrite map.get_put_same. congruence. } eapply exec.weaken. - * eapply IHexec; eassumption. - * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. - sirs. + * eapply IHexec with (f := fun _ => _); try eassumption. + intros. rewrite associate_one_left, H3, sfix_step. + simpl. simpl_rev. repeat rewrite <- app_assoc. reflexivity. + * cbv beta; intros; fwd. eexists _, t1', m1', l1', mc1', _. + ssplit; try eassumption. + -- sirs. + -- align_trace. + -- intros. rewrite sfix_step. simpl. simpl_rev. + repeat rewrite <- app_assoc in *. rewrite H5p4. reflexivity. + 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 IHexec; eassumption. - * cbv beta; intros; fwd. exists t1', m1', l1', mc1'. split. 2: split. all: try eassumption. - sirs. + * eapply IHexec with (f := fun _ => _); try eassumption. + intros. rewrite associate_one_left, H3, sfix_step. + simpl. simpl_rev. repeat rewrite <- app_assoc. reflexivity. + * cbv beta; intros; fwd. eexists _, t1', m1', l1', mc1', _. + ssplit; try eassumption. + -- sirs. + -- align_trace. + -- intros. rewrite sfix_step. simpl. simpl_rev. + repeat rewrite <- app_assoc in *. rewrite H1p6. reflexivity. - (* exec.loop *) rename IHexec into IH1, H3 into IH2, H5 into IH12. eapply exec.loop_cps. eapply exec.seq. - 1: eapply IH1; eassumption. + { eapply IH1; try eassumption. + intros. rewrite H8. rewrite sfix_step. reflexivity. } cbv beta. intros. fwd. unfold prepare_bcond. destr cond; cbn [ForallVars_bcond] in *; fwd. + specialize H0 with (1 := H3p1). cbn in H0. fwd. - eapply exec.seq. { - eapply load_iarg_reg_correct''; (blia || eassumption || idtac). - } + eapply exec.seq. + { eapply load_iarg_reg_correct''; (blia || eassumption || idtac). } cbv beta. intros. fwd. eapply exec.weaken. { eapply load_iarg_reg_correct''; (blia || eassumption || idtac). @@ -1971,19 +2278,37 @@ 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. - -- exact H3p8. + * do 6 eexists. ssplit. + -- exact H3p10. -- eapply H1. 1: eassumption. cbn. rewrite E, E0. congruence. -- sirs. - * eapply exec.weaken. 1: eapply IH2. - -- eassumption. + -- align_trace. + -- intros. rewrite sfix_step. simpl. repeat rewrite <- app_assoc in *. + rewrite H3p4. rewrite List.skipn_app_r by reflexivity. + cbv [Let_In_pf_nd]. simpl. simpl_rev. repeat rewrite <- app_assoc. + reflexivity. + * eapply exec.weaken. 1: eapply IH2; try eassumption. -- cbn. rewrite E, E0. congruence. - -- eassumption. - -- eassumption. + -- intros. rewrite associate_one_left. rewrite app_assoc. rewrite H8. + rewrite sfix_step. simpl. simpl_rev. repeat rewrite <- app_assoc in *. + rewrite H3p4. rewrite List.skipn_app_r by reflexivity. reflexivity. -- 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. - sirs. + ++ eapply IH12 with (f := fun _ => _); try eassumption. + { repeat split; eauto; blia. } + intros. rewrite associate_one_left. repeat rewrite app_assoc. + rewrite H8. rewrite sfix_step. simpl. simpl_rev. rewrite H3p4. + rewrite List.skipn_app_r by reflexivity. cbv [Let_In_pf_nd]. + repeat rewrite <- app_assoc in *. rewrite H5p4. + rewrite List.skipn_app_r by reflexivity. reflexivity. + ++ cbv beta; intros; fwd. eexists _, t1'1, m1'1, l1'1, mc1'1, _. + ssplit; try eassumption. + --- sirs. + --- align_trace. + --- intros. rewrite sfix_step. simpl. simpl_rev. + repeat rewrite <- app_assoc in *. rewrite H3p4. + rewrite List.skipn_app_r by reflexivity. cbv [Let_In_pf_nd]. + simpl. rewrite H5p4. rewrite List.skipn_app_r by reflexivity. + rewrite H5p8. reflexivity. + specialize H0 with (1 := H3p1). cbn in H0. fwd. eapply exec.weaken. { eapply load_iarg_reg_correct''; (blia || eassumption || idtac). @@ -1991,44 +2316,75 @@ 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. 2: split. - -- exact H3p6. + * do 6 eexists. ssplit. + -- exact H3p8. -- eapply H1. 1: eassumption. cbn. rewrite E. congruence. -- sirs. - * eapply exec.weaken. 1: eapply IH2. - -- eassumption. + -- align_trace. + -- intros. rewrite sfix_step. simpl. repeat rewrite <- app_assoc in *. + rewrite H3p4. rewrite List.skipn_app_r by reflexivity. + cbv [Let_In_pf_nd]. simpl. simpl_rev. repeat rewrite <- app_assoc. + reflexivity. + * eapply exec.weaken. 1: eapply IH2; try eassumption. -- cbn. rewrite E. congruence. - -- eassumption. - -- eassumption. - -- 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. - sirs. + -- intros. rewrite associate_one_left. rewrite app_assoc. rewrite H8. + rewrite sfix_step. simpl. simpl_rev. repeat rewrite <- app_assoc in *. + rewrite H3p4. rewrite List.skipn_app_r by reflexivity. reflexivity. + -- cbv beta. intros. fwd. eapply exec.weaken. + ++ eapply IH12 with (f := fun _ => _); try eassumption. + { repeat split; eauto; blia. } + intros. rewrite associate_one_left. repeat rewrite app_assoc. + rewrite H8. rewrite sfix_step. simpl. simpl_rev. rewrite H3p4. + rewrite List.skipn_app_r by reflexivity. cbv [Let_In_pf_nd]. + repeat rewrite <- app_assoc in *. rewrite H5p4. + rewrite List.skipn_app_r by reflexivity. reflexivity. + ++ cbv beta; intros; fwd. eexists _, t1'1, m1'1, l1'1, mc1'1, _. + ssplit; try eassumption. + --- sirs. + --- align_trace. + --- intros. rewrite sfix_step. simpl. simpl_rev. + repeat rewrite <- app_assoc in *. rewrite H3p4. + rewrite List.skipn_app_r by reflexivity. cbv [Let_In_pf_nd]. + simpl. rewrite H5p4. rewrite List.skipn_app_r by reflexivity. + rewrite H5p8. reflexivity. - (* exec.seq *) cbn in *. fwd. rename H1 into IH2, IHexec into IH1. eapply exec.seq. - + eapply IH1. 1: eassumption. eauto 15. + + eapply IH1. 1,2: eassumption. intros. rewrite H4. + rewrite sfix_step. reflexivity. + 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; try eassumption. intros. rewrite app_assoc. rewrite H4. + rewrite sfix_step. simpl. simpl_rev. rewrite H1p4. + rewrite List.skipn_app_r by reflexivity. reflexivity. + * cbv beta. intros. fwd. eexists _, t1'0, m1'0, l1'0, mc1'0, _. + ssplit; try eassumption. + -- solve_MetricLog. + -- align_trace. + -- intros. rewrite sfix_step. simpl. simpl_rev. + repeat rewrite <- app_assoc in *. rewrite H1p4. + rewrite List.skipn_app_r by reflexivity. rewrite H1p8. reflexivity. - (* exec.skip *) - eapply exec.skip. exists t, m, l, mc. repeat split; eauto; solve_MetricLog. + eapply exec.skip. eexists _, t, m, l, mc, _. ssplit; try eassumption. + + solve_MetricLog. + + align_trace. + + intros. rewrite sfix_step. reflexivity. Qed. - Lemma spill_fun_correct: forall e1 e2 argnames1 retnames1 body1 argnames2 retnames2 body2, + Lemma spill_fun_correct: forall e1 e2 pick_sp2 argnames1 retnames1 body1 argnames2 retnames2 body2, spill_functions e1 = Success e2 -> spill_fun (argnames1, retnames1, body1) = Success (argnames2, retnames2, body2) -> - 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'). + forall argvals kH kL t m mcH mcL (post: leakage -> Semantics.trace -> mem -> list word -> MetricLog -> Prop), + call_spec e1 (argnames1, retnames1, body1) (fun k => snd (fun_leakage e1 pick_sp2 (argnames1, retnames1, body1) (skipn (length kH) (rev k)) (rev kL))) kH t m argvals mcH post -> + call_spec_spilled e2 (argnames2, retnames2, body2) pick_sp2 kL t m argvals mcL + (fun kL' t' m' l' mcL' => + exists kH'' mcH', + post (kH'' ++ kH) t' m' l' mcH' /\ + metricsLeq (mcL' - mcL) (mcH' - mcH) /\ + fst (fun_leakage e1 pick_sp2 (argnames1, retnames1, body1) (rev kH'') (rev kL)) = rev kL'). Proof. intros. eapply spill_fun_correct_aux; try eassumption. - unfold spilling_correct_for. - eapply spilling_correct. - assumption. + apply spilling_correct. assumption. Qed. End Spilling. diff --git a/compiler/src/compiler/ToplevelLoop.v b/compiler/src/compiler/ToplevelLoop.v index 99a63f567..33f6986db 100644 --- a/compiler/src/compiler/ToplevelLoop.v +++ b/compiler/src/compiler/ToplevelLoop.v @@ -10,6 +10,7 @@ Require Export compiler.FlattenExprDef. Require Export compiler.FlattenExpr. Require compiler.FlatImp. Require Export riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Export riscv.Spec.Machine. Require Export riscv.Platform.Run. Require Export riscv.Platform.RiscvMachine. @@ -73,10 +74,10 @@ Section Pipeline1. Context {mem: map.map word byte}. Context {Registers: map.map Z word}. Context {string_keyed_map: forall T: Type, map.map string T}. (* abstract T for better reusability *) - Context {ext_spec: Semantics.ExtSpec}. + Context {ext_spec: LeakageSemantics.ExtSpec}. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {word_riscv_ok: RiscvWordProperties.word.riscv_ok word}. Context {string_keyed_map_ok: forall T, map.ok (string_keyed_map T)}. @@ -85,10 +86,11 @@ Section Pipeline1. Context {iset: InstructionSet}. Context {BWM: bitwidth_iset width iset}. Context {mem_ok: map.ok mem}. - Context {ext_spec_ok: Semantics.ext_spec.ok ext_spec}. + Context {ext_spec_ok: LeakageSemantics.ext_spec.ok ext_spec}. Context (compile_ext_call : string_keyed_map Z -> Z -> Z -> FlatImp.stmt Z -> list Instruction). + Context (leak_ext_call: word -> string_keyed_map Z -> Z -> Z -> FlatImp.stmt Z -> list word -> list LeakageEvent). Context (compile_ext_call_correct: forall resvars extcall argvars, - compiles_FlatToRiscv_correctly compile_ext_call compile_ext_call + compiles_FlatToRiscv_correctly compile_ext_call leak_ext_call compile_ext_call (FlatImp.SInteract resvars extcall argvars)). Context (compile_ext_call_length_ignores_positions: forall stackoffset posmap1 posmap2 c pos1 pos2, List.length (compile_ext_call posmap1 pos1 stackoffset c) = @@ -148,6 +150,7 @@ Section Pipeline1. exists mH, isReady spec mach.(getLog) mH /\ goodTrace spec mach.(getLog) /\ mach.(getPc) = word.add loop_pos (word.of_Z (if done then 4 else 0)) /\ + (exists kL0, mach.(getTrace) = Some kL0) /\ machine_ok functions_pos ml.(stack_start) ml.(stack_pastend) functions_instrs mH R (program iset init_sp_pos (init_sp_insts ++ init_insts init_fun_pos ++ @@ -180,6 +183,7 @@ Section Pipeline1. initial.(getNextPc) = word.add initial.(getPc) (word.of_Z 4) /\ regs_initialized initial.(getRegs) /\ initial.(getLog) = nil /\ + initial.(getTrace) = Some nil (*due to the statement of the compiler theorem, this just needs to not be None*) /\ valid_machine initial. Lemma signed_of_Z_small: forall c, @@ -292,28 +296,30 @@ Section Pipeline1. end. intros. destruct_RiscvMachine mid. fwd. (* then, run body of init function (using compiler simulation and correctness of init) *) + pose proof compiler_correct compile_ext_call leak_ext_call compile_ext_call_correct + compile_ext_call_length_ignores_positions as P. + unfold runsTo in P. + specialize P with (fname := "init"%string) (p_funcs := word.add loop_pos (word.of_Z 8)). + edestruct P as [f [pick_spH P'] ]; clear P. + { eassumption. } + { unfold compile. rewrite_match. reflexivity. } eapply runsTo_weaken. - - pose proof compiler_correct compile_ext_call compile_ext_call_correct - compile_ext_call_length_ignores_positions as P. - unfold runsTo in P. - specialize P with (argvals := []) - (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). + - specialize P' with (argvals := []) + (post := fun k' t' m' retvals mc' => isReady spec t' m' /\ goodTrace spec t'). + edestruct P' as (init_rel_pos & G & P''); clear P'; cycle -1. + 1: eapply P'' with (Rdata := R). all: simpl_MetricRiscvMachine_get_set. - 12: { + 11: { unfold hl_inv in init_code_correct. move init_code_correct at bottom. do 4 eexists. split. 1: eassumption. split. 1: reflexivity. - eapply ExprImp.weaken_exec. - - refine (init_code_correct _ _ _). + eapply MetricLeakageSemantics.exec.weaken. + - refine (init_code_correct _ _ _ _). replace (datamem_start spec) with (heap_start ml) by congruence. replace (datamem_pastend spec) with (heap_pastend ml) by congruence. exact HMem. - - cbv beta. intros * _ HP. exists []. split. 1: reflexivity. exact HP. + - cbv beta. intros _ * _ HP. exists []. split. 1: reflexivity. exact HP. } - 11: { unfold compile. rewrite_match. reflexivity. } all: try eassumption. { apply stack_length_divisible. } { cbn. clear CP. @@ -324,8 +330,9 @@ Section Pipeline1. { reflexivity. } { reflexivity. } { reflexivity. } + { reflexivity. } unfold machine_ok. - clear P'. + clear P''. rewrite GetPos in G. fwd. unfold_RiscvMachine_get_set. repeat match goal with @@ -455,6 +462,7 @@ Section Pipeline1. | |- _ => reflexivity end. + cbn. solve_word_eq word_ok. + + cbn. simpl in Hp5p3. subst. simpl. reflexivity. + cbn. destruct mlOk. subst. simpl in *. subst loop_pos init_pos. solve_divisibleBy4. - unfold ll_good, machine_ok. intros. fwd. assumption. @@ -500,22 +508,24 @@ Section Pipeline1. end. 2: solve_word_eq word_ok. subst. + pose proof compiler_correct compile_ext_call leak_ext_call compile_ext_call_correct + compile_ext_call_length_ignores_positions as P. + unfold runsTo in P. + specialize P with (fname := "loop"%string) (p_funcs := word.add loop_pos (word.of_Z 8)) (ret_addr := word.add loop_pos (word.of_Z 4)). + edestruct P as [f [pick_spH P'] ]; clear P. + { eassumption. } + { eassumption. } eapply runsTo_weaken. - + pose proof compiler_correct compile_ext_call compile_ext_call_correct - compile_ext_call_length_ignores_positions as P. - unfold runsTo in P. - specialize P with (argvals := []) - (fname := "loop"%string) - (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)). - 12: { + + specialize P' with (argvals := []) + (post := fun k' 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 (Rdata := R). + 11: { move loop_body_correct at bottom. do 4 eexists. split. 1: eassumption. split. 1: reflexivity. - eapply ExprImp.weaken_exec. + eapply MetricLeakageSemantics.exec.weaken. - eapply loop_body_correct; eauto. - - cbv beta. intros * _ HP. exists []. split. 1: reflexivity. exact HP. + - cbv beta. intros _ * _ HP. exists []. split. 1: reflexivity. exact HP. } all: try eassumption. all: simpl_MetricRiscvMachine_get_set. @@ -527,6 +537,7 @@ Section Pipeline1. { reflexivity. } { reflexivity. } { reflexivity. } + { reflexivity. } unfold loop_pos, init_pos. unfold machine_ok. unfold_RiscvMachine_get_set. @@ -571,6 +582,7 @@ Section Pipeline1. { f_equal. solve_word_eq word_ok. } { solve_word_eq word_ok. } { solve_word_eq word_ok. } + { f_equal. f_equal. f_equal. f_equal. solve_word_eq word_ok. } + cbv beta. intros. destruct_RiscvMachine final. @@ -589,7 +601,7 @@ Section Pipeline1. * wcancel_assumption. * eapply rearrange_footpr_subset. 1: eassumption. wwcancel. - Unshelve. exact MetricLogging.EmptyMetricLog. + Unshelve. 1: exact nil. exact MetricLogging.EmptyMetricLog. Qed. Lemma ll_inv_implies_prefix_of_good: forall st, diff --git a/compiler/src/compiler/UseImmediate.v b/compiler/src/compiler/UseImmediate.v index 7d9114c2e..371b8988e 100644 --- a/compiler/src/compiler/UseImmediate.v +++ b/compiler/src/compiler/UseImmediate.v @@ -1,4 +1,5 @@ Require Import compiler.util.Common. +Require Import bedrock2.LeakageSemantics. Require Import compiler.FlatImp. Require Import Coq.Lists.List. Import ListNotations. Require Import bedrock2.Syntax. @@ -17,7 +18,8 @@ Section WithArguments. Context {env : map.map string (list var * list var * stmt var) } {env_ok : map.ok env}. 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}. + Context {ext_spec : ExtSpec} {ext_spec_ok: ext_spec.ok ext_spec}. + Context {pick_sp: PickSp}. Context (is5BitImmediate : Z -> bool). Context (is12BitImmediate : Z -> bool). @@ -71,11 +73,11 @@ Section WithArguments. 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 -> + forall sH k t m mcH lH post, + exec eH sH k 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'). + exec eL (useImmediate is5BitImmediate is12BitImmediate sH) k t m lH mcL + (fun k' t' m' l' mcL' => exists mcH', metricsLeq (mcL' - mcL) (mcH' - mcH) /\ post k' t' m' l' mcH'). Proof. induction 2; try solve [ simpl; econstructor; eauto; @@ -91,13 +93,13 @@ Section WithArguments. exact Hmap. } cbv beta. intros * (?&?&Houtcome). - destruct (H4 _ _ _ _ Houtcome) as (retvs&l'&Hpost). + destruct (H4 _ _ _ _ _ Houtcome) as (retvs&l'&Hpost). exists retvs, l'. tandem Hpost. finish. - (* SStackalloc *) - simpl; econstructor; eauto. + simpl; econstructor; eauto. simpl in *. tandem H2. eapply exec.weaken; [eauto|]. simpl; intros * (?&?&?&?&?&?&?). @@ -120,7 +122,7 @@ Section WithArguments. { 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). + instantiate (1 := fun k t m l MC1 => exists MC2, MC1 - mcL <= MC2 - mc /\ mid2 k t m l MC2). finish. } { intros * (?&?&?). eapply exec.weaken; [eauto|]. @@ -145,18 +147,18 @@ Section WithArguments. all: eapply @exec.seq_cps; eapply @exec.lit. all: match goal with - | H: exec _ _ _ _ _ _ ?mid, - H': forall t m l mc, - ?mid _ _ _ _ -> exec ?eL _ _ _ _ _ ?post + | H: exec _ _ _ _ _ _ _ ?mid, + H': forall k t m l mc, + ?mid _ _ _ _ _ -> exec ?eL _ _ _ _ _ _ ?post |- _ => inversion H end. all: match goal with - | H: ?mid _ _ _ _, - H0: forall t m l mc, - ?mid t m l mc -> forall mcL, exec ?eL _ _ _ _ mcL _ - |- exec ?eL _ _ _ _ _ _ - => specialize (H0 _ _ _ _ H EmptyMetricLog); inversion H0 + | H: ?mid _ _ _ _ _, + H0: forall k t m l mc, + ?mid k t m l mc -> forall mcL, exec ?eL _ _ _ _ _ mcL _ + |- exec ?eL _ _ _ _ _ _ _ + => specialize (H0 _ _ _ _ _ H EmptyMetricLog); inversion H0 end. all: simpl in *; diff --git a/compiler/src/compiler/load_save_regs_correct.v b/compiler/src/compiler/load_save_regs_correct.v index 5134b8c69..b59d4d4bb 100644 --- a/compiler/src/compiler/load_save_regs_correct.v +++ b/compiler/src/compiler/load_save_regs_correct.v @@ -19,9 +19,9 @@ Section Proofs. Context {mem: map.map word byte}. Context {M: Type -> Type}. Context {MM: Monads.Monad M}. - Context {RVM: Machine.RiscvProgram M word}. + Context {RVM: Machine.RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. - Context {ext_spec: Semantics.ExtSpec}. + Context {ext_spec: LeakageSemantics.ExtSpec}. Context {word_riscv_ok: RiscvWordProperties.word.riscv_ok word}. Context {locals_ok: map.ok locals}. Context {mem_ok: map.ok mem}. @@ -61,20 +61,22 @@ Section Proofs. final.(getMetrics) = Platform.MetricLogging.addMetricInstructions (Z.of_nat (List.length vars)) (Platform.MetricLogging.addMetricStores (Z.of_nat (List.length vars)) - (Platform.MetricLogging.addMetricLoads (Z.of_nat (List.length vars)) initial.(getMetrics))) /\ + (Platform.MetricLogging.addMetricLoads (Z.of_nat (List.length vars)) initial.(getMetrics))) /\ + final.(getTrace) = + option_map (List.app (rev (leakage_events initial.(getPc) (save_regs iset vars offset) (leak_save_regs iset p_sp vars)))) initial.(getTrace) /\ valid_machine final). Proof. unfold map.getmany_of_list. induction vars; intros; subst addr. - simpl in *. simp. destruct oldvalues; simpl in *; [|discriminate]. apply runsToNonDet.runsToDone. repeat split; try assumption; try solve_word_eq word_ok. - destruct_RiscvMachine initial. destruct initial_metrics. MetricsToRiscv.solve_MetricLog. + + destruct_RiscvMachine initial. destruct initial_metrics. MetricsToRiscv.solve_MetricLog. + + destruct_RiscvMachine initial. destruct getTrace; reflexivity. - simpl in *. simp. assert (valid_register RegisterNames.sp) by (cbv; auto). destruct oldvalues as [|oldvalue oldvalues]; simpl in *; [discriminate|]. - replace (Memory.bytes_per_word (Decode.bitwidth iset)) with bytes_per_word in *. 2: { - rewrite bitwidth_matches. reflexivity. - } + replace (Memory.bytes_per_word (Decode.bitwidth iset)) with bytes_per_word in *. + 2: { rewrite bitwidth_matches. reflexivity. } eapply runsToNonDet.runsToStep. { eapply run_store_word. 7: eassumption. @@ -99,8 +101,9 @@ Section Proofs. etransitivity; [eassumption|]. replace (List.length vars) with (List.length oldvalues) by blia. solve_word_eq word_ok. - - rewrite H0p7. MetricsToRiscv.solve_MetricLog. - } + - rewrite H0p7. MetricsToRiscv.solve_MetricLog. + - subst. destruct getTrace; [|reflexivity]. simpl. repeat rewrite <- app_assoc. + reflexivity. } all: try eassumption. + simpl in *. etransitivity. 1: eassumption. ecancel. + simpl. use_sep_assumption. wcancel. @@ -135,12 +138,15 @@ Section Proofs. Platform.MetricLogging.addMetricInstructions (Z.of_nat (List.length vars)) (Platform.MetricLogging.addMetricLoads (Z.of_nat (2 * (List.length vars))) initial.(getMetrics)) /\ + final.(getTrace) = + option_map (List.app (rev (leakage_events initial.(getPc) (load_regs iset vars offset) (leak_load_regs iset p_sp vars)))) initial.(getTrace) /\ valid_machine final). Proof. induction vars; intros. - simpl in *. simp. destruct values; simpl in *; [|discriminate]. apply runsToNonDet.runsToDone. repeat split; try assumption; try solve_word_eq word_ok. - destruct_RiscvMachine initial. destruct initial_metrics. MetricsToRiscv.solve_MetricLog. + + destruct_RiscvMachine initial. destruct initial_metrics. MetricsToRiscv.solve_MetricLog. + + destruct_RiscvMachine initial. destruct getTrace; reflexivity. - simpl in *. simp. assert (valid_register RegisterNames.sp) by (cbv; auto). assert (valid_register a). { @@ -183,7 +189,9 @@ Section Proofs. rewrite Znat.Nat2Z.inj_succ. rewrite <- Z.add_1_r. replace (List.length values) with (List.length vars) by congruence. solve_word_eq word_ok. - * rewrite H1p3. MetricsToRiscv.solve_MetricLog. + * rewrite H1p3. MetricsToRiscv.solve_MetricLog. + * rewrite H1p4. destruct getTrace; [|reflexivity]. simpl. + repeat rewrite <- app_assoc. reflexivity. Qed. Lemma length_load_regs: forall vars offset, diff --git a/compiler/src/compiler/memory_mapped_ext_calls_compiler.v b/compiler/src/compiler/memory_mapped_ext_calls_compiler.v index e9cb05b57..ac31085b2 100644 --- a/compiler/src/compiler/memory_mapped_ext_calls_compiler.v +++ b/compiler/src/compiler/memory_mapped_ext_calls_compiler.v @@ -8,6 +8,7 @@ Require Import riscv.Utility.Monads. Require Import riscv.Utility.FreeMonad. Require Import compiler.FlatImp. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import coqutil.sanity. Require Import riscv.Utility.MkMachineWidth. Require Import riscv.Spec.PseudoInstructions. @@ -46,11 +47,23 @@ Definition compile_load{width: Z}{BW: Bitwidth width} if action =? "memory_mapped_extcall_read32" then (if Z.eqb width 32 then Lw else Lwu) else if action =? "memory_mapped_extcall_read16" then Lhu else Lbu. +Definition leak_load{width: Z}{BW: Bitwidth width}{word: word.word width} + (action: string)(addr : word) : InstructionLeakage := + if action =? "memory_mapped_extcall_read64" then I64Leakage (Ld_leakage addr) else + if action =? "memory_mapped_extcall_read32" then (if Z.eqb width 32 then ILeakage (Lw_leakage addr) else I64Leakage (Lwu_leakage addr)) else + if action =? "memory_mapped_extcall_read16" then ILeakage (Lhu_leakage addr) else ILeakage (Lbu_leakage addr). + Definition compile_store(action: string): Z -> Z -> Z -> Instruction := if action =? "memory_mapped_extcall_write64" then Sd else if action =? "memory_mapped_extcall_write32" then Sw else if action =? "memory_mapped_extcall_write16" then Sh else Sb. +Definition leak_store{width: Z}{BW: Bitwidth width}{word: word.word width} + (action: string)(addr: word): InstructionLeakage := + if action =? "memory_mapped_extcall_write64" then I64Leakage (Sd_leakage addr) else + if action =? "memory_mapped_extcall_write32" then ILeakage (Sw_leakage addr) else + if action =? "memory_mapped_extcall_write16" then ILeakage (Sh_leakage addr) else ILeakage (Sb_leakage addr). + Definition compile_interact{width: Z}{BW: Bitwidth width} (results: list Z)(a: string)(args: list Z): list Instruction := match args with @@ -59,6 +72,18 @@ Definition compile_interact{width: Z}{BW: Bitwidth width} | _ => [] (* invalid, excluded by ext_spec *) end. +Definition leak_interact{width: Z}{BW: Bitwidth width}{word: word.word width} + (abs_pos: word)(results: list Z)(a: string)(args: list Z)(leakage: list word): list LeakageEvent := + match leakage with + | [addr'] => + match args with + | [addr; val] => leakage_events abs_pos [compile_store a addr val 0] [leak_store a addr'] + | [addr] => leakage_events abs_pos [compile_load a (List.hd RegisterNames.a0 results) addr 0] [leak_load a addr'] + | _ => [] (* invalid, excluded by ext_spec *) + end + | _ => [] (*invalid*) + end. + Local Arguments Z.mul: simpl never. Local Arguments Z.add: simpl never. Local Arguments Z.of_nat: simpl never. @@ -110,6 +135,12 @@ Section MMIO. | _ => [] end. + Definition leak_ext_call(program_base: word)(_: funname_env Z)(pos _: Z)(s: stmt Z)(l: list word) := + match s with + | SInteract resvars action argvars => leak_interact (word.add program_base (word.of_Z pos)) resvars action argvars l + | _ => [] + end. + (* Lemma load4bytes_in_MMIO_is_None: forall (m: mem) (addr: word), map.undef_on m isMMIOAddr -> @@ -205,7 +236,12 @@ Section MMIO. cbv iota; unfold ExecuteI.execute, ExecuteI64.execute; eapply go_associativity; - (eapply go_getRegister; [ eassumption | eassumption | ]); + (epose proof go_getRegister as H; cbn [getRegister] in H; eapply H; [eassumption|eassumption|clear H]); + (*we have the above awkwardness because go_getRegister is stated in terms of the + getRegister of a RiscvProgramWithLeakage, while the goal here is (why?) in terms of + the getRegister of the RiscvProgram belonging to the RiscvProgramWithLeakage + ... i'm not sure how best to fix this. should go_getRegister be changed to use + the getRegister of a RiscvProgram?*) eapply go_associativity; eapply go_associativity; unfold mcomp_sat, Primitives.mcomp_sat, primitivesParams, Bind, free.Monad_free in *; @@ -225,6 +261,18 @@ Section MMIO. try (eapply H; assumption). Qed. + Lemma go_ext_call_leak_load: forall action x a addr (initialL: MetricRiscvMachine) post f, + valid_register a -> + map.get initialL.(getRegs) a = Some addr -> + mcomp_sat (f (Some (executeInstr (compile_load action x a 0) (leak_load action addr)))) initialL post -> + mcomp_sat (Bind (leakage_of_instr getRegister (compile_load action x a 0)) f) initialL post. + Proof. + intros. cbv [leakage_of_instr compile_load leak_load] in *. + repeat destruct_one_match. + all: epose proof go_getRegister as H'; cbn [getRegister] in H'; eapply H'; [eassumption|eassumption|clear H']. + all: assumption. + Qed. + Lemma go_ext_call_store: forall n action a addr v val mGive mKeep (initialL: MetricRiscvMachine) post f, action = "memory_mapped_extcall_write" ++ String.of_nat (n * 8) -> @@ -262,7 +310,7 @@ Section MMIO. cbv iota; unfold ExecuteI.execute, ExecuteI64.execute; eapply go_associativity; - (eapply go_getRegister; [ eassumption | eassumption | ]); + (epose proof go_getRegister as H; cbn [getRegister] in H; eapply H; [eassumption|eassumption|clear H]);(*do this for the same reason as in go_ext_call_load*) eapply go_associativity; eapply go_associativity; unfold mcomp_sat, Primitives.mcomp_sat, primitivesParams, Bind, free.Monad_free in *; @@ -287,6 +335,18 @@ Section MMIO. end. Qed. + Lemma go_ext_call_leak_store: forall action a v addr (initialL: MetricRiscvMachine) post f, + valid_register a -> + map.get initialL.(getRegs) a = Some addr -> + mcomp_sat (f (Some (executeInstr (compile_store action a v 0) (leak_store action addr)))) initialL post -> + mcomp_sat (Bind (leakage_of_instr getRegister (compile_store action a v 0)) f) initialL post. + Proof. + intros. cbv [leakage_of_instr compile_store leak_store] in *. + repeat destruct_one_match. + all: epose proof go_getRegister as H'; cbn [getRegister] in H'; eapply H'; [eassumption|eassumption|clear H']. + all: assumption. + Qed. + Lemma disjoint_load_bytes_is_None: forall m (addr: word) n, (0 < n)%nat -> disjoint (map.domain m) (of_list (footprint_list addr n)) -> @@ -338,7 +398,7 @@ Section MMIO. Qed. Lemma compile_ext_call_correct: forall resvars extcall argvars, - FlatToRiscvCommon.compiles_FlatToRiscv_correctly compile_ext_call compile_ext_call + FlatToRiscvCommon.compiles_FlatToRiscv_correctly compile_ext_call leak_ext_call compile_ext_call (FlatImp.SInteract resvars extcall argvars). Proof. unfold FlatToRiscvCommon.compiles_FlatToRiscv_correctly. simpl. intros. @@ -359,7 +419,7 @@ Section MMIO. end. destruct_RiscvMachine initialL. subst. lazymatch goal with - | H: ext_spec _ _ _ _ _ |- _ => destruct H as (n & nValid & HExt) + | H: leakage_ext_spec _ _ _ _ _ |- _ => destruct H as (n & nValid & HExt) end. destruct HExt; fwd. - (* load *) @@ -367,7 +427,8 @@ Section MMIO. destruct argvars. 2: discriminate Hl. clear Hl. fwd. unfold compile_interact in *. cbn [List.length] in *. - eapply go_fetch_inst; simpl_MetricRiscvMachine_get_set. + pose proof go_fetch_inst as gfi. cbn [getRegister] in gfi. + eapply gfi; clear gfi; simpl_MetricRiscvMachine_get_set. { reflexivity. } { eapply rearrange_footpr_subset. 1: eassumption. wwcancel. } { wcancel_assumption. } @@ -375,6 +436,9 @@ Section MMIO. repeat destruct nValid as [nValid|nValid]; [..|apply proj1 in nValid]; subst n. all: try exact I. destruct width_cases as [W | W]; rewrite W; exact I. } + eapply go_ext_call_leak_load. + { unfold valid_FlatImp_var, valid_register in *. lia. } + { cbn. unfold map.extends in *. eauto. } eapply go_ext_call_load. { reflexivity. } { assumption. } @@ -396,7 +460,7 @@ Section MMIO. cbv beta. intros v mRcv HO mL' Hsp. destruct HO as (HO & mExt' & HsE). lazymatch goal with - | H: forall _ _, outcome _ _ -> _ |- _ => + | H: forall _ _ _, outcome _ _ _ -> _ |- _ => specialize H with (1 := HO); destruct H as (l' & Hp & HPost) end. fwd. @@ -423,6 +487,10 @@ Section MMIO. eapply MapEauto.only_differ_put. cbn. left. reflexivity. } { unfold MetricCosts.cost_interact in *; MetricsToRiscv.solve_MetricLog. } + { instantiate (1:= [_]). reflexivity. } + { reflexivity. } + { intros. rewrite fix_step. cbn [stmt_leakage_body rev List.app]. + repeat rewrite <- app_assoc. reflexivity. } { eapply map.put_extends. eassumption. } { eapply map.forall_keys_put; assumption. } { rewrite map.get_put_diff; eauto. unfold RegisterNames.sp. @@ -481,7 +549,8 @@ Section MMIO. destruct argvars. 2: discriminate Hl. clear Hl. fwd. unfold compile_interact in *. cbn [List.length] in *. - eapply go_fetch_inst; simpl_MetricRiscvMachine_get_set. + pose proof go_fetch_inst as gfi. cbn [getRegister] in gfi. + eapply gfi; clear gfi; simpl_MetricRiscvMachine_get_set. { reflexivity. } { eapply rearrange_footpr_subset. 1: eassumption. wwcancel. } { wcancel_assumption. } @@ -500,6 +569,9 @@ Section MMIO. end; cbv; congruence. } + eapply go_ext_call_leak_store. + { unfold valid_FlatImp_var, valid_register in *. lia. } + { cbn. unfold map.extends in *. eauto. } eapply go_ext_call_store with (mKeep := map.putmany mKeep mL) (mGive := mGive). { reflexivity. } { assumption. } @@ -523,7 +595,7 @@ Section MMIO. replace (LittleEndian.split n (LittleEndian.combine n v)) with v. 1: assumption. symmetry. apply LittleEndian.split_combine. } lazymatch goal with - | HO: outcome _ _, H: forall _ _, outcome _ _ -> _ |- _ => + | HO: outcome _ _ _, H: forall _ _ _, outcome _ _ _ -> _ |- _ => specialize H with (1 := HO); destruct H as (l' & Hp & HPost) end. fwd. @@ -537,7 +609,11 @@ Section MMIO. eapply map.split_empty_r. reflexivity. } { reflexivity. } { eapply MapEauto.only_differ_refl. } - { unfold MetricCosts.cost_interact in *; MetricsToRiscv.solve_MetricLog. } + { cbv [id]. unfold MetricCosts.cost_interact in *; MetricsToRiscv.solve_MetricLog. } + { instantiate (1 := [_]). reflexivity. } + { reflexivity. } + { intros. rewrite fix_step. cbn [stmt_leakage_body rev List.app]. + repeat rewrite <- app_assoc. reflexivity. } { eassumption. } { assumption. } { assumption. } diff --git a/compiler/src/compiler/memory_mapped_ext_calls_riscv.v b/compiler/src/compiler/memory_mapped_ext_calls_riscv.v index 35668d7a4..6c73caf78 100644 --- a/compiler/src/compiler/memory_mapped_ext_calls_riscv.v +++ b/compiler/src/compiler/memory_mapped_ext_calls_riscv.v @@ -42,7 +42,7 @@ Section split_mcomp_sane. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {RVS: @Spec.Machine.RiscvMachine M word _ _ RVM}. Context {p: PrimitivesParams M MetricRiscvMachine}. @@ -163,6 +163,7 @@ Section Riscv. postF tt (withNextPc (word.add mach.(getPc) (word.of_Z 4)) mach) | EndCycleNormal => fun postF postA => postF tt (updatePc mach) | EndCycleEarly _ => fun postF postA => postA (updatePc mach) (* ignores postF containing the continuation *) + | LeakEvent e => fun postF postA => postF tt (withLeakageEvent e mach) | MakeReservation _ | ClearReservation _ | CheckReservation _ @@ -270,7 +271,7 @@ Section Riscv. destruct P as (v & mRcv & (N1 & N2)). destruct N2 as (mExt' & Sp). destruct mach. - eexists. eexists (mkMetricRiscvMachine (mkRiscvMachine _ _ _ _ _ _) _). + eexists. eexists (mkMetricRiscvMachine (mkRiscvMachine _ _ _ _ _ _ _) _). cbn -[HList.tuple String.append] in *. eapply N1. clear N1 HI. unfold map.split. split. 1: reflexivity. diff --git a/compiler/src/compilerExamples/AssemblyVerif.v b/compiler/src/compilerExamples/AssemblyVerif.v index 4f7fb99ee..7a8f716fb 100644 --- a/compiler/src/compilerExamples/AssemblyVerif.v +++ b/compiler/src/compilerExamples/AssemblyVerif.v @@ -58,7 +58,7 @@ Section Verif. Context {mem_ok: map.ok mem}. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {PR: MetricPrimitives PRParams}. diff --git a/compiler/src/compilerExamples/SimpleInvariant.v b/compiler/src/compilerExamples/SimpleInvariant.v index beddb5dd2..fb115b721 100644 --- a/compiler/src/compilerExamples/SimpleInvariant.v +++ b/compiler/src/compilerExamples/SimpleInvariant.v @@ -34,9 +34,9 @@ Section Proofs. Context {env: map.map String.string (list Z * list Z * FlatImp.stmt Z)}. Context {M: Type -> Type}. Context {MM: Monads.Monad M}. - Context {RVM: Machine.RiscvProgram M word}. + Context {RVM: Machine.RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. - Context {ext_spec: Semantics.ExtSpec}. + Context {ext_spec: LeakageSemantics.ExtSpec}. Context {word_riscv_ok: RiscvWordProperties.word.riscv_ok word}. Context {locals_ok: map.ok locals}. Context {mem_ok: map.ok mem}. @@ -119,7 +119,7 @@ Section PrintExamples. Context {env: map.map String.string (list Z * list Z * FlatImp.stmt Z)}. Context {M: Type -> Type}. Context {MM: Monads.Monad M}. - Context {RVM: Machine.RiscvProgram M word}. + Context {RVM: Machine.RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {ext_spec: Semantics.ExtSpec}. Context {word_riscv_ok: RiscvWordProperties.word.riscv_ok word}. diff --git a/compiler/src/compilerExamples/memequal.v b/compiler/src/compilerExamples/memequal.v new file mode 100644 index 000000000..436fdb6c0 --- /dev/null +++ b/compiler/src/compilerExamples/memequal.v @@ -0,0 +1,144 @@ +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 coqutil.Tactics.fwd. +Require Import compiler.MemoryLayout. +Require Import compiler.StringNameGen. +Require Import riscv.Utility.InstructionCoercions. +Require Import riscv.Platform.MetricRiscvMachine. +Require Import riscv.Spec.LeakageOfInstr. +Require Import riscv.Spec.Decode. +Require compiler.NaiveRiscvWordProperties. +Require Import bedrock2Examples.memequal. +Require Import compiler.MMIO. +Require Import bedrock2.FE310CSemantics. +Require Import coqutil.Map.SortedListZ. +Require Import compiler.util.Common. + +Open Scope Z_scope. Open Scope string_scope. Open Scope ilist_scope. + +Notation RiscvMachine := MetricRiscvMachine. + +Local Existing Instance coqutil.Map.SortedListString.map. +Local Existing Instance coqutil.Map.SortedListString.ok. +Local Instance mem32 : map.map Words32Naive.word Init.Byte.byte := SortedListWord.map Words32Naive.word byte. +Local Instance mem32_ok : map.ok mem32 := SortedListWord.ok _ _. +Local Instance localsL32 : map.map Z Words32Naive.word := SortedListZ.map Words32Naive.word. +Local Instance localsL32_ok : map.ok localsL32 := SortedListZ.ok _. +Local Instance localsH32: map.map string Words32Naive.word := SortedListString.map Words32Naive.word. +Local Instance localsH32_ok : map.ok localsH32 := SortedListString.ok _. +Local Instance RV32I_bitwidth: FlatToRiscvCommon.bitwidth_iset 32 RV32I := eq_refl. + +Lemma word_ok : RiscvWordProperties.word.riscv_ok Words32Naive.word. +Proof. + cbv [Words32Naive.word]. replace 32 with (2 ^ BinInt.Z.of_nat 5) by reflexivity. + apply NaiveRiscvWordProperties.naive_word_riscv_ok. +Qed. + +Lemma word_ok' : word.ok Words32Naive.word. +Proof. exact Naive.word32_ok. Qed. + +Definition fs_memequal := &[,memequal]. +Definition instrs_memequal := + match (compile compile_ext_call fs_memequal) with + | Success (instrs, _, _) => instrs + | _ => nil + end. +Definition finfo_memequal := + match (compile compile_ext_call fs_memequal) with + | Success (_, finfo, _) => finfo + | _ => nil + end. +Definition req_stack_size_memequal := + match (compile compile_ext_call fs_memequal) with + | Success (_, _, req_stack_size) => req_stack_size + | _ => 0 + end. +Definition fname_memequal := "memequal". +Definition f_rel_pos_memequal := 0. +Definition post : list LogItem -> mem32 -> list Words32Naive.word -> Prop := fun _ _ _ => True. + +Lemma memequal_ct : + forall x y n p_funcs stack_hi ret_addr, + exists finalTrace : list LeakageEvent, + forall Rx Ry xs ys m stack_lo + Rdata Rexec (initial : RiscvMachine), + Separation.sep (Array.array Separation.ptsto (word.of_Z 1) x xs) Rx m /\ + Separation.sep (Array.array Separation.ptsto (word.of_Z 1) y ys) Ry m /\ + Z.of_nat (Datatypes.length xs) = word.unsigned n /\ + Z.of_nat (Datatypes.length ys) = word.unsigned n /\ + req_stack_size_memequal <= word.unsigned (word.sub stack_hi stack_lo) / SeparationLogic.bytes_per_word -> + word.unsigned (word.sub stack_hi stack_lo) mod SeparationLogic.bytes_per_word = 0 -> + getPc initial = word.add p_funcs (word.of_Z f_rel_pos_memequal) -> + initial.(getTrace) = Some [] -> + map.get (getRegs initial) RegisterNames.ra = Some ret_addr -> + word.unsigned ret_addr mod 4 = 0 -> + LowerPipeline.arg_regs_contain (getRegs initial) [x; y; n] -> + LowerPipeline.machine_ok p_funcs stack_lo stack_hi instrs_memequal m Rdata Rexec initial -> + FlatToRiscvCommon.runsTo initial + (fun final : RiscvMachine => + (exists mH' (retvals : list Words32Naive.word), + LowerPipeline.arg_regs_contain (getRegs final) retvals /\ + post (getLog final) mH' retvals /\ + map.only_differ (getRegs initial) reg_class.caller_saved (getRegs final) /\ + getPc final = ret_addr /\ + final.(getTrace) = Some finalTrace /\ + LowerPipeline.machine_ok p_funcs stack_lo stack_hi instrs_memequal mH' + Rdata Rexec final)). +Proof. + assert (spec := @memequal_ok _ _ Words32Naive.word mem32 (SortedListString.map (@Naive.rep 32)) leakage_ext_spec). + intros. + edestruct (@compiler_correct_wp _ _ Words32Naive.word mem32 _ leakage_ext_spec _ _ _ leakage_ext_spec_ok _ _ _ _ _ word_ok _ _ RV32I _ compile_ext_call leak_ext_call compile_ext_call_correct ltac:(reflexivity) fs_memequal instrs_memequal finfo_memequal req_stack_size_memequal fname_memequal p_funcs stack_hi ret_addr f_rel_pos_memequal) as [f_ [pick_sp_ H] ]. + { simpl. reflexivity. } + { vm_compute. reflexivity. } + specialize (spec pick_sp_ word_ok' _ ltac:(apply SortedListString.ok) leakage_ext_spec_ok). + cbv [LeakageProgramLogic.program_logic_goal_for] in spec. + specialize (spec (map.of_list fs_memequal) eq_refl). + cbv [spec_of_memequal] in spec. destruct spec as [f spec]. + eexists. intros. + cbv [FlatToRiscvCommon.runsTo]. + eapply runsToNonDet.runsTo_weaken. + 1: eapply H with (post := (fun k_ t_ m_ r_ mc => k_ = _ /\ + post t_ m_ r_)). + { simpl. repeat constructor. tauto. } + 2: { eapply map.get_of_list_In_NoDup. + { vm_compute. repeat constructor; eauto. } + { vm_compute. left. reflexivity. } } + all: try eassumption. + 2: { fwd. assumption. } + + { subst. cbv [fname_memequal]. move spec at bottom. + specialize (spec x y n xs ys Rx Ry [] (initial.(getLog)) m ltac:(intuition eauto)). + eapply MetricLeakageSemantics.weaken_call. + { eapply MetricLeakageSemantics.to_leakage_call. exact spec. } + cbv beta. intros. fwd. split; [|reflexivity]. reflexivity. } + { reflexivity. } + { reflexivity. } + cbv beta. intros. fwd. do 2 eexists. intuition eauto. + Unshelve. exact EmptyMetricLog. +Qed. +(* Print Assumptions memequal_ct. *) +(* + Axioms: + PropExtensionality.propositional_extensionality : forall P Q : Prop, P <-> Q -> P = Q + FunctionalExtensionality.functional_extensionality_dep : + forall (A : Type) (B : A -> Type) (f g : forall x : A, B x), (forall x : A, f x = g x) -> f = g + *) diff --git a/deps/riscv-coq b/deps/riscv-coq index d0afd4b58..8587bd404 160000 --- a/deps/riscv-coq +++ b/deps/riscv-coq @@ -1 +1 @@ -Subproject commit d0afd4b58178976a2887c07e4f05c15d757fa0fc +Subproject commit 8587bd40439f9f8c995b206b78b4100024bbde4f diff --git a/end2end/src/end2end/End2EndLightbulb.v b/end2end/src/end2end/End2EndLightbulb.v index e1437e21b..15628bdeb 100644 --- a/end2end/src/end2end/End2EndLightbulb.v +++ b/end2end/src/end2end/End2EndLightbulb.v @@ -260,7 +260,7 @@ Proof. - reflexivity. - (* preserve invariant *) intros. - specialize (H ltac:(repeat constructor)). + specialize (H ltac:(repeat constructor) ltac:(repeat constructor)). unfold hl_inv, isReady, goodTrace, goodHlTrace in *. Simp.simp. repeat ProgramLogic.straightline. diff --git a/end2end/src/end2end/End2EndPipeline.v b/end2end/src/end2end/End2EndPipeline.v index dcb547d65..3bd216db8 100644 --- a/end2end/src/end2end/End2EndPipeline.v +++ b/end2end/src/end2end/End2EndPipeline.v @@ -190,7 +190,7 @@ Section Connect. Definition goodTraceE(t: list Event): Prop := exists bedrockTrace, traces_related t bedrockTrace /\ spec.(goodTrace) bedrockTrace. - Definition bedrock2Inv := (fun t m l => forall mc, hl_inv spec t m l mc). + Definition bedrock2Inv := (fun t m l => forall k mc, hl_inv spec k t m l mc). Hypothesis goodTrace_implies_related_to_Events: forall (t: list LogItem), spec.(goodTrace) t -> exists t': list Event, traces_related t' t. @@ -333,7 +333,7 @@ Section Connect. (* destruct spec. TODO why "Error: sat is already used." ?? *) (* 2) riscv-coq to bedrock2 semantics *) - pose proof compiler_invariant_proofs compile_ext_call compile_ext_call_correct as P2. + pose proof compiler_invariant_proofs compile_ext_call leak_ext_call compile_ext_call_correct as P2. specialize_first P2 spec. specialize_first P2 ml. specialize_first P2 mlOk. @@ -356,19 +356,19 @@ Section Connect. * exact V. * exact GetInit. * intros. - eapply ExprImp.weaken_exec. + eapply MetricLeakageSemantics.exec.weaken. -- match goal with | H: LowerPipeline.mem_available ?from ?to _ |- _ => rewrite heap_start_agree in H; rewrite heap_pastend_agree in H end. - eapply MetricSemantics.of_metrics_free. + eapply MetricLeakageSemantics.to_plain_exec. eapply WeakestPreconditionProperties.sound_cmd; eauto. -- simpl. clear. intros. unfold bedrock2Inv in *. eauto. * exact GetLoop. * intros. unfold bedrock2Inv in *. - eapply ExprImp.weaken_exec. - -- eapply MetricSemantics.of_metrics_free. + eapply MetricLeakageSemantics.exec.weaken. + -- eapply MetricLeakageSemantics.to_plain_exec. eapply WeakestPreconditionProperties.sound_cmd; eauto. -- simpl. clear. intros. eauto. + assumption. @@ -512,6 +512,7 @@ Section Connect. | H: forall _, _ -> _ <> None |- _ => eapply H; eauto end. + reflexivity. + + assumption. + simpl. split. * apply @riscv_init_memory_undef_on_MMIO with (instrMemSizeLg:= instrMemSizeLg). { apply instrMemSizeLg_bounds. } diff --git a/etc/test_itauto.sh b/etc/test_itauto.sh deleted file mode 100755 index 51e5def80..000000000 --- a/etc/test_itauto.sh +++ /dev/null @@ -1,44 +0,0 @@ -#!/bin/sh - -printf '\033c' - -set -e -set -o pipefail -set -x - -cd deps/coq -git clean -dfx -git fetch -git reset origin/for_itauto --hard -git show -s --oneline HEAD -./configure -local -make -cd ../.. - -cd deps/itauto/ -git clean -dfx -git pull -git show -s --oneline HEAD -make -make install -cd ../.. - -cd deps/coq-record-update/ -git clean -dfx -cd ../.. - -cd deps/coqutil/ -git clean -dfx -cd ../.. - -cd deps/kami/ -git clean -dfx -cd ../.. - -cd deps/riscv-coq/ -git clean -dfx -cd ../.. - -git clean -dfx - -time TIMED=1 make 2>&1 | tee log.txt | grep -v -e ' ran for 0' -e '^Tactic call *$' diff --git a/processor/src/processor/KamiRiscv.v b/processor/src/processor/KamiRiscv.v index 8b093d500..a2298baf5 100644 --- a/processor/src/processor/KamiRiscv.v +++ b/processor/src/processor/KamiRiscv.v @@ -447,7 +447,8 @@ Section Equiv. RiscvMachine.getNextPc := word.of_Z 4; RiscvMachine.getMem := riscvMemInit; RiscvMachine.getXAddrs := kamiXAddrs instrMemSizeLg; - RiscvMachine.getLog := nil; (* <-- intended to be nil *) |}; + RiscvMachine.getLog := nil; (* <-- intended to be nil *) + RiscvMachine.getTrace := Some nil |}; getMetrics := MetricLogging.EmptyMetricLog; |}. Proof. econstructor; try reflexivity. @@ -530,6 +531,7 @@ Section Equiv. disjoint (of_list m0RV.(RiscvMachine.getXAddrs)) isMMIOAddr -> (forall reg, 0 < reg < 32 -> map.get m0RV.(getRegs) reg <> None) -> m0RV.(getLog) = nil -> + m0RV.(getTrace) = Some nil -> RvInv m0RV) (preserveRvInv: forall (m: RiscvMachine), RvInv m -> mcomp_sat_unit (run1 iset) m RvInv) diff --git a/processor/src/processor/KamiRiscvStep.v b/processor/src/processor/KamiRiscvStep.v index 389946f4b..5c8d685bc 100644 --- a/processor/src/processor/KamiRiscvStep.v +++ b/processor/src/processor/KamiRiscvStep.v @@ -575,7 +575,7 @@ Section Equiv. Inductive states_related: KamiMachine * list Event -> RiscvMachine -> Prop := | relate_states: - forall t t' m riscvXAddrs kpc krf rrf rpc nrpc pinit instrMem kdataMem rdataMem metrics, + forall k t t' m riscvXAddrs kpc krf rrf rpc nrpc pinit instrMem kdataMem rdataMem metrics, traces_related t t' -> KamiProc.RegsToT m = Some (kamiStMk kpc krf pinit instrMem kdataMem) -> (pinit = false -> riscvXAddrs = kamiXAddrs) -> @@ -590,7 +590,8 @@ Section Equiv. RiscvMachine.getNextPc := nrpc; RiscvMachine.getMem := rdataMem; RiscvMachine.getXAddrs := riscvXAddrs; - RiscvMachine.getLog := t'; |}; + RiscvMachine.getLog := t'; + RiscvMachine.getTrace := k; |}; getMetrics := metrics; |}. Inductive KamiLabelR: Kami.Semantics.LabelT -> list Event -> Prop := @@ -1111,7 +1112,7 @@ Section Equiv. fst snd getMetrics getMachine translate - getRegs getPc getNextPc getMem getXAddrs getLog] + getRegs getPc getNextPc getMem getXAddrs getLog getTrace] in (interp_action a state (fun x state' => mcomp_sat (kV x) state' post)) in change TR in HR; subst kV | free.ret ?v => change (post v state) in HR @@ -1212,8 +1213,8 @@ Section Equiv. (cbv beta delta [load store] in H; cbn beta iota delta [ load store fst snd translate - withMetrics updateMetrics getMachine getMetrics getRegs getPc getNextPc getMem getXAddrs getLog withRegs withPc withNextPc withMem withXAddrs withLog withLogItem withLogItems - RiscvMachine.withRegs RiscvMachine.withPc RiscvMachine.withNextPc RiscvMachine.withMem RiscvMachine.withXAddrs RiscvMachine.withLog RiscvMachine.withLogItem RiscvMachine.withLogItems] in H) + withMetrics updateMetrics getMachine getMetrics getRegs getPc getNextPc getMem getXAddrs getLog getTrace withRegs withPc withNextPc withMem withXAddrs withLog withLogItem withLogItems withLeakageEvent withLeakageEvents + RiscvMachine.withRegs RiscvMachine.withPc RiscvMachine.withNextPc RiscvMachine.withMem RiscvMachine.withXAddrs RiscvMachine.withLog RiscvMachine.withLogItem RiscvMachine.withLogItems RiscvMachine.withLeakageEvent RiscvMachine.withLeakageEvents] in H) end. Ltac rt := repeat (r || t). @@ -1642,6 +1643,7 @@ Section Equiv. (** Evaluate (invert) the two fetchers *) rt. eval_kami_fetch. rt. + subst inst'. rewrite H11 in *. (** Begin symbolic evaluation of Kami decode/execute *) kami_cbn_hint Heqic. @@ -1721,14 +1723,15 @@ Section Equiv. | H : Instruction |- _ => clear H end. - all: replace (getReg rrf rs1) with - (if Z.eq_dec rs1 0 then word.of_Z 0 - else match map.get rrf rs1 with - | Some x => x - | None => word.of_Z 0 - end) in *. - 2,4,6,8,10,12: - unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; + all: set (v' := if Z.eq_dec rs1 0 then word.of_Z 0 + else match map.get rrf rs1 with + | Some x => x + | None => word.of_Z 0 + end). + all: try (assert (getReg rrf rs1 = v') as Hv'; [|rewrite Hv' in *]). + + 1,3,5,7,9: + subst v'; unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; [ Lia.lia | pose proof bitSlice_range_ex (@kunsigned 32 kinst) 15 20 as HR; change (bitSlice (@kunsigned 32 kinst) 15 20) with rs1 in HR; @@ -1742,7 +1745,8 @@ Section Equiv. | [H: match Memory.load_bytes ?sz ?m ?a with | Some _ => _ | None => _ end |- _] => destruct (Memory.load_bytes sz m a) as [lv|] eqn:Hlv; [exfalso|] end. - all: try (subst v oimm12; + all: subst v; rewrite Hv' in *. + all: try (subst v' oimm12; regs_get_red Hlv; match goal with | [Heqic: true = evalExpr (isMMIO _ _) |- _] => @@ -1790,7 +1794,7 @@ Section Equiv. prove_states_related. { kami_struct_cbv_goal; cbn [evalExpr evalConstT]. - subst v oimm12 ldVal. + subst v' oimm12 ldVal. regs_get_red_goal. constructor; [|assumption]. apply events_related_mmioLoadEvent. @@ -1817,6 +1821,7 @@ Section Equiv. (** Evaluate (invert) the two fetchers *) rt. eval_kami_fetch. rt. + subst inst'. rewrite H11 in *. (** Symbolic evaluation of Kami decode/execute *) clear Heqic0. @@ -1895,14 +1900,15 @@ Section Equiv. | H : Instruction |- _ => clear H end. - all: replace (getReg rrf rs1) with - (if Z.eq_dec rs1 0 then word.of_Z 0 - else match map.get rrf rs1 with - | Some x => x - | None => word.of_Z 0 - end) in *. - 2,4,6,8,10,12: - unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; + all: set (v' := if Z.eq_dec rs1 0 then word.of_Z 0 + else match map.get rrf rs1 with + | Some x => x + | None => word.of_Z 0 + end). + all: try (assert (getReg rrf rs1 = v') as Hv'; [|rewrite Hv' in *]). + + 1,3,5,7,9: + subst v'; unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; [ Lia.lia | pose proof bitSlice_range_ex (@kunsigned 32 kinst) 15 20 as HR; change (bitSlice (@kunsigned 32 kinst) 15 20) with rs1 in HR; @@ -1916,13 +1922,13 @@ Section Equiv. | [H: match Memory.load_bytes ?sz ?m ?a with | Some _ => _ | None => _ end |- _] => destruct (Memory.load_bytes sz m a) as [lv|] eqn:Hlv end. - + all: try (subst v; rewrite Hv' in *). all: try match goal with | [H: nonmem_load _ _ _ _ _ |- _] => destruct H as [? [[? ?] ?]]; discriminate end. 6: { exfalso. - subst v oimm12. + subst v' oimm12. destruct H13 as [? [? ?]]. regs_get_red H13. apply is_mmio_sound in H13. @@ -1951,7 +1957,7 @@ Section Equiv. all: regs_get_red_goal. all: cbv [int8ToReg int16ToReg uInt8ToReg uInt16ToReg int32ToReg MachineWidth_XLEN word.of_Z word wordW KamiWord.word kofZ]. - all: subst v oimm12 rs1. + all: subst v' oimm12 rs1. all: regs_get_red Hlv. all: cbv [Utility.add ZToReg MachineWidth_XLEN @@ -2086,6 +2092,7 @@ Section Equiv. (** Evaluate (invert) the two fetchers *) rt. eval_kami_fetch. rt. + subst inst'. rewrite H11 in *. (** Begin symbolic evaluation of Kami decode/execute *) kami_cbn_hint Heqic. @@ -2161,14 +2168,15 @@ Section Equiv. | H : Instruction |- _ => clear H end. - all: replace (getReg rrf rs1) with - (if Z.eq_dec rs1 0 then word.of_Z 0 - else match map.get rrf rs1 with - | Some x => x - | None => word.of_Z 0 - end) in *. - 2,4,6,8,10,12: - unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; + all: set (v' := if Z.eq_dec rs1 0 then word.of_Z 0 + else match map.get rrf rs1 with + | Some x => x + | None => word.of_Z 0 + end). + all: try (assert (getReg rrf rs1 = v') as Hv'; [|rewrite Hv' in *]). + + 1,3,5,7,9: + subst v'; unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; [ Lia.lia | pose proof bitSlice_range_ex (@kunsigned 32 kinst) 15 20 as HR; change (bitSlice (@kunsigned 32 kinst) 15 20) with rs1 in HR; @@ -2181,8 +2189,9 @@ Section Equiv. all: try match goal with | [H: match Memory.load_bytes ?sz ?m ?a with | Some _ => _ | None => _ end |- _] => destruct (Memory.load_bytes sz m a) as [lv|] eqn:Hlv; [exfalso|] - end. - all: try (subst v oimm12; + end. + all: subst v; rewrite Hv' in *. + all: try (subst v' oimm12; regs_get_red Hlv; match goal with | [Heqic: true = evalExpr (isMMIO _ _) |- _] => @@ -2217,7 +2226,7 @@ Section Equiv. else map.put rrf rd newval) in * end. 2: { unfold setReg; repeat destruct_one_match; try reflexivity; [ | contradiction ]. - exfalso. clear -E. intuition Lia.lia. + exfalso. clear -E0. intuition Lia.lia. } rt. @@ -2227,7 +2236,7 @@ Section Equiv. prove_states_related. { kami_struct_cbv_goal; cbn [evalExpr evalConstT]. - subst v oimm12. + subst v' oimm12. regs_get_red_goal. constructor; [|assumption]. apply events_related_mmioLoadEvent. @@ -2248,6 +2257,7 @@ Section Equiv. (** Evaluate (invert) the two fetchers *) rt. eval_kami_fetch. rt. + subst inst'. rewrite H11 in *. (** Symbolic evaluation of Kami decode/execute *) clear Heqic0. @@ -2322,14 +2332,15 @@ Section Equiv. | H : Instruction |- _ => clear H end. - all: replace (getReg rrf rs1) with - (if Z.eq_dec rs1 0 then word.of_Z 0 - else match map.get rrf rs1 with - | Some x => x - | None => word.of_Z 0 - end) in *. - 2,4,6,8,10,12: - unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; + all: set (v' := if Z.eq_dec rs1 0 then word.of_Z 0 + else match map.get rrf rs1 with + | Some x => x + | None => word.of_Z 0 + end). + all: try (assert (getReg rrf rs1 = v') as Hv'; [|rewrite Hv' in *]). + + 1,3,5,7,9: + subst v'; unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; [ Lia.lia | pose proof bitSlice_range_ex (@kunsigned 32 kinst) 15 20 as HR; change (bitSlice (@kunsigned 32 kinst) 15 20) with rs1 in HR; @@ -2343,13 +2354,13 @@ Section Equiv. | [H: match Memory.load_bytes ?sz ?m ?a with | Some _ => _ | None => _ end |- _] => destruct (Memory.load_bytes sz m a) as [lv|] eqn:Hlv end. - + all: subst v; rewrite Hv' in *. all: try match goal with | [H: nonmem_load _ _ _ _ _ |- _] => destruct H as [? [[? ?] ?]]; discriminate end. 4: { exfalso. - subst v oimm12. + subst v' oimm12. destruct H13 as [? [? ?]]. regs_get_red H13. apply is_mmio_sound in H13. @@ -2410,6 +2421,7 @@ Section Equiv. (** Evaluate (invert) the two fetchers *) rt. eval_kami_fetch. rt. + subst inst'. rewrite H11 in *. (** Begin symbolic evaluation of Kami decode/execute *) kami_cbn_hint Heqic. @@ -2476,35 +2488,35 @@ Section Equiv. | H : Instruction |- _ => clear H end. - all: replace (getReg rrf rs1) with - (if Z.eq_dec rs1 0 then word.of_Z 0 - else match map.get rrf rs1 with - | Some x => x - | None => word.of_Z 0 - end) in * - by ( - unfold getReg; repeat destruct_one_match; try reflexivity; - exfalso; - [ Lia.lia - | pose proof bitSlice_range_ex (@kunsigned 32 kinst) 15 20 as HR; - change (bitSlice (@kunsigned 32 kinst) 15 20) with rs1 in HR; - Lia.lia ]). - - all: repeat r; t. + all: set (v' := if Z.eq_dec rs1 0 then word.of_Z 0 + else match map.get rrf rs1 with + | Some x => x + | None => word.of_Z 0 + end). + all: try (assert (getReg rrf rs1 = v') as Hv'; [|rewrite Hv' in *]). + 1,3,5: + subst v'; unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; + [ Lia.lia + | pose proof bitSlice_range_ex (@kunsigned 32 kinst) 15 20 as HR; + change (bitSlice (@kunsigned 32 kinst) 15 20) with rs1 in HR; + Lia.lia ]. - all: replace (getReg rrf rs2) with - (if Z.eq_dec rs2 0 then word.of_Z 0 + all: repeat rt. + all: subst v; rewrite Hv' in *. + all: subst v0. + all: set (v2' := if Z.eq_dec rs2 0 then word.of_Z 0 else match map.get rrf rs2 with | Some x => x | None => word.of_Z 0 - end) in * - by ( - unfold getReg; repeat destruct_one_match; try reflexivity; + end). + all: try (assert (getReg rrf rs2 = v2') as Hv2'; [|rewrite Hv2' in *]). + 1,3,5: + subst v2'; unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; [ Lia.lia | pose proof bitSlice_range_ex (@kunsigned 32 kinst) 20 25 as HR; change (bitSlice (@kunsigned 32 kinst) 20 25) with rs2 in HR; - Lia.lia ]). + Lia.lia ]. (** Consistency proof for each instruction *) all: rt. @@ -2518,7 +2530,7 @@ Section Equiv. all: rewrite @kunsigned_combine_shiftl_lor with (sa:= 5%nat) (sb:= 7%nat) in *. all: simpl_bit_manip. - all: try (subst v simm12; + all: try (subst v' simm12; regs_get_red Hst; cbv [Memory.store_bytes] in Hst; destruct (Memory.load_bytes _ _ _) eqn:Hlv in Hst; [clear Hst|discriminate]; @@ -2546,7 +2558,7 @@ Section Equiv. prove_states_related. { kami_struct_cbv_goal; cbn [evalExpr evalConstT]. - subst v simm12. + subst v' simm12. regs_get_red_goal. constructor; [|assumption]. apply events_related_mmioStoreEvent. @@ -2558,7 +2570,7 @@ Section Equiv. rewrite unsigned_split2_as_bitSlice. reflexivity. } - { subst v0; regs_get_red_goal. + { subst v2'; regs_get_red_goal. cbv [regToInt32 MachineWidth_XLEN word.unsigned word wordW KamiWord.word kofZ]. setoid_rewrite signExtend_combine_split_unsigned. @@ -2579,6 +2591,7 @@ Section Equiv. (** Evaluate (invert) the two fetchers *) rt. eval_kami_fetch. rt. + subst inst'. rewrite H11 in *. (** Symbolic evaluation of Kami decode/execute *) clear Heqic0. @@ -2646,33 +2659,35 @@ Section Equiv. | H : Instruction |- _ => clear H end. - all: replace (getReg rrf rs1) with - (if Z.eq_dec rs1 0 then word.of_Z 0 - else match map.get rrf rs1 with - | Some x => x - | None => word.of_Z 0 - end) in * - by ( - unfold getReg; repeat destruct_one_match; try reflexivity; - exfalso; - [ Lia.lia - | pose proof bitSlice_range_ex (@kunsigned 32 kinst) 15 20 as HR; - change (bitSlice (@kunsigned 32 kinst) 15 20) with rs1 in HR; - Lia.lia ]). - all: repeat r; t. - all: replace (getReg rrf rs2) with - (if Z.eq_dec rs2 0 then word.of_Z 0 + all: set (v' := if Z.eq_dec rs1 0 then word.of_Z 0 + else match map.get rrf rs1 with + | Some x => x + | None => word.of_Z 0 + end). + all: try (assert (getReg rrf rs1 = v') as Hv'; [|rewrite Hv' in *]). + 1,3,5: + subst v'; unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; + [ Lia.lia + | pose proof bitSlice_range_ex (@kunsigned 32 kinst) 15 20 as HR; + change (bitSlice (@kunsigned 32 kinst) 15 20) with rs1 in HR; + Lia.lia ]. + + all: repeat rt. + all: subst v; rewrite Hv' in *. + all: subst v0. + all: set (v2' := if Z.eq_dec rs2 0 then word.of_Z 0 else match map.get rrf rs2 with | Some x => x | None => word.of_Z 0 - end) in * - by ( - unfold getReg; repeat destruct_one_match; try reflexivity; + end). + all: try (assert (getReg rrf rs2 = v2') as Hv2'; [|rewrite Hv2' in *]). + 1,3,5: + subst v2'; unfold getReg; repeat destruct_one_match; try reflexivity; exfalso; [ Lia.lia | pose proof bitSlice_range_ex (@kunsigned 32 kinst) 20 25 as HR; change (bitSlice (@kunsigned 32 kinst) 20 25) with rs2 in HR; - Lia.lia ]). + Lia.lia ]. (** Consistency proof for each instruction *) all: rt. @@ -2693,7 +2708,7 @@ Section Equiv. destruct H as [? [[? ?] ?]]; discriminate end. 4: { exfalso. - subst v simm12. + subst v' simm12. destruct H5 as [? [? ?]]. regs_get_red H5. apply is_mmio_sound in H5. @@ -2711,7 +2726,7 @@ Section Equiv. all: prove_states_related. (* -- prove [RiscvXAddrsSafe] after store *) - all: subst v simm12 rs1 v0. + all: subst v2' simm12 rs1 v'. all: regs_get_red_goal; regs_get_red Hnmem. all: cbv [Memory.store_bytes] in Hnmem; @@ -2740,7 +2755,7 @@ Section Equiv. HList.tuple.option_all HList.tuple.map HList.tuple.unfoldn Memory.footprint PrimitivePair.pair._1 PrimitivePair.pair._2] in Hlv. repeat (destruct_one_match_hyp; [|discriminate]). - erewrite H12 in E1. + erewrite H12 in E2. destruct_one_match_hyp; [|discriminate]. repeat apply RiscvXAddrsSafe_removeXAddr_write_ok; assumption. } @@ -2751,7 +2766,7 @@ Section Equiv. HList.tuple.option_all HList.tuple.map HList.tuple.unfoldn Memory.footprint PrimitivePair.pair._1 PrimitivePair.pair._2] in Hlv. repeat (destruct_one_match_hyp; [|discriminate]). - erewrite H12 in E1, E3, E5. + erewrite H12 in E2, E4, E6. repeat (destruct_one_match_hyp; [|discriminate]). repeat apply RiscvXAddrsSafe_removeXAddr_write_ok; assumption. } @@ -2846,6 +2861,7 @@ Section Equiv. (** Evaluate (invert) the two fetchers *) rt. eval_kami_fetch. rt. + subst inst'. rewrite H11 in *. (** Symbolic evaluation of Kami decode/execute *) kami_cbn_all. @@ -2992,7 +3008,7 @@ Section Equiv. { (* [pc_related_and_valid] for `JAL` *) subst newPC jimm20. split. { - apply AddrAligned_consistent. rewrite E. reflexivity. + apply AddrAligned_consistent. rewrite E0. reflexivity. } clear; red. cbv [Utility.add @@ -3005,7 +3021,7 @@ Section Equiv. } { (* [pc_related_and_valid] for `JALR` *) - subst newPC oimm12 v rs1. + subst newPC oimm12 v0 rs1. split. { apply AddrAligned_consistent. rewrite E. reflexivity. } @@ -3194,6 +3210,7 @@ Section Equiv. (** Evaluate (invert) the two fetchers *) rt. eval_kami_fetch. rt. + subst inst'. rewrite H11 in *. (** Symbolic evaluation of Kami decode/execute *) kami_cbn_all. @@ -3356,7 +3373,7 @@ Section Equiv. { (* jal *) subst newPC jimm20. split. { - apply AddrAligned_consistent. rewrite E. reflexivity. + apply AddrAligned_consistent. rewrite E0. reflexivity. } clear; red. cbv [Utility.add @@ -3369,7 +3386,7 @@ Section Equiv. } { (* jalr *) - subst newPC oimm12 v rs1. + subst newPC oimm12 v0 rs1. split. { apply AddrAligned_consistent. rewrite E. reflexivity. } @@ -3385,7 +3402,7 @@ Section Equiv. { (* beq(eq) *) subst newPC sbimm12. split. { - apply AddrAligned_consistent. rewrite E0. reflexivity. + apply AddrAligned_consistent. rewrite E1. reflexivity. } clear; red. cbv [Utility.add @@ -3398,28 +3415,28 @@ Section Equiv. } { (* beq(eq-neq contradiction) *) - exfalso; subst v v0 rs1 rs2. - regs_get_red E. + exfalso; subst v v0 v1 v2 rs1 rs2. + regs_get_red E0. apply N2Z.inj, wordToN_inj in e1; auto. } { (* beq(eq-neq contradiction) *) - exfalso; subst v v0 rs1 rs2. - regs_get_red E; congruence. + exfalso; subst v v0 v1 v2 rs1 rs2. + regs_get_red E0; congruence. } { (* bne(neq) *) match goal with | [ |- context [Z.eqb ?x ?y] ] => destruct (Z.eqb_spec x y) end. - { exfalso; subst v v0 rs1 rs2. - regs_get_red E. + { exfalso; subst v v0 v1 v2 rs1 rs2. + regs_get_red E0. apply N2Z.inj, wordToN_inj in e1; auto. } { cbv [negb]. subst addr sbimm12. split. { - apply AddrAligned_consistent. rewrite E0. reflexivity. + apply AddrAligned_consistent. rewrite E1. reflexivity. } clear; red. cbv [Utility.add @@ -3437,20 +3454,20 @@ Section Equiv. | [ |- context [Z.eqb ?x ?y] ] => destruct (Z.eqb_spec x y) end. { apply pc_related_plus4; red; eauto. } - { exfalso; subst v v0 rs1 rs2. - regs_get_red E. + { exfalso; subst v v0 v1 v2 rs1 rs2. + regs_get_red E0. congruence. } } { (* blt(lt) *) cbv [evalBinBitBool]. - subst v v0 rs1 rs2. - regs_get_red E. - destruct (wslt_dec _ _); [|exfalso; apply n; apply E]. + subst v v0 v1 v2 rs1 rs2. + regs_get_red E0. + destruct (wslt_dec _ _); [|exfalso; apply n; apply E0]. subst addr sbimm12. split. { - apply AddrAligned_consistent. rewrite E0. reflexivity. + apply AddrAligned_consistent. rewrite E1. reflexivity. } clear; red. cbv [Utility.add @@ -3464,22 +3481,22 @@ Section Equiv. { (* blt(not lt) *) cbv [evalBinBitBool]. - subst v v0 rs1 rs2. - regs_get_red E. + subst v v0 v1 v2 rs1 rs2. + regs_get_red E0. destruct (wslt_dec _ _). - { exfalso. eapply Z.le_ngt in E. apply E. apply w. } + { exfalso. eapply Z.le_ngt in E0. apply E0. apply w. } apply pc_related_plus4; red; eauto. } { (* bge(ge) *) cbv [evalBinBitBool]. - subst v v0 rs1 rs2. - regs_get_red E. + subst v v0 v1 v2 rs1 rs2. + regs_get_red E0. destruct (wslt_dec _ _). - { exfalso. eapply Z.le_ngt in E. apply E. apply w. } + { exfalso. eapply Z.le_ngt in E0. apply E0. apply w. } subst addr sbimm12. split. { - apply AddrAligned_consistent. rewrite E0. reflexivity. + apply AddrAligned_consistent. rewrite E1. reflexivity. } clear; red. cbv [negb Utility.add @@ -3493,29 +3510,29 @@ Section Equiv. { (* bge(not ge) *) cbv [evalBinBitBool]. - subst v v0 rs1 rs2. - regs_get_red E. - destruct (wslt_dec _ _); [|exfalso; apply n; apply E]. + subst v v0 v1 v2 rs1 rs2. + regs_get_red E0. + destruct (wslt_dec _ _); [|exfalso; apply n; apply E0]. apply pc_related_plus4; red; eauto. } { (* bltu(ltu) *) cbv [evalBinBitBool]. - subst v v0 rs1 rs2. - regs_get_red E. + subst v v0 v1 v2 rs1 rs2. + regs_get_red E0. destruct (wlt_dec _ _). 2: { exfalso. - lazymatch type of E with + lazymatch type of E0 with | word.unsigned ?x < word.unsigned ?y => - change (Z.of_N (wordToN x) < Z.of_N (wordToN y)) in E + change (Z.of_N (wordToN x) < Z.of_N (wordToN y)) in E0 end. - eapply N2Z.inj_lt in E. - apply n. apply E. + eapply N2Z.inj_lt in E0. + apply n. apply E0. } subst addr sbimm12. split. { - apply AddrAligned_consistent. rewrite E0. reflexivity. + apply AddrAligned_consistent. rewrite E1. reflexivity. } clear; red. cbv [Utility.add @@ -3529,34 +3546,34 @@ Section Equiv. { (* bltu(not ltu) *) cbv [evalBinBitBool]. - subst v v0 rs1 rs2. - regs_get_red E. + subst v v0 v1 v2 rs1 rs2. + regs_get_red E0. destruct (wlt_dec _ _). { exfalso. - lazymatch type of E with + lazymatch type of E0 with | word.unsigned ?x <= word.unsigned ?y => - change (Z.of_N (wordToN x) <= Z.of_N (wordToN y)) in E + change (Z.of_N (wordToN x) <= Z.of_N (wordToN y)) in E0 end. - eapply N2Z.inj_le in E. eapply N.lt_nge in w. apply w. apply E. + eapply N2Z.inj_le in E0. eapply N.lt_nge in w. apply w. apply E0. } apply pc_related_plus4; red; eauto. } { (* bgeu(geu) *) cbv [evalBinBitBool]. - subst v v0 rs1 rs2. - regs_get_red E. + subst v v0 v1 v2 rs1 rs2. + regs_get_red E0. destruct (wlt_dec _ _). { exfalso. - lazymatch type of E with + lazymatch type of E0 with | word.unsigned ?x <= word.unsigned ?y => - change (Z.of_N (wordToN x) <= Z.of_N (wordToN y)) in E + change (Z.of_N (wordToN x) <= Z.of_N (wordToN y)) in E0 end. - eapply N2Z.inj_le in E. eapply N.lt_nge in w. apply w. apply E. + eapply N2Z.inj_le in E0. eapply N.lt_nge in w. apply w. apply E0. } subst addr sbimm12. split. { - apply AddrAligned_consistent. rewrite E0. reflexivity. + apply AddrAligned_consistent. rewrite E1. reflexivity. } clear; red. cbv [negb Utility.add @@ -3570,10 +3587,10 @@ Section Equiv. { (* bgeu(not geu) *) cbv [evalBinBitBool]. - subst v v0 rs1 rs2. - regs_get_red E. + subst v v0 v1 v2 rs1 rs2. + regs_get_red E0. destruct (wlt_dec _ _). - 2: { exfalso. apply n. eapply N2Z.inj_lt. apply E. } + 2: { exfalso. apply n. eapply N2Z.inj_lt. apply E0. } apply pc_related_plus4; red; eauto. }