aboutsummaryrefslogtreecommitdiffstats
path: root/src
diff options
context:
space:
mode:
authorYann Herklotz <git@yannherklotz.com>2021-11-01 09:12:43 +0000
committerYann Herklotz <git@yannherklotz.com>2021-11-01 09:12:43 +0000
commit352c05e9350f84dc8fe3ba5d10b7b76c184e5f84 (patch)
tree01da4a92cc07d759d069f070ee99753dfd2891dd /src
parent04db0fb8a8a891200fd4767d780efc5a897f72a0 (diff)
downloadvericert-352c05e9350f84dc8fe3ba5d10b7b76c184e5f84.tar.gz
vericert-352c05e9350f84dc8fe3ba5d10b7b76c184e5f84.zip
Add work on RTLPargenproof.v
Diffstat (limited to 'src')
-rw-r--r--src/hls/RTLPargen.v282
-rw-r--r--src/hls/RTLPargenproof.v23
2 files changed, 165 insertions, 140 deletions
diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v
index 30a35f3..da8d527 100644
--- a/src/hls/RTLPargen.v
+++ b/src/hls/RTLPargen.v
@@ -131,21 +131,10 @@ Definition ge_preserved {A B C D: Type} (ge: Genv.t A B) (tge: Genv.t C D) : Pro
/\ (forall sp addr vl, Op.eval_addressing ge sp addr vl =
Op.eval_addressing tge sp addr vl).
-Lemma ge_preserved_same:
- forall A B ge, @ge_preserved A B A B ge ge.
-Proof. unfold ge_preserved; auto. Qed.
#[local] Hint Resolve ge_preserved_same : rtlpar.
Ltac rtlpar_crush := crush; eauto with rtlpar.
-Inductive match_states : instr_state -> instr_state -> Prop :=
-| match_states_intro:
- forall ps ps' rs rs' m m',
- (forall x, rs !! x = rs' !! x) ->
- (forall x, ps !! x = ps' !! x) ->
- m = m' ->
- match_states (mk_instr_state rs ps m) (mk_instr_state rs' ps' m').
-
Inductive match_states_ld : instr_state -> instr_state -> Prop :=
| match_states_ld_intro:
forall ps ps' rs rs' m m',
@@ -154,13 +143,13 @@ Inductive match_states_ld : instr_state -> instr_state -> Prop :=
Mem.extends m m' ->
match_states_ld (mk_instr_state rs ps m) (mk_instr_state rs' ps' m').
-Lemma sems_det:
+(*Lemma sems_det:
forall A ge tge sp f rs ps m,
ge_preserved ge tge ->
forall v v' mv mv',
(@sem_value A (mk_ctx rs ps m sp ge) f v /\ @sem_value A (mk_ctx rs ps m sp tge) f v' -> v = v') /\
(@sem_mem A (mk_ctx rs ps m sp ge) f mv /\ @sem_mem A (mk_ctx rs ps m sp tge) f mv' -> mv = mv').
-Proof. Abort.
+Proof. Abort.*)
(*Lemma sem_value_det:
forall A ge tge sp st f v v',
@@ -422,24 +411,11 @@ 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' ->
@@ -543,11 +519,6 @@ Lemma abstr_comp :
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 ->
@@ -583,19 +554,7 @@ Proof.
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',
@@ -790,25 +749,6 @@ Proof.
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 ->
@@ -940,6 +880,62 @@ Lemma step_instr_seq_same :
st = st'.
Proof. inversion 1; auto. 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_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.*)
+*)
+
Lemma match_states_list :
forall A (rs: Regmap.t A) rs',
(forall r, rs !! r = rs' !! r) ->
@@ -958,18 +954,53 @@ 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'.
+ @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 solve [repeat econstructor; try erewrite match_states_list;
try apply PTree_matches; eauto;
match goal with
H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto
- end.
+ end].
+ - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end.
+ repeat econstructor; try erewrite match_states_list; eauto.
+ erewrite <- eval_predf_pr_equiv; eassumption.
+ apply PTree_matches; assumption.
+ repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto.
+ erewrite <- eval_predf_pr_equiv; eassumption.
+ econstructor; auto.
+ match goal with H: eval_pred _ _ _ _ |- _ => inv H end.
+ repeat econstructor; try erewrite match_states_list; eauto.
+ - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end.
+ repeat econstructor; try erewrite match_states_list; eauto.
+ erewrite <- eval_predf_pr_equiv; eassumption.
+ apply PTree_matches; assumption.
+ repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto.
+ erewrite <- eval_predf_pr_equiv; eassumption.
+ econstructor; auto.
+ match goal with H: eval_pred _ _ _ _ |- _ => inv H end.
+ repeat econstructor; try erewrite match_states_list; eauto.
+ - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end.
+ repeat econstructor; try erewrite match_states_list; eauto.
+ match goal with
+ H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto
+ end.
+ erewrite <- eval_predf_pr_equiv; eassumption.
+ repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto.
+ match goal with
+ H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto
+ end.
+ erewrite <- eval_predf_pr_equiv; eassumption.
+ match goal with H: eval_pred _ _ _ _ |- _ => inv H end.
+ repeat econstructor; try erewrite match_states_list; eauto.
+ match goal with
+ H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto
+ end.
Qed.
Lemma step_instr_list_matches :
@@ -980,8 +1011,8 @@ Lemma step_instr_list_matches :
/\ 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.
+ try (exploit step_instr_matches; eauto; []; simplify;
+ exploit IHa; eauto; []; simplify); repeat econstructor; eauto.
Qed.
Lemma step_instr_seq_matches :
@@ -992,8 +1023,8 @@ Lemma step_instr_seq_matches :
/\ 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.
+ try (exploit step_instr_list_matches; eauto; []; simplify;
+ exploit IHa; eauto; []; simplify); repeat econstructor; eauto.
Qed.
Lemma step_instr_block_matches :
@@ -1004,66 +1035,71 @@ Lemma step_instr_block_matches :
/\ 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.
+ try (exploit step_instr_seq_matches; eauto; []; simplify;
+ exploit IHbb; eauto; []; simplify); 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''.
+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.
- Admitted.
+ induction bb.
+ crush. exists st.
+ split. constructor. inv H. inv H6. auto.
+ crush. inv H. exploit IHbb. eassumption. simplify.
+ econstructor. split.
+ econstructor; eauto. eauto.
+Qed.
-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'.
+Lemma abstract_interp_empty A st : @sem A st empty (ctx_is st).
+Proof. destruct st, ctx_is. simpl. repeat econstructor. 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 sem_update :
+ forall A ge sp st x st' st'' st''' f,
+ sem (mk_ctx st sp ge) f st' ->
+ match_states st' st''' ->
+ @step_instr A ge sp st''' x st'' ->
+ exists tst, sem (mk_ctx st sp ge) (update f x) tst /\ match_states st'' tst.
+Proof. Admitted.
+
+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 (mk_ctx tst sp ge) (abstract_sequence empty bb) tst'
+ /\ match_states st' tst'.
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.
+ induction bb using rev_ind; simplify.
+ { econstructor. simplify. apply abstract_interp_empty.
+ inv H. auto. }
+ { apply rtlblock_trans_correct' in H. simplify.
rewrite abstract_seq.
- eapply sem_update2; eauto.
- }
+ exploit IHbb; try eassumption; []; simplify.
+ exploit sem_update. apply H1. symmetry; eassumption.
+ eauto. simplify. econstructor. split. apply H3.
+ auto. }
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 st st' tst,
+ forall bb bb' cfi 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 ->
+ schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi') = true ->
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.
+ unfold schedule_oracle in *. simplify. unfold empty_trees in H4.
+ exploit rtlblock_trans_correct; try eassumption; []; simplify.
exploit abstract_execution_correct';
try solve [eassumption | apply state_lessdef_match_sem; eassumption].
apply match_states_commut. eauto. inv_simp.
@@ -1074,18 +1110,6 @@ Proof.
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
===================
diff --git a/src/hls/RTLPargenproof.v b/src/hls/RTLPargenproof.v
index bb1cf97..fc3272a 100644
--- a/src/hls/RTLPargenproof.v
+++ b/src/hls/RTLPargenproof.v
@@ -16,7 +16,7 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
-(*Require Import compcert.backend.Registers.
+Require Import compcert.backend.Registers.
Require Import compcert.common.AST.
Require Import compcert.common.Errors.
Require Import compcert.common.Linking.
@@ -36,20 +36,22 @@ Definition match_prog (prog : RTLBlock.program) (tprog : RTLPar.program) :=
Inductive match_stackframes: RTLBlock.stackframe -> RTLPar.stackframe -> Prop :=
| match_stackframe:
- forall f tf res sp pc rs rs',
+ forall f tf res sp pc rs rs' ps ps',
transl_function f = OK tf ->
(forall x, rs !! x = rs' !! x) ->
- match_stackframes (Stackframe res f sp pc rs)
- (Stackframe res tf sp pc rs').
+ (forall x, ps !! x = ps' !! x) ->
+ match_stackframes (Stackframe res f sp pc rs ps)
+ (Stackframe res tf sp pc rs' ps).
Inductive match_states: RTLBlock.state -> RTLPar.state -> Prop :=
| match_state:
- forall sf f sp pc rs rs' m sf' tf
+ forall sf f sp pc rs rs' m sf' tf ps ps'
(TRANSL: transl_function f = OK tf)
(STACKS: list_forall2 match_stackframes sf sf')
- (REG: forall x, rs !! x = rs' !! x),
- match_states (State sf f sp pc rs m)
- (State sf' tf sp pc rs' m)
+ (REG: forall x, rs !! x = rs' !! x)
+ (REG: forall x, ps !! x = ps' !! x),
+ match_states (State sf f sp pc rs ps m)
+ (State sf' tf sp pc rs' ps' m)
| match_returnstate:
forall stack stack' v m
(STACKS: list_forall2 match_stackframes stack stack'),
@@ -284,7 +286,7 @@ Proof. induction 2; try rewrite H; eauto with barg. Qed.
match_states s r ->
exists r', step_cf_instr tge r cfi t r' /\ match_states s' r'.
Proof using TRANSL.
- induction 1; repeat semantics_simpl;
+ (*induction 1; repeat semantics_simpl;
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.
@@ -301,7 +303,7 @@ Proof. induction 2; try rewrite H; eauto with barg. Qed.
unfold regmap_optget. destruct or. rewrite REG. constructor; eauto.
constructor; eauto.
}
- Qed.
+ Qed.*) Admitted.
Theorem transl_step_correct :
forall (S1 : RTLBlock.state) t S2,
@@ -377,4 +379,3 @@ Proof. induction 2; try rewrite H; eauto with barg. Qed.
Qed.
End CORRECTNESS.
-*)