aboutsummaryrefslogtreecommitdiffstats
path: root/scheduling/BTL_SEsimuref.v
diff options
context:
space:
mode:
authorSylvain Boulmé <sylvain.boulme@univ-grenoble-alpes.fr>2021-06-09 06:56:31 +0200
committerSylvain Boulmé <sylvain.boulme@univ-grenoble-alpes.fr>2021-06-09 18:18:55 +0200
commit1b1e5cfc70899d759ed099832b8821372e024ad0 (patch)
treedc4b6936d4f08450a3469f73ecdb92c5c7eca66c /scheduling/BTL_SEsimuref.v
parenta3f8b6fc7931e73ca005644124597ff6656cbce7 (diff)
downloadcompcert-kvx-1b1e5cfc70899d759ed099832b8821372e024ad0.tar.gz
compcert-kvx-1b1e5cfc70899d759ed099832b8821372e024ad0.zip
valid_pointer_preserv dans BTL_SEtheory => opportunites de *grosses* simplifications (à venir )
Diffstat (limited to 'scheduling/BTL_SEsimuref.v')
-rw-r--r--scheduling/BTL_SEsimuref.v79
1 files changed, 13 insertions, 66 deletions
diff --git a/scheduling/BTL_SEsimuref.v b/scheduling/BTL_SEsimuref.v
index 5098659d..40c28934 100644
--- a/scheduling/BTL_SEsimuref.v
+++ b/scheduling/BTL_SEsimuref.v
@@ -67,11 +67,9 @@ Local Hint Constructors ris_ok: core.
Record ris_refines ctx (ris: ristate) (sis: sistate): Prop := {
OK_EQUIV: sis_ok ctx sis <-> ris_ok ctx ris;
MEM_EQ: ris_ok ctx ris -> eval_smem ctx ris.(ris_smem) = eval_smem ctx sis.(si_smem);
- REG_EQ: ris_ok ctx ris -> forall r, eval_sval ctx (ris_sreg_get ris r) = eval_sval ctx (si_sreg sis r);
- (* the below invariant allows to evaluate operations in the initial memory of the path instead of the current memory *)
- VALID_PRESERV: forall m b ofs, eval_smem ctx sis.(si_smem) = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer (cm0 ctx) b ofs
+ REG_EQ: ris_ok ctx ris -> forall r, eval_sval ctx (ris_sreg_get ris r) = eval_sval ctx (si_sreg sis r)
}.
-Local Hint Resolve OK_EQUIV MEM_EQ REG_EQ VALID_PRESERV: core.
+Local Hint Resolve OK_EQUIV MEM_EQ REG_EQ: core.
Local Hint Constructors ris_refines: core.
Record ris_simu ris1 ris2: Prop := {
@@ -218,11 +216,11 @@ Inductive rst_refines ctx: rstate -> sstate -> Prop :=
(RIS: ris_refines ctx ris sis)
(RFV: ris_ok ctx ris -> rfv_refines ctx rfv sfv)
: rst_refines ctx (Rfinal ris rfv) (Sfinal sis sfv)
- | Refcond cond rargs args sm rifso rifnot ifso ifnot
- (RCOND: seval_condition ctx cond rargs Sinit = seval_condition ctx cond args sm)
- (REFso: seval_condition ctx cond rargs Sinit = Some true -> rst_refines ctx rifso ifso)
- (REFnot: seval_condition ctx cond rargs Sinit = Some false -> rst_refines ctx rifnot ifnot)
- : rst_refines ctx (Rcond cond rargs rifso rifnot) (Scond cond args sm ifso ifnot)
+ | Refcond cond rargs args rifso rifnot ifso ifnot
+ (RCOND: seval_condition ctx cond rargs = seval_condition ctx cond args)
+ (REFso: seval_condition ctx cond rargs = Some true -> rst_refines ctx rifso ifso)
+ (REFnot: seval_condition ctx cond rargs = Some false -> rst_refines ctx rifnot ifnot)
+ : rst_refines ctx (Rcond cond rargs rifso rifnot) (Scond cond args ifso ifnot)
| Refabort
: rst_refines ctx Rabort Sabort
.
@@ -245,7 +243,7 @@ Proof.
intros ROK1.
exploit ris_simu_ok_preserv; eauto.
- (* cond_simu *)
- destruct (seval_condition (bctx1 ctx) cond args sm) eqn: SCOND; try congruence.
+ destruct (seval_condition (bctx1 ctx) cond args) eqn: SCOND; try congruence.
generalize RCOND0.
erewrite <- seval_condition_preserved, RCOND by eauto.
intros SCOND0; rewrite <- SCOND0 in RCOND0.
@@ -255,34 +253,6 @@ Proof.
* exploit IHrst_simu2; eauto.
Qed.
-
-(* TODO: useless ?
-Record routcome := rout {
- r_sis: ristate;
- r_sfv: sfval;
-}.
-
-Local Open Scope option_monad_scope.
-
-Fixpoint run_sem_irstate ctx (rst:rstate): option routcome :=
- match rst with
- | Rfinal ris sfv => Some (rout ris sfv)
- | Rcond cond args ifso ifnot =>
- SOME b <- seval_condition ctx cond args Sinit IN
- run_sem_irstate ctx (if b then ifso else ifnot)
- | Rabort => None
- end.
-
-(* Non: pas assez de "matching" syntaxique entre rst et st pour rst_simu_correct
-Definition rst_refines ctx rst st:=
- (run_sem_isstate ctx st=None <-> run_sem_irstate ctx rst=None)
- /\ (forall ris rfv sis sfv, run_sem_irstate ctx rst = Some (rout ris rfv) -> run_sem_isstate ctx st = Some (sout sis sfv) ->
- (ris_refines ctx ris sis) /\ (ris_ok ctx ris -> rfv_refines ctx rfv sfv)).
-
-*)
-*)
-
-
(** * Refinement of the symbolic execution *)
Local Hint Constructors rfv_refines optrsv_refines rsvident_refines rsvident_refines: core.
@@ -341,7 +311,6 @@ Proof.
congruence.
+ destruct 1; simpl in *. unfold ris_sreg_get; simpl.
intros; rewrite PTree.gempty; eauto.
- + try_simplify_someHyps.
Qed.
Definition rset_smem rm (ris:ristate): ristate :=
@@ -369,12 +338,11 @@ Qed.
Lemma rset_mem_correct ctx rm sm ris sis:
(eval_smem ctx ris.(ris_smem) = None -> eval_smem ctx rm = None) ->
- (forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer (cm0 ctx) b ofs) ->
ris_refines ctx ris sis ->
(ris_ok ctx ris -> eval_smem ctx rm = eval_smem ctx sm) ->
ris_refines ctx (rset_smem rm ris) (set_smem sm sis).
Proof.
- destruct 3; intros.
+ destruct 2; intros.
econstructor; eauto.
+ rewrite ok_set_mem, ok_rset_mem; intuition congruence.
+ rewrite ok_rset_mem; intuition eauto.
@@ -393,9 +361,6 @@ Lemma rexec_store_correct ctx chunk addr args src ris sis:
Proof.
intros REF; eapply rset_mem_correct; simpl; eauto.
+ intros X; rewrite X. repeat autodestruct; eauto.
- + intros m b ofs; repeat autodestruct.
- intros; erewrite <- Mem.storev_preserv_valid; [| eauto].
- eauto.
+ intros OK; erewrite eval_list_sval_refpreserv, MEM_EQ, REG_EQ; eauto.
Qed.
@@ -448,7 +413,7 @@ Qed.
Definition rexec_op op args dst (ris:ristate): ristate :=
let args := list_sval_inj (List.map ris args) in
- rset_sreg dst (Sop op args Sinit) ris.
+ rset_sreg dst (Sop op args) ris.
Lemma rexec_op_correct ctx op args dst ris sis:
ris_refines ctx ris sis ->
@@ -456,10 +421,6 @@ Lemma rexec_op_correct ctx op args dst ris sis:
Proof.
intros REF; eapply rset_reg_correct; simpl; eauto.
intros OK; erewrite eval_list_sval_refpreserv; eauto.
- do 2 autodestruct; auto.
- + intros. erewrite <- op_valid_pointer_eq; eauto.
- + erewrite <- MEM_EQ; eauto.
- intros; exploit OK_RMEM; eauto. destruct 1.
Qed.
Definition rexec_load trap chunk addr args dst (ris:ristate): ristate :=
@@ -474,24 +435,10 @@ Proof.
intros OK; erewrite eval_list_sval_refpreserv, MEM_EQ; eauto.
Qed.
-Lemma seval_condition_valid_preserv ctx cond args sm
- (OK: eval_smem ctx sm <> None)
- (VALID_PRESERV: forall m b ofs, eval_smem ctx sm = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer (cm0 ctx) b ofs)
- :seval_condition ctx cond args sm = seval_condition ctx cond args Sinit.
-Proof.
- unfold seval_condition; autodestruct; simpl; try congruence.
- intros EVAL_LIST.
- autodestruct; try congruence.
- intros.
- eapply cond_valid_pointer_eq; intros.
- exploit VALID_PRESERV; eauto.
-Qed.
-
Lemma seval_condition_refpreserv ctx cond args ris sis:
ris_refines ctx ris sis ->
ris_ok ctx ris ->
- seval_condition ctx cond (list_sval_inj (map ris args)) Sinit =
- seval_condition ctx cond (list_sval_inj (map sis args)) Sinit.
+ seval_condition ctx cond (list_sval_inj (map ris args)) = seval_condition ctx cond (list_sval_inj (map sis args)).
Proof.
intros; unfold seval_condition.
erewrite eval_list_sval_refpreserv; eauto.
@@ -654,7 +601,7 @@ Fixpoint history ctx ib sis (k:sistate -> option (list sistate)): option (list s
history ctx ib1 sis (fun sis2 => history ctx ib2 sis2 k)
| Bcond cond args ifso ifnot _ =>
let args := list_sval_inj (List.map sis args) in
- SOME b <- seval_condition ctx cond args sis.(si_smem) IN
+ SOME b <- seval_condition ctx cond args IN
SOME oks <- history ctx (if b then ifso else ifnot) sis k IN
Some (sis::oks)
end
@@ -846,7 +793,7 @@ Proof.
exploit (OK2 sis); auto.
assert (rOK: ris_ok ctx ris). { erewrite <- OK_EQUIV; eauto. }
generalize COND.
- erewrite seval_condition_valid_preserv, <- seval_condition_refpreserv; eauto.
+ erewrite <- seval_condition_refpreserv; eauto.
intros;
econstructor; try_simplify_someHyps.
Qed.