aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
-rw-r--r--src/hls/HTLPargen.v2
-rw-r--r--src/hls/IfConversion.v2
-rw-r--r--src/hls/RTLBlock.v15
-rw-r--r--src/hls/RTLBlockInstr.v120
-rw-r--r--src/hls/RTLPar.v15
-rw-r--r--src/hls/RTLPargen.v1139
-rw-r--r--src/hls/RTLPargenproof.v161
7 files changed, 1239 insertions, 215 deletions
diff --git a/src/hls/HTLPargen.v b/src/hls/HTLPargen.v
index 9e1f156..40d1dcc 100644
--- a/src/hls/HTLPargen.v
+++ b/src/hls/HTLPargen.v
@@ -659,7 +659,7 @@ Abort.
Fixpoint pred_expr (preg: reg) (p: pred_op) :=
match p with
| Pvar pred =>
- Vrange preg (Vlit (natToValue pred)) (Vlit (natToValue pred))
+ Vrange preg (Vlit (posToValue pred)) (Vlit (posToValue pred))
| Pnot pred =>
Vunop Vnot (pred_expr preg pred)
| Pand p1 p2 =>
diff --git a/src/hls/IfConversion.v b/src/hls/IfConversion.v
index 39d9fd2..f8d404c 100644
--- a/src/hls/IfConversion.v
+++ b/src/hls/IfConversion.v
@@ -106,7 +106,7 @@ Definition find_blocks_with_cond (c: code) : list (node * bblock) :=
Definition if_convert_code (p: nat * code) (nb: node * bblock) :=
let (n, bb) := nb in
let (p', c) := p in
- let nbb := if_convert_block c p' bb in
+ let nbb := if_convert_block c (Pos.of_nat p') bb in
(S p', PTree.set n nbb c).
Definition transf_function (f: function) : function :=
diff --git a/src/hls/RTLBlock.v b/src/hls/RTLBlock.v
index 6a3487a..aaa3c6c 100644
--- a/src/hls/RTLBlock.v
+++ b/src/hls/RTLBlock.v
@@ -58,11 +58,11 @@ Section RELSEM.
Inductive step: state -> trace -> state -> Prop :=
| exec_bblock:
- forall s f sp pc rs rs' m m' t s' bb,
+ forall s f sp pc rs rs' m m' t s' bb pr pr',
f.(fn_code)!pc = Some bb ->
- step_instr_list sp (InstrState rs m) bb.(bb_body) (InstrState rs' m') ->
- step_cf_instr ge (State s f sp pc rs' m') bb.(bb_exit) t s' ->
- step (State s f sp pc rs m) t s'
+ step_instr_list sp (InstrState rs pr m) bb.(bb_body) (InstrState rs' pr' m') ->
+ step_cf_instr ge (State s f sp pc rs' pr' m') bb.(bb_exit) t s' ->
+ step (State s f sp pc rs pr m) t s'
| exec_function_internal:
forall s f args m m' stk,
Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
@@ -72,6 +72,7 @@ Section RELSEM.
(Vptr stk Ptrofs.zero)
f.(fn_entrypoint)
(init_regs args f.(fn_params))
+ (PMap.init false)
m')
| exec_function_external:
forall s ef args res t m m',
@@ -79,9 +80,9 @@ Section RELSEM.
step (Callstate s (External ef) args m)
t (Returnstate s res m')
| exec_return:
- forall res f sp pc rs s vres m,
- step (Returnstate (Stackframe res f sp pc rs :: s) vres m)
- E0 (State s f sp pc (rs#res <- vres) m).
+ forall res f sp pc rs s vres m pr,
+ step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m)
+ E0 (State s f sp pc (rs#res <- vres) pr m).
End RELSEM.
diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v
index 3fab464..79e3149 100644
--- a/src/hls/RTLBlockInstr.v
+++ b/src/hls/RTLBlockInstr.v
@@ -34,7 +34,7 @@ Require Import vericert.hls.Sat.
Local Open Scope rtl.
Definition node := positive.
-Definition predicate := nat.
+Definition predicate := positive.
Inductive pred_op : Type :=
| Pvar: predicate -> pred_op
@@ -44,7 +44,7 @@ Inductive pred_op : Type :=
Fixpoint sat_predicate (p: pred_op) (a: asgn) : bool :=
match p with
- | Pvar p' => a p'
+ | Pvar p' => a (Pos.to_nat p')
| Pnot p' => negb (sat_predicate p' a)
| Pand p1 p2 => sat_predicate p1 a && sat_predicate p2 a
| Por p1 p2 => sat_predicate p1 a || sat_predicate p2 a
@@ -152,7 +152,7 @@ Fixpoint trans_pred_temp (bound: nat) (p: pred_op) : option formula :=
| O => None
| S n =>
match p with
- | Pvar p' => Some (((true, p') :: nil) :: nil)
+ | Pvar p' => Some (((true, Pos.to_nat p') :: nil) :: nil)
| Pand p1 p2 =>
match trans_pred_temp n p1, trans_pred_temp n p2 with
| Some p1', Some p2' =>
@@ -165,7 +165,7 @@ Fixpoint trans_pred_temp (bound: nat) (p: pred_op) : option formula :=
Some (mult p1' p2')
| _, _ => None
end
- | Pnot (Pvar p') => Some (((false, p') :: nil) :: nil)
+ | Pnot (Pvar p') => Some (((false, Pos.to_nat p') :: nil) :: nil)
| Pnot (Pnot p) => trans_pred_temp n p
| Pnot (Pand p1 p2) => trans_pred_temp n (Por (Pnot p1) (Pnot p2))
| Pnot (Por p1 p2) => trans_pred_temp n (Pand (Pnot p1) (Pnot p2))
@@ -180,7 +180,7 @@ Fixpoint trans_pred (bound: nat) (p: pred_op) :
| O => None
| S n =>
match p with
- | Pvar p' => Some (exist _ (((true, p') :: nil) :: nil) _)
+ | Pvar p' => Some (exist _ (((true, Pos.to_nat p') :: nil) :: nil) _)
| Pand p1 p2 =>
match trans_pred n p1, trans_pred n p2 with
| Some (exist _ p1' _), Some (exist _ p2' _) =>
@@ -331,6 +331,15 @@ Fixpoint max_reg_cfi (m : positive) (i : cf_instr) :=
end.
Definition regset := Regmap.t val.
+Definition predset := PMap.t bool.
+
+Fixpoint eval_predf (pr: predset) (p: pred_op) {struct p} :=
+ match p with
+ | Pvar p' => PMap.get p' pr
+ | Pnot p' => negb (eval_predf pr p')
+ | Pand p' p'' => (eval_predf pr p') && (eval_predf pr p'')
+ | Por p' p'' => (eval_predf pr p') || (eval_predf pr p'')
+ end.
Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset :=
match rl, vl with
@@ -341,6 +350,7 @@ Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset :=
Inductive instr_state : Type :=
| InstrState:
forall (rs: regset)
+ (pr: predset)
(m: mem),
instr_state.
@@ -379,7 +389,8 @@ Section DEFINITION.
(f: function) (**r calling function *)
(sp: val) (**r stack pointer in calling function *)
(pc: node) (**r program point in calling function *)
- (rs: regset), (**r register state in calling function *)
+ (rs: regset) (**r register state in calling function *)
+ (pr: predset), (**r predicate state of the calling function *)
stackframe.
Inductive state : Type :=
@@ -389,6 +400,7 @@ Section DEFINITION.
(sp: val) (**r stack pointer *)
(pc: node) (**r current program point in [c] *)
(rs: regset) (**r register state *)
+ (pr: predset) (**r predicate register state *)
(m: mem), (**r memory state *)
state
| Callstate:
@@ -424,67 +436,89 @@ Section RELSEM.
end
end.
+ Inductive eval_pred: option pred_op -> instr_state -> instr_state -> instr_state -> Prop :=
+ | eval_pred_true:
+ forall (pr: predset) p rs pr m i,
+ eval_predf pr p = true ->
+ eval_pred (Some p) (InstrState rs pr m) i i
+ | eval_pred_false:
+ forall (pr: predset) p rs pr m i,
+ eval_predf pr p = false ->
+ eval_pred (Some p) (InstrState rs pr m) i (InstrState rs pr m)
+ | eval_pred_none:
+ forall i i',
+ eval_pred None i i' i.
+
Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop :=
| exec_RBnop:
- forall rs m sp,
- step_instr sp (InstrState rs m) RBnop (InstrState rs m)
+ forall sp ist,
+ step_instr sp ist RBnop ist
| exec_RBop:
- forall op v res args rs m sp p,
- eval_operation ge sp op rs##args m = Some v ->
- step_instr sp (InstrState rs m)
- (RBop p op args res)
- (InstrState (rs#res <- v) m)
+ forall op v res args rs m sp p ist pr,
+ eval_operation ge sp op rs##args m = Some v ->
+ eval_pred p (InstrState rs pr m) (InstrState (rs#res <- v) pr m) ist ->
+ step_instr sp (InstrState rs pr m) (RBop p op args res) ist
| exec_RBload:
- forall addr rs args a chunk m v dst sp p,
- eval_addressing ge sp addr rs##args = Some a ->
- Mem.loadv chunk m a = Some v ->
- step_instr sp (InstrState rs m)
- (RBload p chunk addr args dst)
- (InstrState (rs#dst <- v) m)
+ forall addr rs args a chunk m v dst sp p pr ist,
+ eval_addressing ge sp addr rs##args = Some a ->
+ Mem.loadv chunk m a = Some v ->
+ eval_pred p (InstrState rs pr m) (InstrState (rs#dst <- v) pr m) ist ->
+ step_instr sp (InstrState rs pr m) (RBload p chunk addr args dst) ist
| exec_RBstore:
- forall addr rs args a chunk m src m' sp p,
- eval_addressing ge sp addr rs##args = Some a ->
- Mem.storev chunk m a rs#src = Some m' ->
- step_instr sp (InstrState rs m)
- (RBstore p chunk addr args src)
- (InstrState rs m').
+ forall addr rs args a chunk m src m' sp p pr ist,
+ eval_addressing ge sp addr rs##args = Some a ->
+ Mem.storev chunk m a rs#src = Some m' ->
+ eval_pred p (InstrState rs pr m) (InstrState rs pr m') ist ->
+ step_instr sp (InstrState rs pr m) (RBstore p chunk addr args src) ist
+ | exec_RBsetpred:
+ forall sp rs pr m p c b args,
+ Op.eval_condition c rs##args m = Some b ->
+ step_instr sp (InstrState rs pr m) (RBsetpred c args p)
+ (InstrState rs (PMap.set p b pr) m).
Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop :=
| exec_RBcall:
- forall s f sp rs m res fd ros sig args pc pc',
+ forall s f sp rs m res fd ros sig args pc pc' pr,
find_function ros rs = Some fd ->
funsig fd = sig ->
- step_cf_instr (State s f sp pc rs m) (RBcall sig ros args res pc')
- E0 (Callstate (Stackframe res f sp pc' rs :: s) fd rs##args m)
+ step_cf_instr (State s f sp pc rs pr m) (RBcall sig ros args res pc')
+ E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m)
| exec_RBtailcall:
- forall s f stk rs m sig ros args fd m' pc,
+ forall s f stk rs m sig ros args fd m' pc pr,
find_function ros rs = Some fd ->
funsig fd = sig ->
Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs m) (RBtailcall sig ros args)
+ step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) (RBtailcall sig ros args)
E0 (Callstate s fd rs##args m')
| exec_RBbuiltin:
- forall s f sp rs m ef args res pc' vargs t vres m' pc,
+ forall s f sp rs m ef args res pc' vargs t vres m' pc pr,
eval_builtin_args ge (fun r => rs#r) sp m args vargs ->
external_call ef ge vargs m t vres m' ->
- step_cf_instr (State s f sp pc rs m) (RBbuiltin ef args res pc')
- t (State s f sp pc' (regmap_setres res vres rs) m')
+ step_cf_instr (State s f sp pc rs pr m) (RBbuiltin ef args res pc')
+ t (State s f sp pc' (regmap_setres res vres rs) pr m')
| exec_RBcond:
- forall s f sp rs m cond args ifso ifnot b pc pc',
+ forall s f sp rs m cond args ifso ifnot b pc pc' pr,
eval_condition cond rs##args m = Some b ->
pc' = (if b then ifso else ifnot) ->
- step_cf_instr (State s f sp pc rs m) (RBcond cond args ifso ifnot)
- E0 (State s f sp pc' rs m)
+ step_cf_instr (State s f sp pc rs pr m) (RBcond cond args ifso ifnot)
+ E0 (State s f sp pc' rs pr m)
| exec_RBjumptable:
- forall s f sp rs m arg tbl n pc pc',
+ forall s f sp rs m arg tbl n pc pc' pr,
rs#arg = Vint n ->
list_nth_z tbl (Int.unsigned n) = Some pc' ->
- step_cf_instr (State s f sp pc rs m) (RBjumptable arg tbl)
- E0 (State s f sp pc' rs m)
- | exec_Ireturn:
- forall s f stk rs m or pc m',
+ step_cf_instr (State s f sp pc rs pr m) (RBjumptable arg tbl)
+ E0 (State s f sp pc' rs pr m)
+ | exec_RBreturn:
+ forall s f stk rs m or pc m' pr,
Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs m) (RBreturn or)
- E0 (Returnstate s (regmap_optget or Vundef rs) m').
+ step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) (RBreturn or)
+ E0 (Returnstate s (regmap_optget or Vundef rs) m')
+ | exec_RBgoto:
+ forall s f sp pc rs pr m pc',
+ step_cf_instr (State s f sp pc rs pr m) (RBgoto pc') E0 (State s f sp pc' rs pr m)
+ | exec_RBpred_cf:
+ forall s f sp pc rs pr m cf1 cf2 st' p t,
+ step_cf_instr (State s f sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' ->
+ step_cf_instr (State s f sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'.
End RELSEM.
diff --git a/src/hls/RTLPar.v b/src/hls/RTLPar.v
index 2e78d36..9d5fc77 100644
--- a/src/hls/RTLPar.v
+++ b/src/hls/RTLPar.v
@@ -80,11 +80,11 @@ Section RELSEM.
Inductive step: state -> trace -> state -> Prop :=
| exec_bblock:
- forall s f sp pc rs rs' m m' t s' bb,
+ forall s f sp pc rs rs' m m' t s' bb pr pr',
f.(fn_code)!pc = Some bb ->
- step_instr_block sp (InstrState rs m) bb.(bb_body) (InstrState rs' m') ->
- step_cf_instr ge (State s f sp pc rs' m') bb.(bb_exit) t s' ->
- step (State s f sp pc rs m) t s'
+ step_instr_block sp (InstrState rs pr m) bb.(bb_body) (InstrState rs' pr' m') ->
+ step_cf_instr ge (State s f sp pc rs' pr' m') bb.(bb_exit) t s' ->
+ step (State s f sp pc rs pr m) t s'
| exec_function_internal:
forall s f args m m' stk,
Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
@@ -94,6 +94,7 @@ Section RELSEM.
(Vptr stk Ptrofs.zero)
f.(fn_entrypoint)
(init_regs args f.(fn_params))
+ (PMap.init false)
m')
| exec_function_external:
forall s ef args res t m m',
@@ -101,9 +102,9 @@ Section RELSEM.
step (Callstate s (External ef) args m)
t (Returnstate s res m')
| exec_return:
- forall res f sp pc rs s vres m,
- step (Returnstate (Stackframe res f sp pc rs :: s) vres m)
- E0 (State s f sp pc (rs#res <- vres) m).
+ forall res f sp pc rs s vres m pr,
+ step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m)
+ E0 (State s f sp pc (rs#res <- vres) pr m).
End RELSEM.
diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v
index 5ad3f90..727ccf3 100644
--- a/src/hls/RTLPargen.v
+++ b/src/hls/RTLPargen.v
@@ -19,7 +19,7 @@
Require Import compcert.backend.Registers.
Require Import compcert.common.AST.
Require Import compcert.common.Globalenvs.
-Require compcert.common.Memory.
+Require Import compcert.common.Memory.
Require Import compcert.common.Values.
Require Import compcert.lib.Floats.
Require Import compcert.lib.Integers.
@@ -44,6 +44,7 @@ Definition reg := positive.
Inductive resource : Set :=
| Reg : reg -> resource
+| Pred : reg -> resource
| Mem : resource.
(*|
@@ -53,7 +54,7 @@ optimised heavily if written manually, as their proofs are not needed.
Lemma resource_eq : forall (r1 r2 : resource), {r1 = r2} + {r1 <> r2}.
Proof.
- decide equality. apply Pos.eq_dec.
+ decide equality; apply Pos.eq_dec.
Defined.
Lemma comparison_eq: forall (x y : comparison), {x = y} + {x <> y}.
@@ -181,7 +182,8 @@ Module R_indexed.
Definition t := resource.
Definition index (rs: resource) : positive :=
match rs with
- | Reg r => xO r
+ | Reg r => xO (xO r)
+ | Pred r => xI (xI r)
| Mem => 1%positive
end.
@@ -205,14 +207,171 @@ Then, to make recursion over expressions easier, expression_list is also defined
that enables mutual recursive definitions over the datatypes.
|*)
-Inductive expression : Set :=
+Definition unsat p := forall a, sat_predicate p a = false.
+Definition sat p := exists a, sat_predicate p a = true.
+
+Inductive expression : Type :=
| Ebase : resource -> expression
-| Eop : Op.operation -> expression_list -> expression
+| Eop : Op.operation -> expression_list -> expression -> expression
| Eload : AST.memory_chunk -> Op.addressing -> expression_list -> expression -> expression
| Estore : expression -> AST.memory_chunk -> Op.addressing -> expression_list -> expression -> expression
-with expression_list : Set :=
+| Esetpred : predicate -> Op.condition -> expression_list -> expression -> expression
+| Econd : expr_pred_list -> expression
+with expression_list : Type :=
| Enil : expression_list
-| Econs : expression -> expression_list -> expression_list.
+| Econs : expression -> expression_list -> expression_list
+with expr_pred_list : Type :=
+| EPnil : expr_pred_list
+| EPcons : pred_op -> expression -> expr_pred_list -> expr_pred_list
+.
+
+Definition pred_list_wf l : Prop :=
+ forall a b, In a l -> In b l -> a <> b -> unsat (Pand a b).
+
+Fixpoint expr_pred_list_to_list e :=
+ match e with
+ | EPnil => nil
+ | EPcons p e l => (p, e) :: expr_pred_list_to_list l
+ end.
+
+Definition pred_list_wf_ep l : Prop :=
+ pred_list_wf (map fst (expr_pred_list_to_list l)).
+
+Lemma unsat_correct1 :
+ forall a b c,
+ unsat (Pand a b) ->
+ sat_predicate a c = true ->
+ sat_predicate b c = false.
+Proof.
+ unfold unsat in *. intros.
+ simplify. specialize (H c).
+ apply andb_false_iff in H. inv H. rewrite H0 in H1. discriminate.
+ auto.
+Qed.
+
+Lemma unsat_correct2 :
+ forall a b c,
+ unsat (Pand a b) ->
+ sat_predicate b c = true ->
+ sat_predicate a c = false.
+Proof.
+ unfold unsat in *. intros.
+ simplify. specialize (H c).
+ apply andb_false_iff in H. inv H. auto. rewrite H0 in H1. discriminate.
+Qed.
+
+Lemma unsat_not a: unsat (Pand a (Pnot a)).
+Proof. unfold unsat; simplify; auto with bool. Qed.
+
+Lemma unsat_commut a b: unsat (Pand a b) -> unsat (Pand b a).
+Proof. unfold unsat; simplify; eauto with bool. Qed.
+
+Lemma sat_dec a n b: sat_pred n a = Some b -> {sat a} + {unsat a}.
+Proof.
+ unfold sat, unsat. destruct b.
+ intros. left. destruct s.
+ exists (Sat.interp_alist x). auto.
+ intros. tauto.
+Qed.
+
+Lemma sat_equiv :
+ forall a b,
+ unsat (Por (Pand a (Pnot b)) (Pand (Pnot a) b)) ->
+ forall c, sat_predicate a c = sat_predicate b c.
+Proof.
+ unfold unsat. intros. specialize (H c); simplify.
+ destruct (sat_predicate b c) eqn:X;
+ destruct (sat_predicate a c) eqn:X2;
+ crush.
+Qed.
+
+(*Parameter op_le : Op.operation -> Op.operation -> bool.
+Parameter chunk_le : AST.memory_chunk -> AST.memory_chunk -> bool.
+Parameter addr_le : Op.addressing -> Op.addressing -> bool.
+Parameter cond_le : Op.condition -> Op.condition -> bool.
+
+Fixpoint pred_le (p1 p2: pred_op) : bool :=
+ match p1, p2 with
+ | Pvar i, Pvar j => (i <=? j)%positive
+ | Pnot p1, Pnot p2 => pred_le p1 p2
+ | Pand p1 p1', Pand p2 p2' => if pred_le p1 p2 then true else pred_le p1' p2'
+ | Por p1 p1', Por p2 p2' => if pred_le p1 p2 then true else pred_le p1' p2'
+ | Pvar _, _ => true
+ | Pnot _, Pvar _ => false
+ | Pnot _, _ => true
+ | Pand _ _, Pvar _ => false
+ | Pand _ _, Pnot _ => false
+ | Pand _ _, _ => true
+ | Por _ _, _ => false
+ end.
+
+Import Lia.
+
+Lemma pred_le_trans :
+ forall p1 p2 p3 b, pred_le p1 p2 = b -> pred_le p2 p3 = b -> pred_le p1 p3 = b.
+Proof.
+ induction p1; destruct p2; destruct p3; crush.
+ destruct b. rewrite Pos.leb_le in *. lia. rewrite Pos.leb_gt in *. lia.
+ firstorder.
+ destruct (pred_le p1_1 p2_1) eqn:?. subst. destruct (pred_le p2_1 p3_1) eqn:?.
+ apply IHp1_1 in Heqb. rewrite Heqb. auto. auto.
+
+
+Fixpoint expr_le (e1 e2: expression) {struct e2}: bool :=
+ match e1, e2 with
+ | Ebase r1, Ebase r2 => (R_indexed.index r1 <=? R_indexed.index r2)%positive
+ | Ebase _, _ => true
+ | Eop op1 elist1 m1, Eop op2 elist2 m2 =>
+ if op_le op1 op2 then true
+ else if elist_le elist1 elist2 then true
+ else expr_le m1 m2
+ | Eop _ _ _, Ebase _ => false
+ | Eop _ _ _, _ => true
+ | Eload chunk1 addr1 elist1 expr1, Eload chunk2 addr2 elist2 expr2 =>
+ if chunk_le chunk1 chunk2 then true
+ else if addr_le addr1 addr2 then true
+ else if elist_le elist1 elist2 then true
+ else expr_le expr1 expr2
+ | Eload _ _ _ _, Ebase _ => false
+ | Eload _ _ _ _, Eop _ _ _ => false
+ | Eload _ _ _ _, _ => true
+ | Estore m1 chunk1 addr1 elist1 expr1, Estore m2 chunk2 addr2 elist2 expr2 =>
+ if expr_le m1 m2 then true
+ else if chunk_le chunk1 chunk2 then true
+ else if addr_le addr1 addr2 then true
+ else if elist_le elist1 elist2 then true
+ else expr_le expr1 expr2
+ | Estore _ _ _ _ _, Ebase _ => false
+ | Estore _ _ _ _ _, Eop _ _ _ => false
+ | Estore _ _ _ _ _, Eload _ _ _ _ => false
+ | Estore _ _ _ _ _, _ => true
+ | Esetpred p1 cond1 elist1 m1, Esetpred p2 cond2 elist2 m2 =>
+ if (p1 <=? p2)%positive then true
+ else if cond_le cond1 cond2 then true
+ else if elist_le elist1 elist2 then true
+ else expr_le m1 m2
+ | Esetpred _ _ _ _, Econd _ => true
+ | Esetpred _ _ _ _, _ => false
+ | Econd eplist1, Econd eplist2 => eplist_le eplist1 eplist2
+ | Econd eplist1, _ => false
+ end
+with elist_le (e1 e2: expression_list) : bool :=
+ match e1, e2 with
+ | Enil, Enil => true
+ | Econs a1 b1, Econs a2 b2 => if expr_le a1 a2 then true else elist_le b1 b2
+ | Enil, _ => true
+ | _, Enil => false
+ end
+with eplist_le (e1 e2: expr_pred_list) : bool :=
+ match e1, e2 with
+ | EPnil, EPnil => true
+ | EPcons p1 a1 b1, EPcons p2 a2 b2 =>
+ if pred_le p1 p2 then true
+ else if expr_le a1 a2 then true else eplist_le b1 b2
+ | EPnil, _ => true
+ | _, EPnil => false
+ end
+.*)
(*|
Using IMap we can create a map from resources to any other type, as resources can be uniquely
@@ -223,8 +382,6 @@ Module Rtree := ITree(R_indexed).
Definition forest : Type := Rtree.t expression.
-Definition regset := Registers.Regmap.t val.
-
Definition get_forest v f :=
match Rtree.get v f with
| None => Ebase v
@@ -234,10 +391,18 @@ Definition get_forest v f :=
Notation "a # b" := (get_forest b a) (at level 1).
Notation "a # b <- c" := (Rtree.set b c a) (at level 1, b at next level).
-Record sem_state := mk_sem_state {
- sem_state_regset : regset;
- sem_state_memory : Memory.mem
- }.
+Definition maybe {A: Type} (vo: A) (pr: predset) p (v: A) :=
+ match p with
+ | Some p' => if eval_predf pr p' then v else vo
+ | None => v
+ end.
+
+Definition get_pr i := match i with InstrState a b c => b end.
+
+Definition get_m i := match i with InstrState a b c => c end.
+
+Definition eval_predf_opt pr p :=
+ match p with Some p' => eval_predf pr p' | None => true end.
(*|
Finally we want to define the semantics of execution for the expressions with symbolic values, so
@@ -246,82 +411,130 @@ the result of executing the expressions will be an expressions.
Section SEMANTICS.
-Context (A : Set) (genv : Genv.t A unit).
+Context {A : Type} (genv : Genv.t A unit).
Inductive sem_value :
- val -> sem_state -> expression -> val -> Prop :=
- | Sbase_reg:
- forall sp st r,
- sem_value sp st (Ebase (Reg r)) (Registers.Regmap.get r (sem_state_regset st))
- | Sop:
- forall st op args v lv sp,
- sem_val_list sp st args lv ->
- Op.eval_operation genv sp op lv (sem_state_memory st) = Some v ->
- sem_value sp st (Eop op args) v
- | Sload :
- forall st mem_exp addr chunk args a v m' lv sp,
- sem_mem sp st mem_exp m' ->
- sem_val_list sp st args lv ->
- Op.eval_addressing genv sp addr lv = Some a ->
- Memory.Mem.loadv chunk m' a = Some v ->
- sem_value sp st (Eload chunk addr args mem_exp) v
+ val -> instr_state -> expression -> val -> Prop :=
+| Sbase_reg:
+ forall sp rs r m pr,
+ sem_value sp (InstrState rs pr m) (Ebase (Reg r)) (rs !! r)
+| Sop:
+ forall rs m op args v lv sp m' mem_exp pr,
+ sem_mem sp (InstrState rs pr m) mem_exp m' ->
+ sem_val_list sp (InstrState rs pr m) args lv ->
+ Op.eval_operation genv sp op lv m' = Some v ->
+ sem_value sp (InstrState rs pr m) (Eop op args mem_exp) v
+| Sload :
+ forall st mem_exp addr chunk args a v m' lv sp,
+ sem_mem sp st mem_exp m' ->
+ sem_val_list sp st args lv ->
+ Op.eval_addressing genv sp addr lv = Some a ->
+ Memory.Mem.loadv chunk m' a = Some v ->
+ sem_value sp st (Eload chunk addr args mem_exp) v
+| Scond :
+ forall sp st e v,
+ sem_val_ep_list sp st e v ->
+ sem_value sp st (Econd e) v
+with sem_pred :
+ val -> instr_state -> expression -> bool -> Prop :=
+| Spred:
+ forall st mem_exp args p c lv m m' v sp,
+ sem_mem sp st mem_exp m' ->
+ sem_val_list sp st args lv ->
+ Op.eval_condition c lv m = Some v ->
+ sem_pred sp st (Esetpred p c args mem_exp) v
+| Sbase_pred:
+ forall rs pr m p sp,
+ sem_pred sp (InstrState rs pr m) (Ebase (Pred p)) (PMap.get p pr)
with sem_mem :
- val -> sem_state -> expression -> Memory.mem -> Prop :=
- | Sstore :
- forall st mem_exp val_exp m'' addr v a m' chunk args lv sp,
- sem_mem sp st mem_exp m' ->
- sem_value sp st val_exp v ->
- sem_val_list sp st args lv ->
- Op.eval_addressing genv sp addr lv = Some a ->
- Memory.Mem.storev chunk m' a v = Some m'' ->
- sem_mem sp st (Estore mem_exp chunk addr args val_exp) m''
- | Sbase_mem :
- forall st m sp,
- sem_mem sp st (Ebase Mem) m
+ val -> instr_state -> expression -> Memory.mem -> Prop :=
+| Sstore :
+ forall st mem_exp val_exp m'' addr v a m' chunk args lv sp,
+ sem_mem sp st mem_exp m' ->
+ sem_value sp st val_exp v ->
+ sem_val_list sp st args lv ->
+ Op.eval_addressing genv sp addr lv = Some a ->
+ Memory.Mem.storev chunk m' a v = Some m'' ->
+ sem_mem sp st (Estore mem_exp chunk addr args val_exp) m''
+| Sbase_mem :
+ forall rs m sp pr,
+ sem_mem sp (InstrState rs pr m) (Ebase Mem) m
with sem_val_list :
- val -> sem_state -> expression_list -> list val -> Prop :=
- | Snil :
- forall st sp,
- sem_val_list sp st Enil nil
- | Scons :
- forall st e v l lv sp,
- sem_value sp st e v ->
- sem_val_list sp st l lv ->
- sem_val_list sp st (Econs e l) (v :: lv).
+ val -> instr_state -> expression_list -> list val -> Prop :=
+| Snil :
+ forall st sp,
+ sem_val_list sp st Enil nil
+| Scons :
+ forall st e v l lv sp,
+ sem_value sp st e v ->
+ sem_val_list sp st l lv ->
+ sem_val_list sp st (Econs e l) (v :: lv)
+with sem_val_ep_list :
+ val -> instr_state -> expr_pred_list -> val -> Prop :=
+| SPnil :
+ forall sp rs r m pr,
+ sem_val_ep_list sp (InstrState rs pr m) EPnil (rs !! r)
+| SPconsTrue :
+ forall pr p sp rs m e v el,
+ eval_predf pr p = true ->
+ sem_value sp (InstrState rs pr m) e v ->
+ sem_val_ep_list sp (InstrState rs pr m) (EPcons p e el) v
+| SPconsFalse :
+ forall pr p sp rs m e v el,
+ eval_predf pr p = false ->
+ sem_val_ep_list sp (InstrState rs pr m) el v ->
+ sem_val_ep_list sp (InstrState rs pr m) (EPcons p e el) v
+.
+
+Inductive sem_predset :
+ val -> instr_state -> forest -> predset -> Prop :=
+| Spredset:
+ forall st f sp rs',
+ (forall x, sem_pred sp st (f # (Pred x)) (PMap.get x rs')) ->
+ sem_predset sp st f rs'.
Inductive sem_regset :
- val -> sem_state -> forest -> regset -> Prop :=
- | Sregset:
- forall st f rs' sp,
- (forall x, sem_value sp st (f # (Reg x)) (Registers.Regmap.get x rs')) ->
- sem_regset sp st f rs'.
+ val -> instr_state -> forest -> regset -> Prop :=
+| Sregset:
+ forall st f sp rs',
+ (forall x, sem_value sp st (f # (Reg x)) (rs' !! x)) ->
+ sem_regset sp st f rs'.
Inductive sem :
- val -> sem_state -> forest -> sem_state -> Prop :=
- | Sem:
- forall st rs' m' f sp,
- sem_regset sp st f rs' ->
- sem_mem sp st (f # Mem) m' ->
- sem sp st f (mk_sem_state rs' m').
+ val -> instr_state -> forest -> instr_state -> Prop :=
+| Sem:
+ forall st rs' m' f sp pr',
+ sem_regset sp st f rs' ->
+ sem_predset sp st f pr' ->
+ sem_mem sp st (f # Mem) m' ->
+ sem sp st f (InstrState rs' pr' m').
End SEMANTICS.
Fixpoint beq_expression (e1 e2: expression) {struct e1}: bool :=
match e1, e2 with
| Ebase r1, Ebase r2 => if resource_eq r1 r2 then true else false
- | Eop op1 el1, Eop op2 el2 =>
- if operation_eq op1 op2 then beq_expression_list el1 el2 else false
+ | Eop op1 el1 exp1, Eop op2 el2 exp2 =>
+ if operation_eq op1 op2 then
+ if beq_expression exp1 exp2 then
+ beq_expression_list el1 el2 else false else false
| Eload chk1 addr1 el1 e1, Eload chk2 addr2 el2 e2 =>
- if memory_chunk_eq chk1 chk2
- then if addressing_eq addr1 addr2
- then if beq_expression_list el1 el2
- then beq_expression e1 e2 else false else false else false
+ if memory_chunk_eq chk1 chk2
+ then if addressing_eq addr1 addr2
+ then if beq_expression_list el1 el2
+ then beq_expression e1 e2 else false else false else false
| Estore m1 chk1 addr1 el1 e1, Estore m2 chk2 addr2 el2 e2=>
- if memory_chunk_eq chk1 chk2
- then if addressing_eq addr1 addr2
- then if beq_expression_list el1 el2
- then if beq_expression m1 m2
- then beq_expression e1 e2 else false else false else false else false
+ if memory_chunk_eq chk1 chk2
+ then if addressing_eq addr1 addr2
+ then if beq_expression_list el1 el2
+ then if beq_expression m1 m2
+ then beq_expression e1 e2 else false else false else false else false
+ | Esetpred p1 c1 el1 m1, Esetpred p2 c2 el2 m2 =>
+ if Pos.eqb p1 p2
+ then if condition_eq c1 c2
+ then if beq_expression_list el1 el2
+ then beq_expression m1 m2 else false else false else false
+ | Econd el1, Econd el2 => beq_expr_pred_list el1 el2
| _, _ => false
end
with beq_expression_list (el1 el2: expression_list) {struct el1} : bool :=
@@ -329,10 +542,19 @@ with beq_expression_list (el1 el2: expression_list) {struct el1} : bool :=
| Enil, Enil => true
| Econs e1 t1, Econs e2 t2 => beq_expression e1 e2 && beq_expression_list t1 t2
| _, _ => false
- end.
+ end
+with beq_expr_pred_list (el1 el2: expr_pred_list) {struct el1} : bool :=
+ match el1, el2 with
+ | EPnil, EPnil => true
+ | EPcons p1 e1 el1', EPcons p2 e2 el2' => true
+ | _, _ => false
+ end
+.
Scheme expression_ind2 := Induction for expression Sort Prop
- with expression_list_ind2 := Induction for expression_list Sort Prop.
+ with expression_list_ind2 := Induction for expression_list Sort Prop
+ with expr_pred_list_ind2 := Induction for expr_pred_list Sort Prop
+.
Lemma beq_expression_correct:
forall e1 e2, beq_expression e1 e2 = true -> e1 = e2.
@@ -342,11 +564,14 @@ Proof.
(P := fun (e1 : expression) =>
forall e2, beq_expression e1 e2 = true -> e1 = e2)
(P0 := fun (e1 : expression_list) =>
- forall e2, beq_expression_list e1 e2 = true -> e1 = e2); simplify;
+ forall e2, beq_expression_list e1 e2 = true -> e1 = e2)
+ (P1 := fun (e1 : expr_pred_list) =>
+ forall e2, beq_expr_pred_list e1 e2 = true -> e1 = e2); simplify;
repeat match goal with
| [ H : context[match ?x with _ => _ end] |- _ ] => destruct x eqn:?
| [ H : context[if ?x then _ else _] |- _ ] => destruct x eqn:?
end; subst; f_equal; crush.
+ eauto using Peqb_true_eq.
Qed.
Definition empty : forest := Rtree.empty _.
@@ -357,7 +582,7 @@ This function checks if all the elements in [fa] are in [fb], but not the other
Definition check := Rtree.beq beq_expression.
-Lemma check_correct: forall (fa fb : forest) (x : resource),
+Lemma check_correct: forall (fa fb : forest),
check fa fb = true -> (forall x, fa # x = fb # x).
Proof.
unfold check, get_forest; intros;
@@ -450,8 +675,8 @@ Lemma tri1:
Proof. crush. Qed.
Definition ge_preserved {A B C D: Type} (ge: Genv.t A B) (tge: Genv.t C D) : Prop :=
- (forall sp op vl, Op.eval_operation ge sp op vl =
- Op.eval_operation tge sp op vl)
+ (forall sp op vl m, Op.eval_operation ge sp op vl m =
+ Op.eval_operation tge sp op vl m)
/\ (forall sp addr vl, Op.eval_addressing ge sp addr vl =
Op.eval_addressing tge sp addr vl).
@@ -460,12 +685,21 @@ Lemma ge_preserved_same:
Proof. unfold ge_preserved; auto. Qed.
Hint Resolve ge_preserved_same : rtlpar.
-Inductive sem_state_ld : sem_state -> sem_state -> Prop :=
-| sem_state_ld_intro:
+Ltac rtlpar_crush := crush; eauto with rtlpar.
+
+Inductive match_states : instr_state -> instr_state -> Prop :=
+| match_states_intro:
forall rs rs' m m',
- regs_lessdef rs rs' ->
+ (forall x, rs !! x = rs' !! x) ->
m = m' ->
- sem_state_ld (mk_sem_state rs m) (mk_sem_state rs' m').
+ match_states (InstrState rs m) (InstrState rs' m').
+
+Inductive match_states_ld : instr_state -> instr_state -> Prop :=
+| match_states_ld_intro:
+ forall rs rs' m m',
+ regs_lessdef rs rs' ->
+ Mem.extends m m' ->
+ match_states_ld (InstrState rs m) (InstrState rs' m').
Lemma sems_det:
forall A ge tge sp st f,
@@ -478,21 +712,21 @@ Proof. Abort.
(*Lemma sem_value_det:
forall A ge tge sp st f v v',
ge_preserved ge tge ->
- sem_value A ge sp st f v ->
- sem_value A tge sp st f v' ->
+ @sem_value A ge sp st f v ->
+ @sem_value A tge sp st f v' ->
v = v'.
Proof.
- intros;
- generalize (sems_det A ge tge sp st f H v v'
- st.(sem_state_memory) st.(sem_state_memory));
+ intros. destruct st.
+ generalize (sems_det A ge tge sp (InstrState rs m) f H v v'
+ m m);
crush.
Qed.
Hint Resolve sem_value_det : rtlpar.
Lemma sem_value_det':
forall FF ge sp s f v v',
- sem_value FF ge sp s f v ->
- sem_value FF ge sp s f v' ->
+ @sem_value FF ge sp s f v ->
+ @sem_value FF ge sp s f v' ->
v = v'.
Proof.
simplify; eauto with rtlpar.
@@ -501,20 +735,20 @@ Qed.
Lemma sem_mem_det:
forall A ge tge sp st f m m',
ge_preserved ge tge ->
- sem_mem A ge sp st f m ->
- sem_mem A tge sp st f m' ->
+ @sem_mem A ge sp st f m ->
+ @sem_mem A tge sp st f m' ->
m = m'.
Proof.
- intros;
- generalize (sems_det A ge tge sp st f H sp sp m m');
+ intros. destruct st.
+ generalize (sems_det A ge tge sp (InstrState rs m0) f H sp sp m m');
crush.
Qed.
Hint Resolve sem_mem_det : rtlpar.
Lemma sem_mem_det':
forall FF ge sp s f m m',
- sem_mem FF ge sp s f m ->
- sem_mem FF ge sp s f m' ->
+ @sem_mem FF ge sp s f m ->
+ @sem_mem FF ge sp s f m' ->
m = m'.
Proof.
simplify; eauto with rtlpar.
@@ -525,9 +759,9 @@ Hint Resolve Val.lessdef_same : rtlpar.
Lemma sem_regset_det:
forall FF ge tge sp st f v v',
ge_preserved ge tge ->
- sem_regset FF ge sp st f v ->
- sem_regset FF tge sp st f v' ->
- regs_lessdef v v'.
+ @sem_regset FF ge sp st f v ->
+ @sem_regset FF tge sp st f v' ->
+ (forall x, v !! x = v' !! x).
Proof.
intros; unfold regs_lessdef.
inv H0; inv H1;
@@ -538,9 +772,9 @@ Hint Resolve sem_regset_det : rtlpar.
Lemma sem_det:
forall FF ge tge sp st f st' st'',
ge_preserved ge tge ->
- sem FF ge sp st f st' ->
- sem FF tge sp st f st'' ->
- sem_state_ld st' st''.
+ @sem FF ge sp st f st' ->
+ @sem FF tge sp st f st'' ->
+ match_states st' st''.
Proof.
intros.
destruct st; destruct st'; destruct st''.
@@ -551,9 +785,9 @@ Hint Resolve sem_det : rtlpar.
Lemma sem_det':
forall FF ge sp st f st' st'',
- sem FF ge sp st f st' ->
- sem FF ge sp st f st'' ->
- sem_state_ld st' st''.
+ @sem FF ge sp st f st' ->
+ @sem FF ge sp st f st'' ->
+ match_states st' st''.
Proof. eauto with rtlpar. Qed.
(*|
@@ -570,7 +804,7 @@ Definition update (f : forest) (i : instr) : forest :=
match i with
| RBnop => f
| RBop p op rl r =>
- f # (Reg r) <- (Eop op (list_translation rl f))
+ f # (Reg r) <- (Eop op (list_translation rl f) (f # Mem))
| RBload p chunk addr rl r =>
f # (Reg r) <- (Eload chunk addr (list_translation rl f) (f # Mem))
| RBstore p chunk addr rl r =>
@@ -588,7 +822,7 @@ Get a sequence from the basic block.
Fixpoint abstract_sequence (f : forest) (b : list instr) : forest :=
match b with
| nil => f
- | i :: l => update (abstract_sequence f l) i
+ | i :: l => abstract_sequence (update f i) l
end.
(*|
@@ -650,14 +884,685 @@ Abstract computations
=====================
|*)
+Definition is_regs i := match i with InstrState rs _ => rs end.
+Definition is_mem i := match i with InstrState _ m => m end.
+
+Inductive state_lessdef : instr_state -> instr_state -> Prop :=
+ state_lessdef_intro :
+ forall rs1 rs2 m1,
+ (forall x, rs1 !! x = rs2 !! x) ->
+ state_lessdef (InstrState rs1 m1) (InstrState rs2 m1).
+
+(*|
+RTLBlock to abstract translation
+--------------------------------
+
+Correctness of translation from RTLBlock to the abstract interpretation language.
+|*)
+
+Lemma match_states_refl x : match_states x x.
+Proof. destruct x; constructor; crush. Qed.
+
+Lemma match_states_commut x y : match_states x y -> match_states y x.
+Proof. inversion 1; constructor; crush. Qed.
+
+Lemma match_states_trans x y z :
+ match_states x y -> match_states y z -> match_states x z.
+Proof. repeat inversion 1; constructor; crush. Qed.
+
+Ltac inv_simp :=
+ repeat match goal with
+ | H: exists _, _ |- _ => inv H
+ end; simplify.
+
+Lemma abstract_interp_empty A ge sp st : @sem A ge sp st empty st.
+Proof. destruct st; repeat constructor. Qed.
+
+Lemma abstract_interp_empty3 :
+ forall A ge sp st st',
+ @sem A ge sp st empty st' ->
+ match_states st st'.
+Proof.
+ inversion 1; subst; simplify.
+ destruct st. inv H1. simplify.
+ constructor. unfold regs_lessdef.
+ intros. inv H0. specialize (H1 x). inv H1; auto.
+ auto.
+Qed.
+
+Definition check_dest i r' :=
+ match i with
+ | RBop p op rl r => (r =? r')%positive
+ | RBload p chunk addr rl r => (r =? r')%positive
+ | _ => false
+ end.
+
+Lemma check_dest_dec i r : {check_dest i r = true} + {check_dest i r = false}.
+Proof. destruct (check_dest i r); tauto. Qed.
+
+Fixpoint check_dest_l l r :=
+ match l with
+ | nil => false
+ | a :: b => check_dest a r || check_dest_l b r
+ end.
+
+Lemma check_dest_l_forall :
+ forall l r,
+ check_dest_l l r = false ->
+ Forall (fun x => check_dest x r = false) l.
+Proof. induction l; crush. Qed.
+
+Lemma check_dest_l_ex :
+ forall l r,
+ check_dest_l l r = true ->
+ exists a, In a l /\ check_dest a r = true.
+Proof.
+ induction l; crush.
+ destruct (check_dest a r) eqn:?; try solve [econstructor; crush].
+ simplify.
+ exploit IHl. apply H. inv_simp. econstructor. simplify. right. eassumption.
+ auto.
+Qed.
+
+Lemma check_dest_l_dec i r : {check_dest_l i r = true} + {check_dest_l i r = false}.
+Proof. destruct (check_dest_l i r); tauto. Qed.
+
+Lemma check_dest_l_dec2 l r :
+ {Forall (fun x => check_dest x r = false) l}
+ + {exists a, In a l /\ check_dest a r = true}.
+Proof.
+ destruct (check_dest_l_dec l r); [right | left];
+ auto using check_dest_l_ex, check_dest_l_forall.
+Qed.
+
+Lemma check_dest_l_forall2 :
+ forall l r,
+ Forall (fun x => check_dest x r = false) l ->
+ check_dest_l l r = false.
+Proof.
+ induction l; crush.
+ inv H. apply orb_false_intro; crush.
+Qed.
+
+Lemma check_dest_l_ex2 :
+ forall l r,
+ (exists a, In a l /\ check_dest a r = true) ->
+ check_dest_l l r = true.
+Proof.
+ induction l; crush.
+ specialize (IHl r). inv H.
+ apply orb_true_intro; crush.
+ apply orb_true_intro; crush.
+ right. apply IHl. exists x. auto.
+Qed.
+
+Lemma check_dest_update :
+ forall f i r,
+ check_dest i r = false ->
+ (update f i) # (Reg r) = f # (Reg r).
+Proof.
+ destruct i; crush; try apply Pos.eqb_neq in H; apply genmap1; crush.
+Qed.
+
+Lemma check_dest_update2 :
+ forall f r rl op p,
+ (update f (RBop p op rl r)) # (Reg r) = Eop op (list_translation rl f) (f # Mem).
+Proof. crush; rewrite map2; auto. Qed.
+
+Lemma check_dest_update3 :
+ forall f r rl p addr chunk,
+ (update f (RBload p chunk addr rl r)) # (Reg r) = Eload chunk addr (list_translation rl f) (f # Mem).
+Proof. crush; rewrite map2; auto. Qed.
+
+Lemma abstr_comp :
+ forall l i f x x0,
+ abstract_sequence f (l ++ i :: nil) = x ->
+ abstract_sequence f l = x0 ->
+ x = update x0 i.
+Proof. induction l; intros; crush; eapply IHl; eauto. Qed.
+
+Lemma abstract_seq :
+ forall l f i,
+ abstract_sequence f (l ++ i :: nil) = update (abstract_sequence f l) i.
+Proof. induction l; crush. Qed.
+
+Lemma check_list_l_false :
+ forall l x r,
+ check_dest_l (l ++ x :: nil) r = false ->
+ check_dest_l l r = false /\ check_dest x r = false.
+Proof.
+ simplify.
+ apply check_dest_l_forall in H. apply Forall_app in H.
+ simplify. apply check_dest_l_forall2; auto.
+ apply check_dest_l_forall in H. apply Forall_app in H.
+ simplify. inv H1. auto.
+Qed.
+
+Lemma check_list_l_true :
+ forall l x r,
+ check_dest_l (l ++ x :: nil) r = true ->
+ check_dest_l l r = true \/ check_dest x r = true.
+Proof.
+ simplify.
+ apply check_dest_l_ex in H; inv_simp.
+ apply in_app_or in H. inv H. left.
+ apply check_dest_l_ex2. exists x0. auto.
+ inv H0; auto.
+Qed.
+
+Lemma abstract_sequence_update :
+ forall l r f,
+ check_dest_l l r = false ->
+ (abstract_sequence f l) # (Reg r) = f # (Reg r).
+Proof.
+ induction l using rev_ind; crush.
+ rewrite abstract_seq. rewrite check_dest_update. apply IHl.
+ apply check_list_l_false in H. tauto.
+ apply check_list_l_false in H. tauto.
+Qed.
+
+Lemma rtlblock_trans_correct' :
+ forall bb ge sp st x st'',
+ RTLBlock.step_instr_list ge sp st (bb ++ x :: nil) st'' ->
+ exists st', RTLBlock.step_instr_list ge sp st bb st'
+ /\ step_instr ge sp st' x st''.
+Proof.
+ induction bb.
+ crush. exists st.
+ split. constructor. inv H. inv H6. auto.
+ crush. inv H. exploit IHbb. eassumption. inv_simp.
+ econstructor. split.
+ econstructor; eauto. eauto.
+Qed.
+
+Lemma sem_update_RBnop :
+ forall A ge sp st f st',
+ @sem A ge sp st f st' -> sem ge sp st (update f RBnop) st'.
+Proof. crush. Qed.
+
+Lemma gen_list_base:
+ forall FF ge sp l rs exps st1,
+ (forall x, @sem_value FF ge sp st1 (exps # (Reg x)) (rs !! x)) ->
+ sem_val_list ge sp st1 (list_translation l exps) rs ## l.
+Proof.
+ induction l.
+ intros. simpl. constructor.
+ intros. simpl. eapply Scons; eauto.
+Qed.
+
+Lemma abstract_seq_correct_aux:
+ forall FF ge sp i st1 st2 st3 f,
+ @step_instr FF ge sp st3 i st2 ->
+ sem ge sp st1 f st3 ->
+ sem ge sp st1 (update f i) st2.
+Proof.
+ intros; inv H; simplify.
+ { simplify; eauto. } (*apply match_states_refl. }*)
+ { inv H0. inv H6. destruct st1. econstructor. simplify.
+ constructor. intros.
+ destruct (resource_eq (Reg res) (Reg x)). inv e.
+ rewrite map2. econstructor. eassumption. apply gen_list_base; eauto.
+ rewrite Regmap.gss. eauto.
+ assert (res <> x). { unfold not in *. intros. apply n. rewrite H0. auto. }
+ rewrite Regmap.gso by auto.
+ rewrite genmap1 by auto. auto.
+
+ rewrite genmap1; crush. }
+ { inv H0. inv H7. constructor. constructor. intros.
+ destruct (Pos.eq_dec dst x); subst.
+ rewrite map2. econstructor; eauto.
+ apply gen_list_base. auto. rewrite Regmap.gss. auto.
+ rewrite genmap1. rewrite Regmap.gso by auto. auto.
+ unfold not in *; intros. inv H0. auto.
+ rewrite genmap1; crush.
+ }
+ { inv H0. inv H7. constructor. constructor; intros.
+ rewrite genmap1; crush.
+ rewrite map2. econstructor; eauto.
+ apply gen_list_base; auto.
+ }
+Qed.
+
+Lemma regmap_list_equiv :
+ forall A (rs1: Regmap.t A) rs2,
+ (forall x, rs1 !! x = rs2 !! x) ->
+ forall rl, rs1##rl = rs2##rl.
+Proof. induction rl; crush. Qed.
+
+Lemma sem_update_Op :
+ forall A ge sp st f st' r l o0 o m rs v,
+ @sem A ge sp st f st' ->
+ Op.eval_operation ge sp o0 rs ## l m = Some v ->
+ match_states st' (InstrState rs m) ->
+ exists tst,
+ sem ge sp st (update f (RBop o o0 l r)) tst /\ match_states (InstrState (Regmap.set r v rs) m) tst.
+Proof.
+ intros. inv H1. simplify.
+ destruct st.
+ econstructor. simplify.
+ { constructor.
+ { constructor. intros. destruct (Pos.eq_dec x r); subst.
+ { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto.
+ { inv H9. eapply gen_list_base; eauto. }
+ { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } }
+ { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } }
+ { inv H. rewrite genmap1; crush. eauto. } }
+ { constructor; eauto. intros.
+ destruct (Pos.eq_dec r x);
+ subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. }
+Qed.
+
+Lemma sem_update_load :
+ forall A ge sp st f st' r o m a l m0 rs v a0,
+ @sem A ge sp st f st' ->
+ Op.eval_addressing ge sp a rs ## l = Some a0 ->
+ Mem.loadv m m0 a0 = Some v ->
+ match_states st' (InstrState rs m0) ->
+ exists tst : instr_state,
+ sem ge sp st (update f (RBload o m a l r)) tst
+ /\ match_states (InstrState (Regmap.set r v rs) m0) tst.
+Proof.
+ intros. inv H2. pose proof H. inv H. inv H9.
+ destruct st.
+ econstructor; simplify.
+ { constructor.
+ { constructor. intros.
+ destruct (Pos.eq_dec x r); subst.
+ { rewrite map2. econstructor; eauto. eapply gen_list_base. intros.
+ rewrite <- H6. eauto.
+ instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. auto. }
+ { rewrite Regmap.gso by auto. rewrite genmap1; crush. } }
+ { rewrite genmap1; crush. eauto. } }
+ { constructor; auto; intros. destruct (Pos.eq_dec r x);
+ subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. }
+Qed.
+
+Lemma sem_update_store :
+ forall A ge sp a0 m a l r o f st m' rs m0 st',
+ @sem A ge sp st f st' ->
+ Op.eval_addressing ge sp a rs ## l = Some a0 ->
+ Mem.storev m m0 a0 rs !! r = Some m' ->
+ match_states st' (InstrState rs m0) ->
+ exists tst, sem ge sp st (update f (RBstore o m a l r)) tst
+ /\ match_states (InstrState rs m') tst.
+Proof.
+ intros. inv H2. pose proof H. inv H. inv H9.
+ destruct st.
+ econstructor; simplify.
+ { econstructor.
+ { econstructor; intros. rewrite genmap1; crush. }
+ { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. rewrite <- H6.
+ eauto. specialize (H6 r). rewrite H6. eauto. } }
+ { econstructor; eauto. }
+Qed.
+
+Lemma sem_update :
+ forall A ge sp st x st' st'' st''' f,
+ sem ge sp st f st' ->
+ match_states st' st''' ->
+ @step_instr A ge sp st''' x st'' ->
+ exists tst, sem ge sp st (update f x) tst /\ match_states st'' tst.
+Proof.
+ intros. destruct x; inv H1.
+ { econstructor. split.
+ apply sem_update_RBnop. eassumption.
+ apply match_states_commut. auto. }
+ { eapply sem_update_Op; eauto. }
+ { eapply sem_update_load; eauto. }
+ { eapply sem_update_store; eauto. }
+Qed.
+
+Lemma sem_update2_Op :
+ forall A ge sp st f r l o0 o m rs v,
+ @sem A ge sp st f (InstrState rs m) ->
+ Op.eval_operation ge sp o0 rs ## l m = Some v ->
+ sem ge sp st (update f (RBop o o0 l r)) (InstrState (Regmap.set r v rs) m).
+Proof.
+ intros. destruct st. constructor.
+ inv H. inv H6.
+ { constructor; intros. simplify.
+ destruct (Pos.eq_dec r x); subst.
+ { rewrite map2. econstructor. eauto.
+ apply gen_list_base. eauto.
+ rewrite Regmap.gss. auto. }
+ { rewrite genmap1; crush. rewrite Regmap.gso; auto. } }
+ { simplify. rewrite genmap1; crush. inv H. eauto. }
+Qed.
+
+Lemma sem_update2_load :
+ forall A ge sp st f r o m a l m0 rs v a0,
+ @sem A ge sp st f (InstrState rs m0) ->
+ Op.eval_addressing ge sp a rs ## l = Some a0 ->
+ Mem.loadv m m0 a0 = Some v ->
+ sem ge sp st (update f (RBload o m a l r)) (InstrState (Regmap.set r v rs) m0).
+Proof.
+ intros. simplify. inv H. inv H7. constructor.
+ { constructor; intros. destruct (Pos.eq_dec r x); subst.
+ { rewrite map2. rewrite Regmap.gss. econstructor; eauto.
+ apply gen_list_base; eauto. }
+ { rewrite genmap1; crush. rewrite Regmap.gso; eauto. }
+ }
+ { simplify. rewrite genmap1; crush. }
+Qed.
+
+Lemma sem_update2_store :
+ forall A ge sp a0 m a l r o f st m' rs m0,
+ @sem A ge sp st f (InstrState rs m0) ->
+ Op.eval_addressing ge sp a rs ## l = Some a0 ->
+ Mem.storev m m0 a0 rs !! r = Some m' ->
+ sem ge sp st (update f (RBstore o m a l r)) (InstrState rs m').
+Proof.
+ intros. simplify. inv H. inv H7. constructor; simplify.
+ { econstructor; intros. rewrite genmap1; crush. }
+ { rewrite map2. econstructor; eauto. apply gen_list_base; eauto. }
+Qed.
+
+Lemma sem_update2 :
+ forall A ge sp st x st' st'' f,
+ sem ge sp st f st' ->
+ @step_instr A ge sp st' x st'' ->
+ sem ge sp st (update f x) st''.
+Proof.
+ intros.
+ destruct x; inv H0;
+ eauto using sem_update_RBnop, sem_update2_Op, sem_update2_load, sem_update2_store.
+Qed.
+
+Lemma rtlblock_trans_correct :
+ forall bb ge sp st st',
+ RTLBlock.step_instr_list ge sp st bb st' ->
+ forall tst,
+ match_states st tst ->
+ exists tst', sem ge sp tst (abstract_sequence empty bb) tst'
+ /\ match_states st' tst'.
+Proof.
+ induction bb using rev_ind; simplify.
+ { econstructor. simplify. apply abstract_interp_empty.
+ inv H. auto. }
+ { apply rtlblock_trans_correct' in H. inv_simp.
+ rewrite abstract_seq.
+ exploit IHbb; try eassumption; []; inv_simp.
+ exploit sem_update. apply H1. apply match_states_commut; eassumption.
+ eauto. inv_simp. econstructor. split. apply H3.
+ auto. }
+Qed.
+
+Lemma abstr_sem_val_mem :
+ forall A B ge tge st tst sp a,
+ ge_preserved ge tge ->
+ forall v m,
+ (@sem_mem A ge sp st a m /\ match_states st tst -> @sem_mem B tge sp tst a m) /\
+ (@sem_value A ge sp st a v /\ match_states st tst -> @sem_value B tge sp tst a v).
+Proof.
+ intros * H.
+ apply expression_ind2 with
+
+ (P := fun (e1: expression) =>
+ forall v m,
+ (@sem_mem A ge sp st e1 m /\ match_states st tst -> @sem_mem B tge sp tst e1 m) /\
+ (@sem_value A ge sp st e1 v /\ match_states st tst -> @sem_value B tge sp tst e1 v))
+
+ (P0 := fun (e1: expression_list) =>
+ forall lv, @sem_val_list A ge sp st e1 lv /\ match_states st tst -> @sem_val_list B tge sp tst e1 lv);
+ simplify; intros; simplify.
+ { inv H1. inv H2. constructor. }
+ { inv H2. inv H1. rewrite H0. constructor. }
+ { inv H3. }
+ { inv H3. inv H4. econstructor. apply H1; auto. simplify. eauto. constructor. auto. auto.
+ apply H0; simplify; eauto. constructor; eauto.
+ unfold ge_preserved in *. simplify. rewrite <- H2. auto.
+ }
+ { inv H3. }
+ { inv H3. inv H4. econstructor. apply H1; eauto; simplify; eauto. constructor; eauto.
+ apply H0; simplify; eauto. constructor; eauto.
+ inv H. rewrite <- H4. eauto.
+ auto.
+ }
+ { inv H4. inv H5. econstructor. apply H0; eauto. simplify; eauto. constructor; eauto.
+ apply H2; eauto. simplify; eauto. constructor; eauto.
+ apply H1; eauto. simplify; eauto. constructor; eauto.
+ inv H. rewrite <- H5. eauto. auto.
+ }
+ { inv H4. }
+ { inv H1. constructor. }
+ { inv H3. constructor; auto. apply H0; eauto. apply Mem.empty. }
+Qed.
+
+Lemma abstr_sem_value :
+ forall a A B ge tge sp st tst v,
+ @sem_value A ge sp st a v ->
+ ge_preserved ge tge ->
+ match_states st tst ->
+ @sem_value B tge sp tst a v.
+Proof. intros; eapply abstr_sem_val_mem; eauto; apply Mem.empty. Qed.
+
+Lemma abstr_sem_mem :
+ forall a A B ge tge sp st tst v,
+ @sem_mem A ge sp st a v ->
+ ge_preserved ge tge ->
+ match_states st tst ->
+ @sem_mem B tge sp tst a v.
+Proof. intros; eapply abstr_sem_val_mem; eauto. Qed.
+
+Lemma abstr_sem_regset :
+ forall a a' A B ge tge sp st tst rs,
+ @sem_regset A ge sp st a rs ->
+ ge_preserved ge tge ->
+ (forall x, a # x = a' # x) ->
+ match_states st tst ->
+ exists rs', @sem_regset B tge sp tst a' rs' /\ (forall x, rs !! x = rs' !! x).
+Proof.
+ inversion 1; intros.
+ inv H7.
+ econstructor. simplify. econstructor. intros.
+ eapply abstr_sem_value; eauto. rewrite <- H6.
+ eapply H0. constructor; eauto.
+ auto.
+Qed.
+
+Lemma abstr_sem :
+ forall a a' A B ge tge sp st tst st',
+ @sem A ge sp st a st' ->
+ ge_preserved ge tge ->
+ (forall x, a # x = a' # x) ->
+ match_states st tst ->
+ exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'.
+Proof.
+ inversion 1; subst; intros.
+ inversion H4; subst.
+ exploit abstr_sem_regset; eauto; inv_simp.
+ do 3 econstructor; eauto.
+ rewrite <- H3.
+ eapply abstr_sem_mem; eauto.
+Qed.
+
+Lemma abstract_execution_correct':
+ forall A B ge tge sp st' a a' st tst,
+ @sem A ge sp st a st' ->
+ ge_preserved ge tge ->
+ check a a' = true ->
+ match_states st tst ->
+ exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'.
+Proof.
+ intros;
+ pose proof (check_correct a a' H1);
+ eapply abstr_sem; eauto.
+Qed.
+
+Lemma states_match :
+ forall st1 st2 st3 st4,
+ match_states st1 st2 ->
+ match_states st2 st3 ->
+ match_states st3 st4 ->
+ match_states st1 st4.
+Proof.
+ intros * H1 H2 H3; destruct st1; destruct st2; destruct st3; destruct st4.
+ inv H1. inv H2. inv H3; constructor.
+ unfold regs_lessdef in *. intros.
+ repeat match goal with
+ | H: forall _, _, r : positive |- _ => specialize (H r)
+ end.
+ congruence.
+ auto.
+Qed.
+
+Lemma step_instr_block_same :
+ forall ge sp st st',
+ step_instr_block ge sp st nil st' ->
+ st = st'.
+Proof. inversion 1; auto. Qed.
+
+Lemma step_instr_seq_same :
+ forall ge sp st st',
+ step_instr_seq ge sp st nil st' ->
+ st = st'.
+Proof. inversion 1; auto. Qed.
+
+Lemma match_states_list :
+ forall A (rs: Regmap.t A) rs',
+ (forall r, rs !! r = rs' !! r) ->
+ forall l, rs ## l = rs' ## l.
+Proof. induction l; crush. Qed.
+
+Lemma PTree_matches :
+ forall A (v: A) res rs rs',
+ (forall r, rs !! r = rs' !! r) ->
+ forall x, (Regmap.set res v rs) !! x = (Regmap.set res v rs') !! x.
+Proof.
+ intros; destruct (Pos.eq_dec x res); subst;
+ [ repeat rewrite Regmap.gss by auto
+ | repeat rewrite Regmap.gso by auto ]; auto.
+Qed.
+
+Lemma step_instr_matches :
+ forall A a ge sp st st',
+ @step_instr A ge sp st a st' ->
+ forall tst, match_states st tst ->
+ exists tst', step_instr ge sp tst a tst'
+ /\ match_states st' tst'.
+Proof.
+ induction 1; simplify;
+ match goal with H: match_states _ _ |- _ => inv H end;
+ repeat econstructor; try erewrite match_states_list;
+ try apply PTree_matches; eauto;
+ match goal with
+ H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto
+ end.
+Qed.
+
+Lemma step_instr_list_matches :
+ forall a ge sp st st',
+ step_instr_list ge sp st a st' ->
+ forall tst, match_states st tst ->
+ exists tst', step_instr_list ge sp tst a tst'
+ /\ match_states st' tst'.
+Proof.
+ induction a; intros; inv H;
+ try (exploit step_instr_matches; eauto; []; inv_simp;
+ exploit IHa; eauto; []; inv_simp); repeat econstructor; eauto.
+Qed.
+
+Lemma step_instr_seq_matches :
+ forall a ge sp st st',
+ step_instr_seq ge sp st a st' ->
+ forall tst, match_states st tst ->
+ exists tst', step_instr_seq ge sp tst a tst'
+ /\ match_states st' tst'.
+Proof.
+ induction a; intros; inv H;
+ try (exploit step_instr_list_matches; eauto; []; inv_simp;
+ exploit IHa; eauto; []; inv_simp); repeat econstructor; eauto.
+Qed.
+
+Lemma step_instr_block_matches :
+ forall bb ge sp st st',
+ step_instr_block ge sp st bb st' ->
+ forall tst, match_states st tst ->
+ exists tst', step_instr_block ge sp tst bb tst'
+ /\ match_states st' tst'.
+Proof.
+ induction bb; intros; inv H;
+ try (exploit step_instr_seq_matches; eauto; []; inv_simp;
+ exploit IHbb; eauto; []; inv_simp); repeat econstructor; eauto.
+Qed.
+
+Lemma sem_update' :
+ forall A ge sp st a x st',
+ sem ge sp st (update (abstract_sequence empty a) x) st' ->
+ exists st'',
+ @step_instr A ge sp st'' x st' /\
+ sem ge sp st (abstract_sequence empty a) st''.
+Proof.
+ Admitted.
+
+Lemma sem_separate :
+ forall A (ge: @RTLBlockInstr.genv A) b a sp st st',
+ sem ge sp st (abstract_sequence empty (a ++ b)) st' ->
+ exists st'',
+ sem ge sp st (abstract_sequence empty a) st''
+ /\ sem ge sp st'' (abstract_sequence empty b) st'.
+Proof.
+ induction b using rev_ind; simplify.
+ { econstructor. simplify. rewrite app_nil_r in H. eauto. apply abstract_interp_empty. }
+ { simplify. rewrite app_assoc in H. rewrite abstract_seq in H.
+ exploit sem_update'; eauto; inv_simp.
+ exploit IHb; eauto; inv_simp.
+ econstructor; split; eauto.
+ rewrite abstract_seq.
+ eapply sem_update2; eauto.
+ }
+Qed.
+
+Lemma rtlpar_trans_correct :
+ forall bb ge sp sem_st' sem_st st,
+ sem ge sp sem_st (abstract_sequence empty (concat (concat bb))) sem_st' ->
+ match_states sem_st st ->
+ exists st', RTLPar.step_instr_block ge sp st bb st'
+ /\ match_states sem_st' st'.
+Proof.
+ induction bb using rev_ind.
+ { repeat econstructor. eapply abstract_interp_empty3 in H.
+ inv H. inv H0. constructor; congruence. }
+ { simplify. inv H0. repeat rewrite concat_app in H. simplify.
+ rewrite app_nil_r in H.
+ exploit sem_separate; eauto; inv_simp.
+ repeat econstructor. admit. admit.
+ }
+Admitted.
+
Lemma abstract_execution_correct:
- forall bb bb' cfi ge tge sp rs m rs' m',
+ forall bb bb' cfi ge tge sp st st' tst,
+ RTLBlock.step_instr_list ge sp st bb st' ->
ge_preserved ge tge ->
schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi) = true ->
- RTLBlock.step_instr_list ge sp (InstrState rs m) bb (InstrState rs' m') ->
- exists rs'', RTLPar.step_instr_block tge sp (InstrState rs m) bb' (InstrState rs'' m')
- /\ regs_lessdef rs' rs''.
-Proof. Abort.
+ match_states st tst ->
+ exists tst', RTLPar.step_instr_block tge sp tst bb' tst'
+ /\ match_states st' tst'.
+Proof.
+ intros.
+ unfold schedule_oracle in *. simplify.
+ exploit rtlblock_trans_correct; try eassumption; []; inv_simp.
+ exploit abstract_execution_correct';
+ try solve [eassumption | apply state_lessdef_match_sem; eassumption].
+ apply match_states_commut. eauto. inv_simp.
+ exploit rtlpar_trans_correct; try eassumption; []; inv_simp.
+ exploit step_instr_block_matches; eauto. apply match_states_commut; eauto. inv_simp.
+ repeat match goal with | H: match_states _ _ |- _ => inv H end.
+ do 2 econstructor; eauto.
+ econstructor; congruence.
+Qed.
+
+(*Lemma abstract_execution_correct_ld:
+ forall bb bb' cfi ge tge sp st st' tst,
+ RTLBlock.step_instr_list ge sp st bb st' ->
+ ge_preserved ge tge ->
+ schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi) = true ->
+ match_states_ld st tst ->
+ exists tst', RTLPar.step_instr_block tge sp tst bb' tst'
+ /\ match_states st' tst'.
+Proof.
+ intros.*)
+
(*|
Top-level functions
@@ -677,15 +1582,7 @@ Definition transl_function (f: RTLBlock.function) : Errors.res RTLPar.function :
else
Errors.Error (Errors.msg "RTLPargen: Could not prove the blocks equivalent.").
-Definition transl_function_temp (f: RTLBlock.function) : Errors.res RTLPar.function :=
- let tfcode := fn_code (schedule f) in
- Errors.OK (mkfunction f.(fn_sig)
- f.(fn_params)
- f.(fn_stacksize)
- tfcode
- f.(fn_entrypoint)).
-
-Definition transl_fundef := transf_partial_fundef transl_function_temp.
+Definition transl_fundef := transf_partial_fundef transl_function.
Definition transl_program (p : RTLBlock.program) : Errors.res RTLPar.program :=
transform_partial_program transl_fundef p.
diff --git a/src/hls/RTLPargenproof.v b/src/hls/RTLPargenproof.v
index a0c01fa..bb1cf97 100644
--- a/src/hls/RTLPargenproof.v
+++ b/src/hls/RTLPargenproof.v
@@ -38,34 +38,29 @@ Inductive match_stackframes: RTLBlock.stackframe -> RTLPar.stackframe -> Prop :=
| match_stackframe:
forall f tf res sp pc rs rs',
transl_function f = OK tf ->
- regs_lessdef rs rs' ->
+ (forall x, rs !! x = rs' !! x) ->
match_stackframes (Stackframe res f sp pc rs)
(Stackframe res tf sp pc rs').
Inductive match_states: RTLBlock.state -> RTLPar.state -> Prop :=
| match_state:
- forall sf f sp pc rs rs' m m' sf' tf
+ forall sf f sp pc rs rs' m sf' tf
(TRANSL: transl_function f = OK tf)
(STACKS: list_forall2 match_stackframes sf sf')
- (REG: regs_lessdef rs rs')
- (MEM: Mem.extends m m'),
+ (REG: forall x, rs !! x = rs' !! x),
match_states (State sf f sp pc rs m)
- (State sf' tf sp pc rs' m')
+ (State sf' tf sp pc rs' m)
| match_returnstate:
- forall stack stack' v v' m m'
- (STACKS: list_forall2 match_stackframes stack stack')
- (MEM: Mem.extends m m')
- (LD: Val.lessdef v v'),
+ forall stack stack' v m
+ (STACKS: list_forall2 match_stackframes stack stack'),
match_states (Returnstate stack v m)
- (Returnstate stack' v' m')
+ (Returnstate stack' v m)
| match_callstate:
- forall stack stack' f tf args args' m m'
+ forall stack stack' f tf args m
(TRANSL: transl_fundef f = OK tf)
- (STACKS: list_forall2 match_stackframes stack stack')
- (LD: Val.lessdef_list args args')
- (MEM: Mem.extends m m'),
+ (STACKS: list_forall2 match_stackframes stack stack'),
match_states (Callstate stack f args m)
- (Callstate stack' tf args' m').
+ (Callstate stack' tf args m).
Section CORRECTNESS.
@@ -121,7 +116,7 @@ Section CORRECTNESS.
Lemma find_function_translated:
forall ros rs rs' f,
- regs_lessdef rs rs' ->
+ (forall x, rs !! x = rs' !! x) ->
find_function ge ros rs = Some f ->
exists tf, find_function tge ros rs' = Some tf
/\ transl_fundef f = OK tf.
@@ -134,7 +129,7 @@ Section CORRECTNESS.
| [ H: Genv.find_funct _ Vundef = Some _ |- _] => solve [inv H]
| _ => solve [exploit functions_translated; eauto]
end.
- unfold regs_lessdef; destruct ros; simplify; try rewrite <- H;
+ destruct ros; simplify; try rewrite <- H;
[| rewrite symbols_preserved; destruct_match;
try (apply function_ptr_translated); crush ];
intros;
@@ -160,8 +155,8 @@ Section CORRECTNESS.
Qed.
Lemma eval_op_eq:
- forall (sp0 : Values.val) (op : Op.operation) (vl : list Values.val),
- Op.eval_operation ge sp0 op vl = Op.eval_operation tge sp0 op vl.
+ forall (sp0 : Values.val) (op : Op.operation) (vl : list Values.val) m,
+ Op.eval_operation ge sp0 op vl m = Op.eval_operation tge sp0 op vl m.
Proof using TRANSL.
intros.
destruct op; auto; unfold Op.eval_operation, Genv.symbol_address, Op.eval_addressing32;
@@ -197,6 +192,16 @@ Section CORRECTNESS.
Proof using. destruct or; crush. Qed.
Hint Resolve lessdef_regmap_optget : rtlgp.
+ Lemma regmap_equiv_lessdef:
+ forall rs rs',
+ (forall x, rs !! x = rs' !! x) ->
+ regs_lessdef rs rs'.
+ Proof using.
+ intros; unfold regs_lessdef; intros.
+ rewrite H. apply Val.lessdef_refl.
+ Qed.
+ Hint Resolve regmap_equiv_lessdef : rtlgp.
+
Lemma int_lessdef:
forall rs rs',
regs_lessdef rs rs' ->
@@ -227,8 +232,8 @@ Section CORRECTNESS.
let H2 := fresh "SCHED" in
learn H as H2;
apply schedule_oracle_nil in H2
- | [ H: find_function _ _ _ = Some _ |- _ ] =>
- learn H; exploit find_function_translated; eauto; inversion 1
+ | [ H: find_function _ _ _ = Some _, H2: forall x, ?rs !! x = ?rs' !! x |- _ ] =>
+ learn H; exploit find_function_translated; try apply H2; eauto; inversion 1
| [ H: Mem.free ?m _ _ _ = Some ?m', H2: Mem.extends ?m ?m'' |- _ ] =>
learn H; exploit Mem.free_parallel_extends; eauto; intros
| [ H: Events.eval_builtin_args _ _ _ _ _ _, H2: regs_lessdef ?rs ?rs' |- _ ] =>
@@ -249,6 +254,29 @@ Section CORRECTNESS.
Hint Resolve set_reg_lessdef : rtlgp.
Hint Resolve Op.eval_condition_lessdef : rtlgp.
+ Hint Constructors Events.eval_builtin_arg: barg.
+
+ Lemma eval_builtin_arg_eq:
+ forall A ge a v1 m1 e1 e2 sp,
+ (forall x, e1 x = e2 x) ->
+ @Events.eval_builtin_arg A ge e1 sp m1 a v1 ->
+ Events.eval_builtin_arg ge e2 sp m1 a v1.
+Proof. induction 2; try rewrite H; eauto with barg. Qed.
+
+ Lemma eval_builtin_args_eq:
+ forall A ge e1 sp m1 e2 al vl1,
+ (forall x, e1 x = e2 x) ->
+ @Events.eval_builtin_args A ge e1 sp m1 al vl1 ->
+ Events.eval_builtin_args ge e2 sp m1 al vl1.
+ Proof.
+ induction 2.
+ - econstructor; split.
+ - exploit eval_builtin_arg_eq; eauto. intros.
+ destruct IHlist_forall2 as [| y]. constructor; eauto.
+ constructor. constructor; auto.
+ constructor; eauto.
+ Qed.
+
Lemma step_cf_instr_correct:
forall cfi t s s',
step_cf_instr ge s cfi t s' ->
@@ -257,7 +285,22 @@ Section CORRECTNESS.
exists r', step_cf_instr tge r cfi t r' /\ match_states s' r'.
Proof using TRANSL.
induction 1; repeat semantics_simpl;
- repeat (econstructor; eauto with rtlgp).
+ try solve [repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp)].
+ { do 3 (try erewrite match_states_list by eauto; econstructor; eauto with rtlgp).
+ eapply eval_builtin_args_eq. eapply REG.
+ eapply Events.eval_builtin_args_preserved. eapply symbols_preserved.
+ eauto.
+ intros.
+ unfold regmap_setres. destruct res.
+ destruct (Pos.eq_dec x0 x); subst.
+ repeat rewrite Regmap.gss; auto.
+ repeat rewrite Regmap.gso; auto.
+ eapply REG. eapply REG.
+ }
+ { repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp).
+ unfold regmap_optget. destruct or. rewrite REG. constructor; eauto.
+ constructor; eauto.
+ }
Qed.
Theorem transl_step_correct :
@@ -269,21 +312,69 @@ Section CORRECTNESS.
Proof.
induction 1; repeat semantics_simpl.
- Abort.
-
-(* { destruct bb as [bbc bbe]; destruct x as [bbc' bbe'].
- assert (bbe = bbe') by admit.
- rewrite H3 in H5.
- eapply abstract_execution_correct in H5; eauto with rtlgp.
- repeat econstructor; eauto with rtlgp. simplify.
- exploit step_cf_instr_correct. eauto.
- econstructor; eauto with rtlgp.
+
+ { destruct bb; destruct x.
+ assert (bb_exit = bb_exit0).
+ { unfold schedule_oracle in *. simplify.
+ unfold check_control_flow_instr in *.
+ destruct_match; crush.
+ }
+ subst.
+
+ exploit abstract_execution_correct; try eassumption. eapply ge_preserved_lem.
+ econstructor; eauto.
+ inv_simp. destruct x. inv H7.
+
+ exploit step_cf_instr_correct; try eassumption. econstructor; eauto.
+ inv_simp.
+
+ econstructor. econstructor. eapply Smallstep.plus_one. econstructor.
+ eauto. eauto. simplify. eauto. eauto.
+ }
+ { unfold bind in *. inv TRANSL0. clear Learn. inv H0. destruct_match; crush.
+ inv H2. unfold transl_function in Heqr. destruct_match; crush.
+ inv Heqr.
+ repeat econstructor; eauto.
+ unfold bind in *. destruct_match; crush.
}
- { unfold bind in *. destruct_match; try discriminate. repeat semantics_simpl. inv TRANSL0.
- repeat econstructor; eauto. }
{ inv TRANSL0. repeat econstructor; eauto using Events.external_call_symbols_preserved, symbols_preserved, senv_preserved, Events.E0_right. }
- { inv STACKS. inv H2. repeat econstructor; eauto. }
- Qed.*)
+ { inv STACKS. inv H2. repeat econstructor; eauto.
+ intros. apply PTree_matches; eauto.
+ }
+ Qed.
+
+ Lemma transl_initial_states:
+ forall S,
+ RTLBlock.initial_state prog S ->
+ exists R, RTLPar.initial_state tprog R /\ match_states S R.
+ Proof.
+ induction 1.
+ exploit function_ptr_translated; eauto. intros [tf [A B]].
+ econstructor; split.
+ econstructor. apply (Genv.init_mem_transf_partial TRANSL); eauto.
+ replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved; eauto.
+ symmetry; eapply match_program_main; eauto.
+ eexact A.
+ rewrite <- H2. apply sig_transl_function; auto.
+ constructor. auto. constructor.
+ Qed.
+
+ Lemma transl_final_states:
+ forall S R r,
+ match_states S R -> RTLBlock.final_state S r -> RTLPar.final_state R r.
+ Proof.
+ intros. inv H0. inv H. inv STACKS. constructor.
+ Qed.
+
+ Theorem transf_program_correct:
+ Smallstep.forward_simulation (RTLBlock.semantics prog) (RTLPar.semantics tprog).
+ Proof.
+ eapply Smallstep.forward_simulation_plus.
+ apply senv_preserved.
+ eexact transl_initial_states.
+ eexact transl_final_states.
+ exact transl_step_correct.
+ Qed.
End CORRECTNESS.
*)