From 86e64fd05cea7b1da996701cd3653db5f471f8d1 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 10 Aug 2023 12:09:03 +0100 Subject: Finish final forward simulation correctness --- src/hls/IfConversionproof.v | 82 ++++++++++++++++++++++++--------------------- 1 file changed, 44 insertions(+), 38 deletions(-) (limited to 'src/hls/IfConversionproof.v') diff --git a/src/hls/IfConversionproof.v b/src/hls/IfConversionproof.v index 3c9d2bd..eedb318 100644 --- a/src/hls/IfConversionproof.v +++ b/src/hls/IfConversionproof.v @@ -50,8 +50,8 @@ Require Import vericert.hls.Predicate. Variant match_stackframe : stackframe -> stackframe -> Prop := | match_stackframe_init : - forall res f tf sp pc rs p l - (TF: transf_function l f = tf), + forall res f tf sp pc rs p + (TF: transf_function f = tf), match_stackframe (Stackframe res f sp pc rs p) (Stackframe res tf sp pc rs p). Definition bool_order (b: bool): nat := if b then 1 else 0. @@ -234,8 +234,8 @@ Proof. Qed. Lemma transf_spec_correct : - forall f pc l, - if_conv_spec f.(fn_code) (transf_function l f).(fn_code) pc. + forall f pc, + if_conv_spec f.(fn_code) (transf_function f).(fn_code) pc. Proof. intros; unfold transf_function; destruct_match; cbn. unfold if_convert_code. @@ -309,8 +309,8 @@ Section CORRECTNESS. Variant match_states : option SeqBB.t -> state -> state -> Prop := | match_state_some : - forall stk stk' f tf sp pc rs p m b pc0 rs0 p0 m0 l - (TF: transf_function l f = tf) + forall stk stk' f tf sp pc rs p m b pc0 rs0 p0 m0 + (TF: transf_function f = tf) (STK: Forall2 match_stackframe stk stk') (* This can be improved with a recursive relation for a more general structure of the if-conversion proof. *) @@ -322,13 +322,13 @@ Section CORRECTNESS. (SIM: step ge (State stk f sp pc0 rs0 p0 m0) E0 (State stk f sp pc rs p m)), match_states (Some b) (State stk f sp pc rs p m) (State stk' tf sp pc0 rs0 p0 m0) | match_state_none : - forall stk stk' f tf sp pc rs p m l - (TF: transf_function l f = tf) + forall stk stk' f tf sp pc rs p m + (TF: transf_function f = tf) (STK: Forall2 match_stackframe stk stk'), match_states None (State stk f sp pc rs p m) (State stk' tf sp pc rs p m) | match_callstate : - forall cs cs' f tf args m l - (TF: transf_fundef l f = tf) + forall cs cs' f tf args m + (TF: transf_fundef f = tf) (STK: Forall2 match_stackframe cs cs'), match_states None (Callstate cs f args m) (Callstate cs' tf args m) | match_returnstate : @@ -337,7 +337,13 @@ Section CORRECTNESS. match_states None (Returnstate cs v m) (Returnstate cs' v m). Definition match_prog (p: program) (tp: program) := - match_program (fun _ f tf => exists l, transf_fundef l f = tf) eq p tp. + match_program (fun _ f tf => tf = transf_fundef f) eq p tp. + + Lemma transf_program_match: + forall prog, match_prog prog (transf_program prog). + Proof. + intros. eapply Linking.match_transform_program_contextual. auto. + Qed. Context (TRANSL : match_prog prog tprog). @@ -354,34 +360,34 @@ Section CORRECTNESS. Qed. Lemma function_ptr_translated: - forall b f l, + forall b f, Genv.find_funct_ptr ge b = Some f -> - Genv.find_funct_ptr tge b = Some (transf_fundef l f). + Genv.find_funct_ptr tge b = Some (transf_fundef f). Proof. intros. exploit (Genv.find_funct_ptr_match TRANSL); eauto. crush. Qed. Lemma sig_transf_function: - forall (f tf: fundef) l, - funsig (transf_fundef l f) = funsig f. + forall (f tf: fundef), + funsig (transf_fundef f) = funsig f. Proof using. unfold transf_fundef. unfold AST.transf_fundef; intros. destruct f. unfold transf_function. destruct_match. auto. auto. Qed. Lemma functions_translated: - forall (v: Values.val) (f: GibleSeq.fundef) l, + forall (v: Values.val) (f: GibleSeq.fundef), Genv.find_funct ge v = Some f -> - Genv.find_funct tge v = Some (transf_fundef l f). + Genv.find_funct tge v = Some (transf_fundef f). Proof using TRANSL. intros. exploit (Genv.find_funct_match TRANSL); eauto. simplify. eauto. Qed. Lemma find_function_translated: - forall ros rs f l, + forall ros rs f, find_function ge ros rs = Some f -> - find_function tge ros rs = Some (transf_fundef l f). + find_function tge ros rs = Some (transf_fundef f). Proof using TRANSL. Ltac ffts := match goal with | [ H: forall _, Val.lessdef _ _, r: Registers.reg |- _ ] => @@ -410,7 +416,7 @@ Section CORRECTNESS. replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved; eauto. symmetry; eapply Linking.match_program_main; eauto. eauto. erewrite sig_transf_function; eauto. econstructor. auto. auto. - Unshelve. exact None. + Unshelve. Qed. Lemma transf_final_states : @@ -507,8 +513,8 @@ Section CORRECTNESS. Qed. Lemma fn_all_eq : - forall f tf l, - transf_function l f = tf -> + forall f tf, + transf_function f = tf -> fn_stacksize f = fn_stacksize tf /\ fn_sig f = fn_sig tf /\ fn_params f = fn_params tf @@ -521,16 +527,16 @@ Section CORRECTNESS. Ltac func_info := match goal with - H: transf_function _ _ = _ |- _ => + H: transf_function _ = _ |- _ => let H2 := fresh "ALL_EQ" in - pose proof (fn_all_eq _ _ _ H) as H2; simplify + pose proof (fn_all_eq _ _ H) as H2; simplify end. Lemma step_cf_eq : - forall stk stk' f tf sp pc rs pr m cf s t pc' l, + forall stk stk' f tf sp pc rs pr m cf s t pc', step_cf_instr ge (State stk f sp pc rs pr m) cf t s -> Forall2 match_stackframe stk stk' -> - transf_function l f = tf -> + transf_function f = tf -> exists s', step_cf_instr tge (State stk' tf sp pc' rs pr m) cf t s' /\ match_states None s s'. Proof. @@ -778,19 +784,19 @@ Section CORRECTNESS. Qed. Lemma match_none_correct : - forall t s1' stk f sp pc rs m pr rs' m' bb pr' cf stk' l, + forall t s1' stk f sp pc rs m pr rs' m' bb pr' cf stk', (fn_code f) ! pc = Some bb -> SeqBB.step ge sp (Iexec (mki rs pr m)) bb (Iterm (mki rs' pr' m') cf) -> step_cf_instr ge (State stk f sp pc rs' pr' m') cf t s1' -> Forall2 match_stackframe stk stk' -> exists b' s2', - (plus step tge (State stk' (transf_function l f) sp pc rs pr m) t s2' \/ - star step tge (State stk' (transf_function l f) sp pc rs pr m) t s2' + (plus step tge (State stk' (transf_function f) sp pc rs pr m) t s2' \/ + star step tge (State stk' (transf_function f) sp pc rs pr m) t s2' /\ ltof (option SeqBB.t) measure b' None) /\ match_states b' s1' s2'. Proof. intros * H H0 H1 STK. - pose proof (transf_spec_correct f pc l) as X; inv X. + pose proof (transf_spec_correct f pc) as X; inv X. - apply sim_plus. rewrite H in H2. symmetry in H2. exploit step_cf_eq; eauto; simplify. do 3 econstructor. apply plus_one. econstructor; eauto. @@ -800,7 +806,7 @@ Section CORRECTNESS. destruct (cf_wf_dec x b' cf pc'); subst; simplify. + inv H1. exploit exec_if_conv; eauto; simplify. - apply sim_star. exists (Some b'). exists (State stk' (transf_function l f) sp pc rs pr m). + apply sim_star. exists (Some b'). exists (State stk' (transf_function f) sp pc rs pr m). simplify; try (unfold ltof; auto). apply star_refl. econstructor; auto. simplify. econstructor; eauto. @@ -820,21 +826,21 @@ Section CORRECTNESS. Qed. Lemma match_some_correct: - forall t s1' s f sp pc rs m pr rs' m' bb pr' cf stk' b0 pc1 rs1 p0 m1 l, + forall t s1' s f sp pc rs m pr rs' m' bb pr' cf stk' b0 pc1 rs1 p0 m1, step ge (State s f sp pc rs pr m) t s1' -> (fn_code f) ! pc = Some bb -> SeqBB.step ge sp (Iexec (mki rs pr m)) bb (Iterm (mki rs' pr' m') cf) -> step_cf_instr ge (State s f sp pc rs' pr' m') cf t s1' -> Forall2 match_stackframe s stk' -> (fn_code f) ! pc = Some b0 -> - sem_extrap (transf_function l f) pc1 sp (Iexec (mki rs pr m)) (Iexec (mki rs1 p0 m1)) b0 -> + sem_extrap (transf_function f) pc1 sp (Iexec (mki rs pr m)) (Iexec (mki rs1 p0 m1)) b0 -> (forall b', f.(fn_code)!pc1 = Some b' -> - exists tb, (transf_function l f).(fn_code)!pc1 = Some tb /\ if_conv_replace pc b0 b' tb) -> + exists tb, (transf_function f).(fn_code)!pc1 = Some tb /\ if_conv_replace pc b0 b' tb) -> step ge (State s f sp pc1 rs1 p0 m1) E0 (State s f sp pc rs pr m) -> exists b' s2', - (plus step tge (State stk' (transf_function l f) sp pc1 rs1 p0 m1) t s2' \/ - star step tge (State stk' (transf_function l f) sp pc1 rs1 p0 m1) t s2' /\ + (plus step tge (State stk' (transf_function f) sp pc1 rs1 p0 m1) t s2' \/ + star step tge (State stk' (transf_function f) sp pc1 rs1 p0 m1) t s2' /\ ltof (option SeqBB.t) measure b' (Some b0)) /\ match_states b' s1' s2'. Proof. intros * H H0 H1 H2 STK IS_B EXTRAP IS_TB SIM. @@ -863,14 +869,14 @@ Section CORRECTNESS. match goal with H: context[match_states] |- _ => inv H end. - eauto using match_some_correct. - eauto using match_none_correct. - - apply sim_plus. remember (transf_function l f) as tf. symmetry in Heqtf. func_info. + - apply sim_plus. remember (transf_function f) as tf. symmetry in Heqtf. func_info. exists (@None (list instr)). eexists. split. apply plus_one. constructor; eauto. rewrite <- H1. eassumption. rewrite <- H4. rewrite <- H2. econstructor; auto. - apply sim_plus. exists (@None (list instr)). eexists. split. apply plus_one. constructor; eauto. constructor; auto. - - apply sim_plus. remember (transf_function None f) as tf. symmetry in Heqtf. func_info. + - apply sim_plus. remember (transf_function f) as tf. symmetry in Heqtf. func_info. exists (@None (list instr)). inv STK. inv H7. eexists. split. apply plus_one. constructor. econstructor; auto. Qed. -- cgit