include "Clight/switchRemoval.ma". include "Clight/labelSimulation.ma". (* Combine the results about switch removal and labelling to provide a single result linking the input Clight program with the cost labelled Clight program. *) (* First lift the step result for labelling to a finite run of steps. *) lemma steps_with_labels : ∀ge,ge'. related_globals_gen … label_fundef ge ge' → ∀n,s1,s1',tr,s2. state_with_labels s1 s1' → after_n_steps (S n) … clight_exec ge s1 (λ_.true) (λtr',s'.〈tr',s'〉 = 〈tr,s2〉) → ∃m. after_n_steps (S m) … clight_exec ge' s1' (λ_.true) (λtr',s2'. trace_with_extra_labels tr tr' ∧ state_with_labels s2 s2'). #ge #ge' #RG #n elim n [ #s1 #s1' #tr #s2 #SWL #AFTER cases (after_1_step … AFTER) #tr' * #s' * * #F1 #ST1 #E destruct @(step_with_labels … RG … ST1) // | #n #IH #s1 #s1' #tr #s2 #SWL #AFTER cases (after_1_of_n_steps' … AFTER) #tr1 * #tr2 * #s' * * * #F1 #ST1 #E #AFTER' destruct cases (step_with_labels … RG … SWL ST1) #m1 #AF1 cases (after_n_result … AF1) #tr2 * #s2' * * #TWEL #SWL' #AF'' cases (IH … SWL' AFTER') #m2 #AF2 %{(m1+S m2)} @(after_n_m (S m1) (S m2) … AF2) // @(after_n_covariant … AF'') #tr #s #E destruct %{(refl ??)} #tr'' #s'' * #TWEL #SWL % /2/ ] qed. lemma interactive_switch_step : ∀ge,ge'. related_globals … fundef_switch_removal ge ge' → ∀s1,s1',o,k. exec_step ge s1 = Interact … o k → switch_state_sim ge s1 s1' → ∃k'. exec_step ge' s1' = Interact … o k' ∧ ∀i. ∃tr,s2. k i = Value … 〈tr,s2〉 ∧ ∃s2'. k' i = Value … 〈tr,s2'〉 ∧ switch_state_sim ge s2 s2'. #ge #ge' #RG #s1 #s1' #o #k #EX cases (exec_step_interaction … EX) #vf * #fn * #argtys * #retty * #vargs * #k' * #m #E destruct #S inversion S [1,3,4: #H1 #H2 #H3 #H4 destruct #H5 #H6 #H7 #H8 #H9 #H10 #H11 #H12 destruct #H13 #H14 #H15 #H16 #H17 #H18 #H19 #H20 #H21 #H22 #H23 #H24 #H25 #H26 #H27 #H28 #H29 destruct ] #s_vf #s_fd #s_args #s_k #s_k_ext #s_m #s_m_ext #s_writeable #s_me #s_H #E1 #E2 #E3 destruct whd in EX:(??%?); @(bindIO_res_interact ?????????? EX) -EX #vs #CHECK #EX whd in EX:(??%?); destruct % [2: % [ whd in ⊢ (??%?); >CHECK in ⊢ (??%?); whd in ⊢ (??%?); @refl | #i % [2: % [2: % [ @refl | % [2: % [ @refl | @(sws_returnstate … s_me) // ] | skip ] ] | skip ] ] ]| skip ] qed. lemma Value_eq_l : ∀tr,s,tr',s'. Value io_out io_in (trace×state) 〈tr,s〉 = Value io_out io_in (trace×state) 〈tr',s'〉 → tr = tr'. #tr #s #tr' #s' #E destruct % qed. lemma Value_eq_r : ∀tr,s,tr',s'. Value io_out io_in (trace×state) 〈tr,s〉 = Value io_out io_in (trace×state) 〈tr',s'〉 → s = s'. #tr #s #tr' #s' #E destruct % qed. (* Now build the coinduction simulation, taking one step to many after switch removal, then use steps_with_labels to take those to the cost labelled version. *) let corec build_exec (ge1,ge2,ge3:genv) (SRG:related_globals … fundef_switch_removal ge1 ge2) (RG:related_globals_gen … label_fundef ge2 ge3) (s1,s2,s3:state) (S1:switch_state_sim ge1 s1 s2) (S2:state_with_labels s2 s3) (NW:not_wrong state (exec_inf_aux … clight_fullexec ge1 (exec_step ge1 s1))) : sim_with_labels (exec_inf_aux … clight_fullexec ge1 (exec_step ge1 s1)) (exec_inf_aux … clight_fullexec ge3 (exec_step ge3 s3)) ≝ match NW return λe,NW. exec_inf_aux … clight_fullexec ?? = e → sim_with_labels e ? with [ nw_stop tr i s ⇒ ? | nw_step tr1 s1' e1 NW1 ⇒ ? | nw_interact o k NWk ⇒ ? ] (refl ? (exec_inf_aux … clight_fullexec ge1 (exec_step ge1 s1))). [ #E1 cases (exec_inf_stops_at_final ?? clight_fullexec … E1) #EX1 #F1 cases (switch_removal_correction … SRG … S1 EX1) #n1 #A1 cases (after_n_result … A1) #tr1' * #s1' * * #E destruct #S1' #A1' cases (steps_with_labels … RG … S2 A1') #n2 #A2 cases (after_inv clight_fullexec ????? A2) #tr' * #s' * * #TWL #S' lapply F1 whd in ⊢ (??%? → ?); >(switch_final_related … S1') >(final_related … S') #F3 whd in match (is_final … s'); >F3 * #tr2' #S2 @(swl_stop … S2 TWL) | #E1 cases (extract_step ?? clight_fullexec … E1) #EX1 #E1' cases (switch_removal_correction … SRG … S1 EX1) #n1 #A1 cases (after_n_result … A1) #tr1' * #s2' * * #E @(match E with [refl ⇒ ?]) #S1' #A1' cases (steps_with_labels … RG … S2 A1') #n #AF cases (after_inv clight_fullexec ????? AF) #tr' * #s' * * #TWL #S' whd in match (is_final … s'); lapply (refl ? (is_final s')) cases (is_final s') in ⊢ (???% → %); [ #NF #H whd in H; @(swl_steps … H) // @(match sym_eq … E1' return λe,E1'. sim_with_labels e ? with [refl ⇒ ?]) @(build_exec … SRG RG … S1' S') >E1' in NW1; // | #r <(final_related … S') <(switch_final_related … S1') change with (is_final … clight_fullexec ge1 s1') in ⊢ (??%? → ?); >(exec_e_step_not_final ?? clight_fullexec … E1) #E' destruct ] | #E1 cases (extract_interact ?? clight_fullexec … E1) #k' * #EX1 #Ek lapply NWk -NWk @(match sym_eq … Ek with [refl ⇒ ?]) #NWk cases (interactive_switch_step … SRG … EX1 S1) #k2' * #EX2 #H2 cases (interactive_step_with_labels … RG … EX2 S2) #k3' * #EX3 #H3 @(match sym_eq … EX3 with [refl ⇒ ?]) @(match sym_eq … (exec_inf_aux_unfold …) with [refl ⇒ ?]) (* XXX why is this necessary? *) whd in ⊢ (??%); @swl_interact #i cases (H3 i) #tr2 * #s2' * #K2 * #tr3 * #s3' * * #K3 #TR2 #S2' cases (H2 i) #tr1 * #s1' * #K1 * #s2x' * @(match sym_eq … K2 with [refl ⇒ ?]) #Ex @(match (Value_eq_r … Ex) with [refl ⇒ ?]) #S1' lapply (NWk i) @(match sym_eq … K1 with [refl ⇒ ?]) @(match sym_eq … (exec_inf_aux_unfold …) with [refl ⇒ ?]) whd in ⊢ (??% → ?%%); whd in match (is_final ?????); @(match sym_eq … (switch_final_related … S1') with [refl ⇒ ?]) @(match sym_eq … (final_related … S2') with [refl ⇒ ?]) @(match sym_eq … K2 with [refl ⇒ ?]) @(match sym_eq … (exec_inf_aux_unfold ?? clight_fullexec ge3 …) with [refl ⇒ ?]) whd in ⊢ (? → ??%); @(match sym_eq … K3 with [refl ⇒ ?]) whd in ⊢ (? → ??%); whd in match (is_final ?????); cases (is_final s3') [ whd in ⊢ (??% → ?%%); #NW' @(swl_steps … (steps_step …)) [ destruct // | @build_exec // inversion NW' [ #H1 #H2 #H3 #H4 #H5 destruct | #H7 #H8 #H9 #H10 #H11 #H12 destruct // | #H14 #H15 #H16 #H17 #H18 destruct ] ] | #r whd in ⊢ (??% → ?%%); #NW' @(swl_stop … (steps_stop …)) destruct // ] ] qed. (* And deal with the start of the program to tie up everything nicely. *) theorem switch_and_labelling_sim : ∀p. let e1 ≝ exec_inf … clight_fullexec p in let e2 ≝ exec_inf … clight_fullexec (\fst (clight_label (program_switch_removal p))) in not_wrong state e1 → sim_with_labels e1 e2. #p letin p' ≝ (program_switch_removal p) #NW cases (not_wrong_init clight_fullexec p NW) whd in ⊢ (% → ??%? → ?); #s1 #I1 cases (initial_state_in_switch_simulation … I1) #s' * * change with p' in match (program_switch_removal p); #I' #RG' #S' cases (initial_state_in_simulation … I') #s2 * #I2 whd in match (exec_inf ?? clight_fullexec p) in NW ⊢ %; letin ge1 ≝ (make_global ?? clight_fullexec p) in NW ⊢ %; change with (make_initial_state p) in match (make_initial_state ?? clight_fullexec p); >I1 whd in match (exec_inf ????); letin ge2 ≝ (make_global ?? clight_fullexec (\fst (clight_label p'))) change with (make_initial_state (\fst (clight_label p'))) in match (make_initial_state ?? clight_fullexec (\fst (clight_label p'))); >I2 whd in ⊢ (??% → ? → ?%%); >exec_inf_aux_unfold >exec_inf_aux_unfold whd in ⊢ (??% → ? → ?%%); whd in match (is_final ?????); >(initial_state_is_not_final … I1) whd in match (is_final ?????); >(initial_state_is_not_final … I2) whd in ⊢ (??% → ? → ?%%); #NW #S @(swl_steps E0 E0) [ 2: @steps_step | skip | // | @(build_exec … RG' … S' S) [ whd in match ge2; >unfold_clight_label @transform_program_gen_related | inversion NW [ #tr #i #s #E1 #E2 destruct | #trX #sX #eX #NWX #E1X #E2X destruct // | #H1 #H2 #H3 #H4 #H5 destruct ] ] ] qed.