From 1b1e5cfc70899d759ed099832b8821372e024ad0 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Wed, 9 Jun 2021 06:56:31 +0200 Subject: valid_pointer_preserv dans BTL_SEtheory => opportunites de *grosses* simplifications (à venir ) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- scheduling/BTL_SEtheory.v | 50 +++++++++++++++++++++++++++-------------------- 1 file changed, 29 insertions(+), 21 deletions(-) (limited to 'scheduling/BTL_SEtheory.v') diff --git a/scheduling/BTL_SEtheory.v b/scheduling/BTL_SEtheory.v index 464bb0f0..97dc7fc8 100644 --- a/scheduling/BTL_SEtheory.v +++ b/scheduling/BTL_SEtheory.v @@ -29,7 +29,7 @@ Record iblock_exec_context := Bctx { Inductive sval := | Sundef | Sinput (r: reg) - | Sop (op:operation) (lsv: list_sval) (sm: smem) + | Sop (op:operation) (lsv: list_sval) | Sload (sm: smem) (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) with list_sval := | Snil @@ -55,10 +55,9 @@ Fixpoint eval_sval ctx (sv: sval): option val := match sv with | Sundef => Some Vundef | Sinput r => Some ((crs0 ctx)#r) - | Sop op l sm => + | Sop op l => SOME args <- eval_list_sval ctx l IN - SOME m <- eval_smem ctx sm IN - eval_operation (cge ctx) (csp ctx) op args m + eval_operation (cge ctx) (csp ctx) op args (cm0 ctx) | Sload sm trap chunk addr lsv => match trap with | TRAP => @@ -98,6 +97,16 @@ with eval_smem ctx (sm: smem): option mem := Mem.storev chunk m a sv end. + +Lemma valid_pointer_preserv ctx sm: + forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer (cm0 ctx) b ofs = Mem.valid_pointer m b ofs. +Proof. + induction sm; simpl; intros; try_simplify_someHyps; auto. + repeat autodestruct; intros; erewrite IHsm by reflexivity. + eapply Mem.storev_preserv_valid; eauto. +Qed. +Local Hint Resolve valid_pointer_preserv: core. + Lemma eval_list_sval_inj ctx l (sreg: reg -> sval) rs: (forall r : reg, eval_sval ctx (sreg r) = Some (rs # r)) -> eval_list_sval ctx (list_sval_inj (map sreg l)) = Some (rs ## l). @@ -105,10 +114,9 @@ Proof. intros H; induction l as [|r l]; simpl; repeat autodestruct; auto. Qed. -Definition seval_condition ctx (cond: condition) (lsv: list_sval) (sm: smem) : option bool := +Definition seval_condition ctx (cond: condition) (lsv: list_sval): option bool := SOME args <- eval_list_sval ctx lsv IN - SOME m <- eval_smem ctx sm IN - eval_condition cond args m. + eval_condition cond args (cm0 ctx). (** * Auxiliary definitions on Builtins *) @@ -565,7 +573,7 @@ Qed. Definition sexec_op op args dst sis: sistate := let args := list_sval_inj (List.map sis.(si_sreg) args) in - set_sreg dst (Sop op args sis.(si_smem)) sis. + set_sreg dst (Sop op args) sis. Lemma sexec_op_correct ctx op args dst sis rs m v (EVAL: eval_operation (cge ctx) (csp ctx) op rs ## args m = Some v) @@ -577,6 +585,7 @@ Proof. rewrite Regmap.gss; simpl; auto. erewrite eval_list_sval_inj; simpl; auto. try_simplify_someHyps. + intros; erewrite op_valid_pointer_eq; eauto. - intros; rewrite Regmap.gso; auto. Qed. @@ -619,11 +628,11 @@ Qed. Lemma seval_condition_eq ctx cond args sis rs m (SIS : sem_sistate ctx sis rs m) - :seval_condition ctx cond (list_sval_inj (map (si_sreg sis) args)) (si_smem sis) = eval_condition cond rs ## args m. + :seval_condition ctx cond (list_sval_inj (map (si_sreg sis) args)) = eval_condition cond rs ## args m. Proof. destruct SIS as (PRE&MEM®); unfold seval_condition; simpl. erewrite eval_list_sval_inj; simpl; auto. - erewrite MEM; auto. + eapply cond_valid_pointer_eq; eauto. Qed. (** * Compute sistate associated to final values *) @@ -698,7 +707,7 @@ Qed. (* symbolic state *) Inductive sstate := | Sfinal (sis: sistate) (sfv: sfval) - | Scond (cond: condition) (args: list_sval) (sm: smem) (ifso ifnot: sstate) + | Scond (cond: condition) (args: list_sval) (ifso ifnot: sstate) | Sabort . @@ -711,8 +720,8 @@ Record soutcome := sout { Fixpoint run_sem_isstate ctx (st:sstate): option soutcome := match st with | Sfinal sis sfv => Some (sout sis sfv) - | Scond cond args sm ifso ifnot => - SOME b <- seval_condition ctx cond args sm IN + | Scond cond args ifso ifnot => + SOME b <- seval_condition ctx cond args IN run_sem_isstate ctx (if b then ifso else ifnot) | Sabort => None end. @@ -723,10 +732,10 @@ Inductive sem_sstate ctx stk t s: sstate -> Prop := (SIS: sem_sistate ctx sis (str_regs (cf ctx) sfv rs) m) (SFV: sem_sfval ctx stk sfv rs m t s) : sem_sstate ctx stk t s (Sfinal sis sfv) - | sem_Scond b cond args sm ifso ifnot - (SEVAL: seval_condition ctx cond args sm = Some b) + | sem_Scond b cond args ifso ifnot + (SEVAL: seval_condition ctx cond args = Some b) (SELECT: sem_sstate ctx stk t s (if b then ifso else ifnot)) - : sem_sstate ctx stk t s (Scond cond args sm ifso ifnot) + : sem_sstate ctx stk t s (Scond cond args ifso ifnot) (* NB: Sabort: fails to produce a transition *) . @@ -784,7 +793,7 @@ Fixpoint sexec_rec f ib sis (k: sistate -> sstate): sstate := let args := list_sval_inj (List.map sis.(si_sreg) args) in let ifso := sexec_rec f ifso sis k in let ifnot := sexec_rec f ifnot sis k in - Scond cond args sis.(si_smem) ifso ifnot + Scond cond args ifso ifnot end . @@ -954,6 +963,7 @@ Proof. simpl. destruct SIS as (PRE&MEM®). erewrite eval_list_sval_inj; simpl; auto. try_simplify_someHyps. + intros; erewrite op_valid_pointer_eq; eauto. Qed. Lemma sexec_load_TRAP_abort ctx chunk addr args dst sis rs m @@ -1177,7 +1187,6 @@ Proof. induction sv using sval_mut with (P0 := fun lsv => eval_list_sval (bctx1 ctx) lsv = eval_list_sval (bctx2 ctx) lsv) (P1 := fun sm => eval_smem (bctx1 ctx) sm = eval_smem (bctx2 ctx) sm); simpl; auto. + rewrite IHsv; clear IHsv. destruct (eval_list_sval _ _); auto. - rewrite IHsv0; clear IHsv0. destruct (eval_smem _ _); auto. erewrite eval_operation_preserved; eauto. + rewrite IHsv0; clear IHsv0. destruct (eval_list_sval _ _); auto. erewrite eval_addressing_preserved; eauto. @@ -1228,12 +1237,11 @@ Proof. reflexivity. Qed. -Lemma seval_condition_preserved cond lsv sm: - seval_condition (bctx1 ctx) cond lsv sm = seval_condition (bctx2 ctx) cond lsv sm. +Lemma seval_condition_preserved cond lsv: + seval_condition (bctx1 ctx) cond lsv = seval_condition (bctx2 ctx) cond lsv. Proof. unfold seval_condition. rewrite list_sval_eval_preserved. destruct (eval_list_sval _ _); auto. - rewrite smem_eval_preserved; auto. Qed. (* TODO: useless ? -- cgit