Changeset 3549
- Timestamp:
- Apr 2, 2015, 3:44:19 PM (6 years ago)
- Location:
- LTS
- Files:
-
- 1 added
- 7 edited
Legend:
- Unmodified
- Added
- Removed
-
LTS/Final.ma
r3540 r3549 18 18 19 19 theorem cerco: 20 ∀l_p : label_params. 20 21 (* Given the operational semantics of a source program *) 21 ∀S1: abstract_status .22 ∀S1: abstract_status l_p. 22 23 23 24 (* Given the compiled assembly program *) … … 25 26 26 27 (* Let S2 be the operational semantics of the compiled code *) 27 let S2 ≝ asm_operational_semantics p p' prog in28 let S2 ≝ asm_operational_semantics p l_p p' prog in 28 29 29 30 (* If the simulation conditions between the source and compiled 30 31 code holds (i.e. if the compiler is correct) *) 31 ∀rel: relations S1 S2. simulation_conditions … rel →32 ∀rel: relations … S1 S2. simulation_conditions … rel → 32 33 33 34 (* Given an abstract interpretation framework for the compiled code *) 34 ∀R: asm_abstract_interpretation p p' prog.35 ∀R: asm_abstract_interpretation p p' … prog. 35 36 36 37 (* If the static analysis does not fail *) … … 40 41 equivalently, whose state after the initial labelled transition is s1 41 42 and whose state after the final labelled transition is s2 *) 42 ∀si,s1,s2,sn. ∀t: raw_trace S1 … si sn.43 ∀si,s1,s2,sn. ∀t: raw_trace … S1 … si sn. 43 44 measurable … s1 s2 … t → 44 45 45 46 (* For each target non I/O state si' in relation with the source state si *) 46 ∀si': S2. as_classify … si' ≠ cl_io → Srel ??rel si si' →47 ∀si': S2. as_classify … si' ≠ cl_io → Srel … rel si si' → 47 48 48 49 (* There exists a corresponding target trace starting from si' *) … … 75 76 >(proj1 … (static_analisys_ok … EQmap … Hmem)) 76 77 @(proj2 … (static_analisys_ok … EQmap … Hmem))] 77 # S1 #p #p' #prog letin S2 ≝ (asm_operational_semantics???) #rel #sim_conds78 #l_p #S1 #p #p' #prog letin S2 ≝ (asm_operational_semantics ????) #rel #sim_conds 78 79 #R #mT #map #EQmap #si #s1 #s2 #sn #t #Hmeas #si' #Hclass #Rsisi' 79 cases (simulates_measurable S1 S2 rel sim_conds … Hmeas … Rsisi')80 cases (simulates_measurable … S1 S2 rel sim_conds … Hmeas … Rsisi') 80 81 [2: #H cases Hclass >H #ABS cases(ABS (refl …)) ] 81 82 #sn' * #t' ** #Rsnsn' #EQcostlabs * #s1' * #s2' #Hmeas' … … 105 106 #A #l1 #l2 #x #Hmem normalize #H @mem_neq_zero_is_multiset <H @is_multiset_mem_neq_zero assumption 106 107 qed. 107 108 definition is_abelian ≝ λM : monoid.109 ∀x,y : M.op … M x y = op … M y x.110 108 111 109 lemma mem_list : ∀A : Type[0]. … … 211 209 ∀p_asm,p_asm',prog_asm. 212 210 (* Let S3 be the operational semantics of the compiled code *) 213 let S3 ≝ asm_operational_semantics p_asm p _asm' prog_asm in211 let S3 ≝ asm_operational_semantics p_asm p p_asm' prog_asm in 214 212 215 213 (* If the simulation conditions between the source with call post-labelled and compiled 216 214 code holds (i.e. if the compiler is correct after the first pass) *) 217 ∀rel: relations S2 S3. simulation_conditions … rel →215 ∀rel: relations … S2 S3. simulation_conditions … rel → 218 216 219 217 (* let REL be semantical relation between the states of the source program and … … 223 221 224 222 (* Given an abstract interpretation framework for the compiled code *) 225 ∀R: asm_abstract_interpretation p_asm p_asm' prog_asm.223 ∀R: asm_abstract_interpretation p_asm p_asm' … prog_asm. 226 224 227 225 (* If the abstract instructions monoid is abelian *) … … 234 232 equivalently, whose state after the initial labelled transition is s1 235 233 and whose state after the final labelled transition is s2 *) 236 ∀si,s1,s2,sn. ∀t: raw_trace S1 … si sn.234 ∀si,s1,s2,sn. ∀t: raw_trace … S1 … si sn. 237 235 measurable … s1 s2 … t → 238 236 -
LTS/Lang_meas.ma
r3542 r3549 13 13 (**************************************************************************) 14 14 15 include "Lang uage.ma".15 include "Lang_corr.ma". 16 16 17 17 … … 21 21 ∀s1,s2,s1' : state p.∀abs_top,abs_tail. 22 22 execute_l … p' (env … prog) l s1 s2 → 23 is_costlabelled_act l →24 (is_call_act l → is_call_post_label(operational_semantics p p' prog) … s1) →23 is_costlabelled_act p l → 24 (is_call_act p l → is_call_post_label p (operational_semantics p p' prog) … s1) → 25 25 state_rel … m keep abs_top abs_tail s1 s1' → abs_top = nil ? ∧ 26 (is_call_act l → is_call_post_label(operational_semantics p p' t_prog) … s1').26 (is_call_act p l → is_call_post_label p (operational_semantics p p' t_prog) … s1'). 27 27 #p #p' #prog @pair_elim * #t_prog #m #keep #EQt_prog normalize nodelta 28 28 #l #s1 #s2 #s1' #abs_top #abs_tail #H elim H -s1 -s2 -l … … 33 33 >EQcont in EQcheck; whd in ⊢ (??%? → ?); inversion(cont … s1') normalize nodelta 34 34 [ #_ #EQ destruct] * #lab' #code_cont' #stack' #_ #EQcont' 35 change with (check_continuations ????? ) in match (foldr2 ???????);35 change with (check_continuations ??????) in match (foldr2 ???????); 36 36 inversion(check_continuations ?????) [ #_ normalize #EQ destruct] 37 37 ** #H2 #a_top1 #a_tail1 #EQcheck >m_return_bind normalize nodelta … … 134 134 #a_top3 #istr3 #EQi_true' inversion(?==?) #EQget1 inversion(?==?) #EQget2 normalize nodelta 135 135 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct #_ #_ #EQ destruct % // * #f * #c #EQ destruct 136 | cases daemon 136 | #s1 #s2 #exp #ltrue #i_true #lfalse #i_false #instrs #new_m #EQcode_s1 #EQeval #EQcont_s2 #EQ1 #EQ2 destruct 137 #Hio1 #Hio2 * #_ whd in match state_rel; normalize nodelta inversion(check_continuations ?????) normalize nodelta 138 [ #_ *] ** #H1 #a_top #a_tail #EQcheck normalize nodelta **** #HH1 >EQcode_s1 inversion(code ? s1') 139 [1,2,3,5,6,7: 140 [ | #ret | #seq #opt_l #instr #_ | #cond #ltrue #itrue #lfalse #ifalse #_ #_| #f #pars #opt_l #instr #_ 141 | #lin #io #lout #instr #_ ] 142 #_ whd in ⊢ (??%% → ?); 143 [ | 144 | cases(call_post_clean ?????); normalize nodelta [2: #x cases opt_l normalize nodelta [| #label cases(?==?) normalize nodelta]] 145 | cases(call_post_clean ?????) normalize nodelta 146 [| #x whd in ⊢ (??%? → ?); cases(call_post_clean ?????) normalize nodelta 147 [| #y cases(?∧?) normalize nodelta 148 ] 149 ] 150 | cases(call_post_clean ?????) normalize nodelta 151 [| #x cases opt_l normalize nodelta [| #lbls cases(memb ???) normalize nodelta 152 [ cases (?==?) normalize nodelta ] 153 ]] 154 | cases(call_post_clean ?????) normalize nodelta 155 [| #x cases(? ∧ ?) normalize nodelta ] 156 ] 157 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct 158 ] 159 #exp' #ltrue' #i_true' #lfalse' #i_false' #instrs' #_ #_ #_ #EQcode_s1' whd in ⊢ (??%% → ?); 160 inversion(call_post_clean ?????) normalize nodelta [#_ #EQ destruct] * #a_top1 #istr1 #EQinstr' 161 whd in ⊢ (??%%→ ?); inversion(call_post_clean ?????) normalize nodelta [ #_ #EQ destruct] * #a_top2 #istr2 162 #EQi_false' whd in ⊢ (??%% → ?); inversion(call_post_clean ?????) normalize nodelta [#_ #EQ destruct] * 163 #a_top3 #istr3 #EQi_true' inversion(?==?) #EQget1 inversion(?==?) #EQget2 normalize nodelta 164 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct #_ #_ #EQ destruct % // * #f * #c #EQ destruct 137 165 | cases daemon 138 166 | cases daemon … … 156 184 cases(?==?) normalize nodelta 157 185 [2: #EQ destruct] whd in ⊢ (??%% → ?); #EQ1 lapply(eq_to_jmeq ??? EQ1) -EQ1 #EQ1 destruct #_#_ #EQ1 destruct % // 158 #_ whd in match (is_call_post_label ?? ); >EQcode_s1' %186 #_ whd in match (is_call_post_label ???); >EQcode_s1' % 159 187 ] 160 188 #ABS whd in ⊢ (??%% → ?); #EQ1 lapply(eq_to_jmeq ??? EQ1) -EQ1 #EQ1 destruct lapply(Hcall ?) [ %{f} %{(f_lab … env_it)} %] 161 whd in ⊢ (% → ?); whd in match (is_call_post_label ?? ); normalize nodelta >EQcode *189 whd in ⊢ (% → ?); whd in match (is_call_post_label ???); normalize nodelta >EQcode * 162 190 ] 163 191 | #s1 #s2 #r_t #mem #stack * [|#lbl] #i #EQcode #EQcont #EQ destruct #_ #_ #EQeval #EQ1 #EQ2 destruct * #_ … … 168 196 >EQcont in EQcheck; inversion (cont ? s1') [ #_ whd in ⊢ (??%% → ?); #EQ destruct] 169 197 * #LAB #instrs #stack' #_ #EQcont_s1' whd in match check_continuations; normalize nodelta 170 whd in match (foldr2 ???????); change with (check_continuations ????? ) in match (foldr2 ???????);198 whd in match (foldr2 ???????); change with (check_continuations ??????) in match (foldr2 ???????); 171 199 cases(check_continuations ?????) normalize nodelta [#EQ destruct] ** #H2 #a_top1 #a_tail1 172 200 normalize nodelta whd in ⊢ (??%? → ?); cases(call_post_clean ?????) normalize nodelta [ #EQ destruct cases HH1] … … 196 224 let 〈t_prog,m,keep〉 ≝ trans_prog … prog in 197 225 ∀s1,s2,s1' : state p.∀abs_top,abs_tail. 198 ∀l.is_costlabelled_act l →226 ∀l.is_costlabelled_act p l → 199 227 ∀prf :execute_l p p' (env … prog) l s1 s2. 200 228 state_rel … m keep abs_top abs_tail s1 s1' → … … 202 230 execute_l p p' (env … t_prog) l s1' s2' ∧ 203 231 state_rel … m keep abs_top' abs_tail s2 s2' ∧ 204 map_labels_on_trace … m … (actionlabel_to_costlabel l) =205 actionlabel_to_costlabel l @ abs_top'.232 map_labels_on_trace … m … (actionlabel_to_costlabel p l) = 233 actionlabel_to_costlabel p l @ abs_top'. 206 234 #p #p' #prog #no_dup @pair_elim * #t_prog #m #keep #EQt_prog normalize nodelta 207 235 #s1 #s2 #s1' #abs_top #abs_tail #l #Hl #H lapply Hl -Hl elim H -s1 -s2 -l … … 209 237 #Hlbl #cost_lbl whd in match state_rel in ⊢ (% → ?); normalize nodelta inversion(check_continuations ?????) 210 238 normalize nodelta [ #_ *] ** #H1 #a_top #a_tail >EQcont_s1 inversion(cont … s1') [ #_ whd in ⊢ (??%% → ?); #EQ destruct] 211 * #lbl' #i' #stack' #_ #EQcont_s1' whd in ⊢ (??%% → ?); change with (check_continuations ????? ) in match (foldr2 ???????);239 * #lbl' #i' #stack' #_ #EQcont_s1' whd in ⊢ (??%% → ?); change with (check_continuations ??????) in match (foldr2 ???????); 212 240 inversion(check_continuations ?????) normalize nodelta [ #_ #EQ destruct] ** #H2 #a_top1 #a_tail1 #EQcheck 213 241 normalize nodelta inversion(call_post_clean ?????) normalize nodelta [#_ whd in ⊢ (??%% → ?); #EQ destruct ***** ] … … 247 275 * #genlab #cleaned #EQcleaned inversion(?∧?) normalize nodelta #EQget whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ 248 276 #EQ destruct #EQstore_s1_s1' #EQioinfo #EQ destruct %{[]} 249 %{(mk_state … (EMPTY …) (〈cost_act (Some ? lout),i'〉::(cont … s1')) (store … s2) true)} % [%277 %{(mk_state … (EMPTY …) (〈cost_act p (Some ? lout),i'〉::(cont … s1')) (store … s2) true)} % [% 250 278 [ @(io_in … EQcode_s1') // ] 251 whd >EQcont_s2 whd in match (check_continuations ????? ); change with (check_continuations?????) in match (foldr2 ???????);279 whd >EQcont_s2 whd in match (check_continuations ??????); change with (check_continuations ??????) in match (foldr2 ???????); 252 280 >EQcheck normalize nodelta >EQcleaned normalize nodelta % // % // % // % [2: >EQcode_s2 %] % 253 281 [2: cases(andb_Prop_true … (bool_true … EQget)) #H #_ >(\P (bool_true … H)) % ] % // % //] … … 265 293 let 〈t_prog,m,keep〉 ≝ trans_prog … prog in 266 294 ∀si,s1,s2,sn,si' : state p.∀abs_top,abs_tail. 267 ∀t : raw_trace (operational_semantics … p' prog) si sn.268 measurable (operational_semantics …p' prog) … s1 s2 … t →295 ∀t : raw_trace p (operational_semantics … p' prog) si sn. 296 measurable p (operational_semantics p p' prog) … s1 s2 … t → 269 297 state_rel … m keep abs_top abs_tail si si' → 270 ∃abs_top',abs_tail'.∃sn'.∃t' : raw_trace (operational_semantics … p' t_prog) si' sn'.298 ∃abs_top',abs_tail'.∃sn'.∃t' : raw_trace p (operational_semantics … p' t_prog) si' sn'. 271 299 state_rel … m keep abs_top' abs_tail' sn sn' ∧ 272 is_permutation ? (map_labels_on_trace m (chop … (get_costlabels_of_trace … t)))300 is_permutation ? (map_labels_on_trace … m (chop … (get_costlabels_of_trace … t))) 273 301 (chop … (get_costlabels_of_trace … t')) ∧ 274 302 ∃s1',s2'. 275 measurable (operational_semantics … p' t_prog) … s1' s2' … t'.303 measurable p (operational_semantics … p' t_prog) … s1' s2' … t'. 276 304 #p #p' #prog #no_dup @pair_elim * #t_prog #m #keep #EQt_prog normalize nodelta 277 305 #si #s1 #s3 #sn #si' #abs_top #abs_tail #t * #s0 * #s2 * #ti0 * #t12 * #t3n * #l1 * #l2 … … 300 328 >append_nil cases(actionlabel_ok … Hl2) #c2 #EQc2 >EQc2 <associative_append <associative_append 301 329 <associative_append >chop_append_singleton <associative_append <associative_append >chop_append_singleton 302 >append_nil whd in match (get_costlabels_of_trace ??? (t_ind(operational_semantics p p' t_prog) … prf1' (t_base …)));330 >append_nil whd in match (get_costlabels_of_trace ???? (t_ind ? (operational_semantics p p' t_prog) … prf1' (t_base …))); 303 331 >append_nil >map_labels_on_trace_append >EQmap_l1 >associative_append @is_permutation_append_left_cancellable 304 332 lapply(top_move_source_is_empty … p' prog) >EQt_prog normalize nodelta #H -
LTS/Language.ma
r3535 r3549 18 18 include "../src/utilities/option.ma". 19 19 include "basics/jmeq.ma". 20 include "utils.ma". 20 21 21 22 discriminator option. … … 31 32 32 33 33 inductive Instructions (p : instr_params) : Type[0] ≝ 34 | EMPTY : Instructions p 35 | RETURN : return_type p → Instructions p 36 | SEQ : (seq_instr p) → option NonFunctionalLabel → Instructions p → Instructions p 37 | COND : (cond_instr p) → NonFunctionalLabel → Instructions p → 38 NonFunctionalLabel → Instructions p → Instructions p → 39 Instructions p 40 | LOOP : (loop_instr p) → NonFunctionalLabel → Instructions p → 41 NonFunctionalLabel → Instructions p → Instructions p 42 | CALL : FunctionName → (act_params_type p) → option ReturnPostCostLabel → 43 Instructions p → Instructions p 44 | IO : NonFunctionalLabel → (io_instr p) → NonFunctionalLabel → Instructions p → 45 Instructions p. 46 47 let rec eq_instructions (p : instr_params) (i : Instructions p) 48 on i : (Instructions p) → bool ≝ 34 inductive Instructions (p : instr_params) 35 (l_p :label_params) : Type[0] ≝ 36 | EMPTY : Instructions p l_p 37 | RETURN : return_type p → Instructions p l_p 38 | SEQ : (seq_instr p) → option (NonFunctionalLabel l_p) → Instructions p l_p → Instructions p l_p 39 | COND : (cond_instr p) → (NonFunctionalLabel l_p) → Instructions p l_p → 40 (NonFunctionalLabel l_p) → Instructions p l_p → Instructions p l_p → 41 Instructions p l_p 42 | LOOP : (loop_instr p) → NonFunctionalLabel l_p → Instructions p l_p → 43 NonFunctionalLabel l_p → Instructions p l_p → Instructions p l_p 44 | CALL : FunctionName → (act_params_type p) → option (ReturnPostCostLabel l_p) → 45 Instructions p l_p → Instructions p l_p 46 | IO : NonFunctionalLabel l_p → (io_instr p) → NonFunctionalLabel l_p → Instructions p l_p → 47 Instructions p l_p. 48 49 let rec eq_instructions (p : instr_params) (l_p : label_params) (i : Instructions p l_p) 50 on i : (Instructions p l_p) → bool ≝ 49 51 match i with 50 52 [ EMPTY ⇒ λi'.match i' with [ EMPTY ⇒ true | _ ⇒ false ] … … 53 55 [ SEQ y lab' instr' ⇒ x == y ∧ eq_instructions … instr instr' ∧ 54 56 match lab with [ None ⇒ match lab' with [ None ⇒ true | _ ⇒ false ] 55 | Some l1 ⇒ match lab' with [Some l2 ⇒ eq_nf_label l1 l2 | _ ⇒ false]57 | Some l1 ⇒ match lab' with [Some l2 ⇒ eq_nf_label … l1 l2 | _ ⇒ false] 56 58 ] 57 59 | _ ⇒ false … … 59 61 | COND x ltrue i1 lfalse i2 i3 ⇒ λi'.match i' with 60 62 [ COND y ltrue' i1' lfalse' i2' i3' ⇒ 61 x == y ∧ eq_nf_label ltrue ltrue' ∧62 eq_instructions … i1 i1' ∧ eq_nf_label lfalse lfalse' ∧63 x == y ∧ eq_nf_label … ltrue ltrue' ∧ 64 eq_instructions … i1 i1' ∧ eq_nf_label … lfalse lfalse' ∧ 63 65 eq_instructions … i2 i2' ∧ eq_instructions … i3 i3' 64 66 | _ ⇒ false … … 66 68 | LOOP x ltrue i1 lfalse i2 ⇒ λi'.match i' with 67 69 [ LOOP y ltrue' i1' lfalse' i2' ⇒ x == y ∧ 68 eq_instructions … i1 i1' ∧ eq_nf_label ltrue ltrue' ∧70 eq_instructions … i1 i1' ∧ eq_nf_label … ltrue ltrue' ∧ 69 71 eq_instructions … i2 i2' 70 72 | _ ⇒ false … … 74 76 act_p == act_p' ∧ eq_instructions … i1 i1' ∧ 75 77 match r_lb with [ None ⇒ match r_lb' with [None ⇒ true | _ ⇒ false] 76 | Some z ⇒ match r_lb' with [Some w ⇒ eq_return_cost_lab z w | _ ⇒ false ]78 | Some z ⇒ match r_lb' with [Some w ⇒ eq_return_cost_lab … z w | _ ⇒ false ] 77 79 ] 78 80 | _ ⇒ false 79 81 ] 80 82 | IO lin io lout i1 ⇒ λi'.match i' with 81 [ IO lin' io' lout' i1' ⇒ eq_nf_label lin lin' ∧ io == io' ∧82 eq_nf_label lout lout' ∧ eq_instructions … i1 i1'83 [ IO lin' io' lout' i1' ⇒ eq_nf_label … lin lin' ∧ io == io' ∧ 84 eq_nf_label … lout lout' ∧ eq_instructions … i1 i1' 83 85 | _ ⇒ false 84 86 ] … … 117 119 }. 118 120 119 record env_item (p : env_params) (p' : instr_params) : Type[0] ≝121 record env_item (p : env_params) (p' : instr_params) (l_p : label_params) : Type[0] ≝ 120 122 { f_sig :> signature p p' 121 ; f_lab : CallCostLabel 122 ; f_body : Instructions p' 123 ; f_lab : CallCostLabel l_p 124 ; f_body : Instructions p' l_p 123 125 }. 124 126 … … 126 128 { i_pars :> instr_params 127 129 ; e_pars :> env_params 130 ; l_pars :> label_params 128 131 ; store_type : DeqSet 129 132 }. 130 133 131 134 record state (p : state_params) : Type[0] ≝ 132 { code : Instructions p 133 ; cont : list (ActionLabel × (Instructionsp))135 { code : Instructions p p 136 ; cont : list (ActionLabel p × (Instructions p p)) 134 137 ; store : store_type p 135 138 ; io_info : bool … … 149 152 150 153 151 let rec lookup (p : env_params) (p' : instr_params) (l : list (env_item p p'))152 on l : FunctionName → option (env_item p p' ) ≝154 let rec lookup (p : env_params) (p' : instr_params) (l_p : label_params) (l : list (env_item p p' l_p)) 155 on l : FunctionName → option (env_item p p' l_p) ≝ 153 156 match l with 154 157 [ nil ⇒ λ_.None ? … … 158 161 ]. 159 162 160 definition is_ret_call_act : ActionLabel → Prop ≝ 161 λa.match a with [ret_act _ ⇒ True | call_act _ _ ⇒ True | _ ⇒ False ]. 162 163 inductive execute_l (p : state_params) (p' : sem_state_params p) (env : list (env_item p p)) : 164 ActionLabel → relation (state p) ≝ 165 | empty : ∀st,st',hd,tl.(code ? st) = (EMPTY p)→ (cont ? st) = hd :: tl → 163 inductive execute_l (p : state_params) (p' : sem_state_params p) (env : list (env_item p p p)) : 164 ActionLabel p → relation (state p) ≝ 165 | empty : ∀st,st',hd,tl.(code ? st) = (EMPTY p p)→ (cont ? st) = hd :: tl → 166 166 (code ? st') = \snd hd → (cont … st') = tl → (store … st) = (store … st') → 167 (io_info … st = true → is_non_silent_cost_act (\fst hd)) → (io_info … st') = false → ¬ is_ret_call_act(\fst hd) → execute_l … (\fst hd) st st'167 (io_info … st = true → is_non_silent_cost_act … (\fst hd)) → (io_info … st') = false → ¬ is_ret_call_act … (\fst hd) → execute_l … (\fst hd) st st' 168 168 | seq_sil : ∀st,st',i,cd,s,opt_l.(code ? st) = SEQ … i opt_l cd → 169 169 eval_seq … p' i (store … st) = return s → (code ? st') = cd → 170 170 (cont … st) = (cont … st') → (store … st') = s → 171 io_info … st = false → io_info ? st' = false → execute_l … (cost_act opt_l) st st'171 io_info … st = false → io_info ? st' = false → execute_l … (cost_act … opt_l) st st' 172 172 | cond_true : ∀st,st',exp,ltrue,i_true,lfalse,i_false,cd,new_m. 173 173 (code ? st) = COND … exp ltrue i_true lfalse i_false cd → eval_cond_cond … p' exp (store … st) = return 〈true,new_m〉 → 174 cont ? st' = 〈cost_act (None ?),cd〉 ::(cont … st) → code … st' = i_true → store … st' = new_m →175 io_info … st = false → io_info … st' = false → execute_l … (cost_act (Some ? ltrue)) st st'174 cont ? st' = 〈cost_act … (None ?),cd〉 ::(cont … st) → code … st' = i_true → store … st' = new_m → 175 io_info … st = false → io_info … st' = false → execute_l … (cost_act … (Some ? ltrue)) st st' 176 176 | cond_false : ∀st,st',exp,ltrue,i_true,lfalse,i_false,cd,new_m. 177 177 (code ? st) = COND … exp ltrue i_true lfalse i_false cd → eval_cond_cond … p' exp (store … st) = return 〈false,new_m〉 → 178 cont ? st' = 〈cost_act (None ?),cd〉 ::(cont … st) → code … st' = i_false → store … st' = new_m →179 io_info … st = false → io_info … st' = false → execute_l … (cost_act (Some ? lfalse)) st st'178 cont ? st' = 〈cost_act … (None ?),cd〉 ::(cont … st) → code … st' = i_false → store … st' = new_m → 179 io_info … st = false → io_info … st' = false → execute_l … (cost_act … (Some ? lfalse)) st st' 180 180 | loop_true : ∀st,st',exp,ltrue,i_true,lfalse,i_false,new_m. 181 181 code ? st = LOOP … exp ltrue i_true lfalse i_false → eval_loop_cond … p' exp (store … st) = return 〈true,new_m〉 → 182 cont ? st' = 〈cost_act (None ?),LOOP … exp ltrue i_true lfalse i_false〉 :: (cont … st) →182 cont ? st' = 〈cost_act … (None ?),LOOP … exp ltrue i_true lfalse i_false〉 :: (cont … st) → 183 183 code … st' = i_true → store … st' = new_m → io_info … st = false → io_info … st' = false → 184 execute_l … (cost_act (Some ? ltrue)) st st'184 execute_l … (cost_act … (Some ? ltrue)) st st' 185 185 | loop_false : ∀st,st',exp,ltrue,i_true,lfalse,i_false,new_m. 186 186 code ? st = LOOP … exp ltrue i_true lfalse i_false → eval_loop_cond … p' exp (store … st) = return 〈false,new_m〉 → 187 187 cont ? st' = cont … st → code … st' = i_false → store … st' = new_m → 188 io_info … st = false → io_info … st' = false → execute_l … (cost_act (Some ? lfalse)) st st'188 io_info … st = false → io_info … st' = false → execute_l … (cost_act … (Some ? lfalse)) st st' 189 189 | io_in : ∀st,st',lin,io,lout,cd,mem.(code ? st) = IO … lin io lout cd → 190 eval_io … p' io (store … st) = return mem → code ? st' = EMPTY p →191 cont … st' = 〈cost_act (Some ? lout),cd〉 :: (cont … st) → store … st' = mem →192 io_info … st' = true → execute_l … (cost_act (Some ? lin)) st st'190 eval_io … p' io (store … st) = return mem → code ? st' = EMPTY p p → 191 cont … st' = 〈cost_act … (Some ? lout),cd〉 :: (cont … st) → store … st' = mem → 192 io_info … st' = true → execute_l … (cost_act … (Some ? lin)) st st' 193 193 | call : ∀st,st',f,act_p,r_lb,cd,mem,env_it.(code ? st) = CALL … f act_p r_lb cd → 194 194 lookup … env f = return env_it → … … 196 196 store ? st' = mem → code … st' = f_body … env_it → 197 197 cont … st' = 198 〈(ret_act r_lb),cd〉 :: (cont … st) →198 〈(ret_act … r_lb),cd〉 :: (cont … st) → 199 199 io_info … st = false → (io_info … st') = false → 200 execute_l … (call_act f (f_lab ??env_it)) st st'200 execute_l … (call_act … f (f_lab … env_it)) st st' 201 201 | ret_instr : ∀st,st',r_t,mem,tl',rb,cd.code ? st = RETURN … r_t → 202 cont … st = 〈ret_act rb,cd〉 :: tl' → cont ? st' = tl' →202 cont … st = 〈ret_act … rb,cd〉 :: tl' → cont ? st' = tl' → 203 203 io_info … st = false → io_info ? st' = false → 204 204 eval_after_return … p' r_t (store … st) = return mem → code … st' = cd → 205 store … st' = mem → execute_l … (ret_act rb) st st'.205 store … st' = mem → execute_l … (ret_act … rb) st st'. 206 206 207 let rec get_labels_of_code (p : instr_params) ( i : Instructions p) on i : list CostLabel≝207 let rec get_labels_of_code (p : instr_params) (l_p : label_params) (i : Instructions p l_p) on i : list (CostLabel l_p) ≝ 208 208 match i with 209 209 [ EMPTY ⇒ [ ] 210 210 | RETURN x ⇒ [ ] 211 211 | SEQ x lab instr ⇒ let ih ≝ get_labels_of_code … instr in 212 match lab with [ None ⇒ ih | Some lbl ⇒ a_non_functional_label lbl :: ih ]212 match lab with [ None ⇒ ih | Some lbl ⇒ a_non_functional_label … lbl :: ih ] 213 213 | COND x ltrue i1 lfalse i2 i3 ⇒ 214 214 let ih3 ≝ get_labels_of_code … i3 in … … 219 219 let ih2 ≝ get_labels_of_code … i2 in 220 220 let ih1 ≝ get_labels_of_code … i1 in 221 a_non_functional_label ltrue :: a_non_functional_labellfalse :: (ih1 @ ih2)221 a_non_functional_label … ltrue :: a_non_functional_label … lfalse :: (ih1 @ ih2) 222 222 | CALL f act_p r_lb i1 ⇒ 223 223 let ih1 ≝ get_labels_of_code … i1 in 224 match r_lb with [ None ⇒ ih1 | Some lbl ⇒ a_return_post lbl :: ih1]224 match r_lb with [ None ⇒ ih1 | Some lbl ⇒ a_return_post … lbl :: ih1] 225 225 | IO lin io lout i1 ⇒ 226 226 let ih1 ≝ get_labels_of_code … i1 in 227 a_non_functional_label lin :: a_non_functional_labellout :: ih1227 a_non_functional_label … lin :: a_non_functional_label … lout :: ih1 228 228 ]. 229 230 include "basics/lists/listb.ma".231 include "../src/utilities/hide.ma".232 233 let rec no_duplicates (A : DeqSet) (l : list A) on l : Prop ≝234 match l with235 [ nil ⇒ True236 | cons x xs ⇒ ¬ (bool_to_Prop (x ∈ xs)) ∧ no_duplicates … xs237 ].238 239 lemma no_duplicates_append_r : ∀A : DeqSet.∀l1,l2 : list A.no_duplicates … (l1 @ l2) →240 no_duplicates … l2.241 #A #l1 elim l1 // #x #xs normalize #IH #l2 * /2/242 qed.243 244 lemma no_duplicates_append_l : ∀A : DeqSet.∀l1,l2 : list A.no_duplicates … (l1 @ l2) →245 no_duplicates … l1.246 #A #l1 elim l1 // #x #xs normalize #IH #l2 * #H1 #H2 % [2: /2/ ]247 inversion(x ∈ xs @l2) in H1; normalize [ #_ * #H @⊥ @H %] #H1 #_248 % inversion(x ∈ xs) normalize [2: //] #H3 #_ >(memb_append_l1 … H3) in H1;249 #EQ destruct(EQ)250 qed.251 229 252 record Program (p : env_params) (p' : instr_params) : Type[0] ≝253 { env : list (env_item p p' )254 ; main : Instructions p' 230 record Program (p : env_params) (p' : instr_params) (l_p : label_params) : Type[0] ≝ 231 { env : list (env_item p p' l_p) 232 ; main : Instructions p' l_p 255 233 }. 256 234 257 235 258 definition no_duplicates_labels : ∀p,p'.Program p p' → Prop ≝ 259 λp,p',prog. 236 237 definition no_duplicates_labels : ∀p,p',l_p.Program p p' l_p → Prop ≝ 238 λp,p',l_p,prog. 260 239 no_duplicates … 261 (foldr … (λitem,acc.((a_call (f_lab … item)) :: get_labels_of_code … (f_body … item)) @ acc) (get_labels_of_code … (main … prog)) (env … prog)). 240 (foldr … 241 (λitem,acc.((a_call … (f_lab … item)) :: get_labels_of_code … (f_body … item)) @ acc) 242 (get_labels_of_code … (main … prog)) (env … prog)). 262 243 263 244 lemma no_duplicates_domain_of_fun: 264 ∀p,p', prog.no_duplicates_labels … prog →265 ∀f,env_it.lookup p p' (env … prog) f = return env_it →245 ∀p,p',l_p,prog.no_duplicates_labels … prog → 246 ∀f,env_it.lookup p p' l_p (env … prog) f = return env_it → 266 247 no_duplicates … (get_labels_of_code … (f_body … env_it)). 267 #p #p' * #env elim env [ #main normalize #_ #f #env_it #EQ destruct(EQ)]248 #p #p' #l_p * #env elim env [ #main normalize #_ #f #env_it #EQ destruct(EQ)] 268 249 #x #xs #IH #main whd in ⊢ (% → ?); whd in match (foldr ?????); #H #f #env_it 269 250 whd in ⊢ (??%? → ?); @eq_function_name_elim normalize nodelta … … 279 260 ]. 280 261 281 definition operational_semantics : ∀p : state_params.∀p'.Program p p → abstract_status≝282 λp,p',prog.mk_abstract_status 262 definition operational_semantics : ∀p : state_params.∀p'.Program p p p → abstract_status p ≝ 263 λp,p',prog.mk_abstract_status … 283 264 (state p) 284 265 (execute_l ? p' (env … prog)) … … 377 358 qed. 378 359 379 let rec eqb_list (D : DeqSet) (l1 : list D) on l1 : list D → bool ≝ 380 match l1 with 381 [ nil ⇒ λl2.match l2 with [nil ⇒ true | _ ⇒ false ] 382 | cons x xs ⇒ λl2.match l2 with [ cons y ys ⇒ x == y ∧ eqb_list … xs ys | _ ⇒ false ] 383 ]. 384 385 definition DeqSet_List : DeqSet → DeqSet ≝ 386 λX.mk_DeqSet (list X) (eqb_list …) ?. 387 #x elim x [ * /2 by refl, conj/ #y #ys normalize % #EQ destruct] -x 388 #x #xs #IH * [ normalize % #EQ destruct] #y #ys normalize % inversion(x == y) 389 #EQ normalize nodelta 390 [ #H >(proj1 … (IH ys) H) @eq_f2 [2: %] @(proj1 … (eqb_true …)) assumption 391 | #EQ destruct 392 | #EQ1 destruct @(proj2 … (IH …)) % 393 | #EQ1 destruct <EQ @(proj2 … (eqb_true …)) % 394 ] 395 qed. 396 397 unification hint 0 ≔ C; 398 X ≟ DeqSet_List C 399 (* ---------------------------------------- *) ⊢ 400 list C ≡ carr X. 401 402 403 unification hint 0 ≔ D,p1,p2; 404 X ≟ DeqSet_List D 405 (* ---------------------------------------- *) ⊢ 406 eqb_list D p1 p2 ≡ eqb X p1 p2. 407 408 definition associative_list : DeqSet → Type[0] → Type[0] ≝ 409 λA,B.list (A × (list B)). 410 411 let rec update_list (A : DeqSet) (B : Type[0]) (l : associative_list A B) 412 on l : A → list B → associative_list A B ≝ 413 λa,b.match l with 414 [ nil ⇒ [〈a,b〉] 415 | cons x xs ⇒ if (a == (\fst x)) then 〈a,b〉 :: xs 416 else x :: (update_list … xs a b) 417 ]. 418 419 let rec get_element (A :DeqSet) (B : Type[0]) (l: associative_list A B) on l : A → list B ≝ 420 λa.match l with [ nil ⇒ nil ? 421 | cons x xs ⇒ if (a == \fst x) then \snd x else get_element … xs a 422 ]. 423 424 let rec domain_of_associative_list (A :DeqSet) (B : Type[0]) (l: associative_list A B) on l : list A ≝ 425 match l with 426 [ nil ⇒ [] 427 | cons x xs ⇒ \fst x :: domain_of_associative_list … xs 428 ]. 429 430 lemma get_element_append_l: 431 ∀A,B. ∀l1,l2: associative_list A B. ∀x. 432 x ∈ domain_of_associative_list … l1 → 433 get_element … (l1@l2) x = get_element … l1 x. 434 #A #B #l1 elim l1 normalize [ #l2 #x * ] #hd #tl #IH #l2 #x cases (dec_eq … x (\fst hd)) 435 #H [ >(\b H) | >(\bf H) ] normalize /2/ 436 qed. 437 438 lemma get_element_append_r: 439 ∀A,B. ∀l1,l2: associative_list A B. ∀x. 440 ¬ (bool_to_Prop (x ∈ domain_of_associative_list … l1)) → 441 get_element ?? (l1@l2) x = get_element … l2 x. 442 #A #B #l1 elim l1 normalize [ #l2 #x // ] #hd #tl #IH #l2 #x cases (dec_eq … x (\fst hd)) 443 #H [ >(\b H) | >(\bf H) ] normalize /2 by/ * #K cases (K I) 444 qed. 445 446 lemma get_element_append_l1 : 447 ∀A,B. ∀l1,l2: associative_list A B. ∀x. 448 ¬ (bool_to_Prop (x ∈ domain_of_associative_list … l2)) → 449 get_element ?? (l1@l2) x = get_element … l1 x. 450 #A #B #l1 elim l1 normalize [2: #x #xs #IH #l2 #a #H >IH // ] 451 #l2 elim l2 // #y #ys #IH #a normalize cases(a == \fst y) normalize 452 [ * #H @⊥ @H % ] #H @IH assumption 453 qed. 454 455 lemma get_element_append_r1 : 456 ∀A,B. ∀l1,l2: associative_list A B. ∀x. 457 ¬ (bool_to_Prop (x ∈ domain_of_associative_list … l1)) → 458 get_element ?? (l1@l2) x = get_element … l2 x. 459 #A #B #l1 elim l1 normalize // #x #xs #IH #l2 #a cases (?==?) 460 normalize [* #H cases H //] #H >IH normalize // 461 qed. 462 463 definition fresh_nf_label : ℕ → NonFunctionalLabel × ℕ ≝ 464 λx.〈a_non_functional_label x,S x〉. 465 466 definition fresh_cc_labe : ℕ → CallCostLabel × ℕ ≝ 467 λx.〈a_call_label x,S x〉. 468 469 definition fresh_rc_label : ℕ → ReturnPostCostLabel × ℕ ≝ 470 λx.〈a_return_cost_label (S x),S x〉. 471 472 record call_post_info (p : instr_params) : Type[0] ≝ 473 { gen_labels : list CostLabel 474 ; t_code : Instructions p 475 ; fresh : ℕ 476 ; lab_map : associative_list DEQCostLabel CostLabel 477 ; lab_to_keep : list ReturnPostCostLabel 360 record call_post_info (p : instr_params) (l_p : label_params) : Type[0] ≝ 361 { gen_labels : list (CostLabel l_p) 362 ; t_code : Instructions p l_p 363 ; fresh : l_p 364 ; lab_map : associative_list (DEQCostLabel l_p) (CostLabel l_p) 365 ; lab_to_keep : list (ReturnPostCostLabel l_p) 478 366 }. 479 367 480 let rec call_post_trans (p : instr_params) ( i : Instructions p) (n : ℕ) on i :481 list CostLabel → call_post_infop ≝368 let rec call_post_trans (p : instr_params) (l_p : label_params) (i : Instructions p l_p) (n : l_p) on i : 369 list (CostLabel l_p) → call_post_info p l_p ≝ 482 370 λabs. 483 371 match i with 484 [ EMPTY ⇒ mk_call_post_info ?abs (EMPTY …) n (nil ?) (nil ?)485 | RETURN x ⇒ mk_call_post_info ?abs (RETURN … x) n (nil ?) (nil ?)372 [ EMPTY ⇒ mk_call_post_info … abs (EMPTY …) n (nil ?) (nil ?) 373 | RETURN x ⇒ mk_call_post_info … abs (RETURN … x) n (nil ?) (nil ?) 486 374 | SEQ x lab instr ⇒ 487 375 let ih ≝ call_post_trans … instr n abs in 488 376 match lab with 489 [ None ⇒ mk_call_post_info ? (gen_labels …ih) (SEQ … x (None ?) (t_code … ih))377 [ None ⇒ mk_call_post_info … (gen_labels ?? ih) (SEQ … x (None ?) (t_code … ih)) 490 378 (fresh … ih) (lab_map … ih) (lab_to_keep … ih) 491 379 | Some lbl ⇒ 492 mk_call_post_info ?(nil ?) (SEQ … x (Some ? lbl) (t_code … ih)) (fresh … ih)493 (〈a_non_functional_label lbl,(a_non_functional_label lbl:: (gen_labels … ih))〉 :: (lab_map … ih))380 mk_call_post_info … (nil ?) (SEQ … x (Some ? lbl) (t_code … ih)) (fresh … ih) 381 (〈a_non_functional_label … lbl,((a_non_functional_label … lbl) :: (gen_labels … ih))〉 :: (lab_map … ih)) 494 382 (lab_to_keep … ih) 495 383 ] … … 498 386 let ih2 ≝ call_post_trans … i2 (fresh … ih3) (gen_labels … ih3) in 499 387 let ih1 ≝ call_post_trans … i1 (fresh … ih2) (gen_labels … ih3) in 500 mk_call_post_info ?(nil ?) (COND … x ltrue (t_code … ih1) lfalse (t_code … ih2) (t_code … ih3))388 mk_call_post_info p l_p (nil ?) (COND … x ltrue (t_code … ih1) lfalse (t_code … ih2) (t_code … ih3)) 501 389 (fresh … ih1) 502 (〈a_non_functional_label ltrue,(a_non_functional_labelltrue :: (gen_labels … ih1))〉::503 〈a_non_functional_label lfalse,(a_non_functional_labellfalse :: (gen_labels … ih2))〉::390 (〈a_non_functional_label … ltrue,(a_non_functional_label … ltrue :: (gen_labels … ih1))〉:: 391 〈a_non_functional_label … lfalse,(a_non_functional_label … lfalse :: (gen_labels … ih2))〉:: 504 392 ((lab_map … ih1) @ (lab_map … ih2) @ (lab_map … ih3))) 505 393 ((lab_to_keep … ih1) @ (lab_to_keep … ih2) @ (lab_to_keep … ih3)) … … 507 395 let ih2 ≝ call_post_trans … i2 n abs in 508 396 let ih1 ≝ call_post_trans … i1 (fresh … ih2) (nil ?) in 509 mk_call_post_info ?(nil ?) (LOOP … x ltrue (t_code … ih1) lfalse (t_code … ih2)) (fresh … ih1)510 (〈a_non_functional_label lfalse,(a_non_functional_labellfalse :: (gen_labels … ih2))〉 ::511 〈a_non_functional_label ltrue,(a_non_functional_labelltrue :: (gen_labels … ih1))〉 ::397 mk_call_post_info p l_p (nil ?) (LOOP … x ltrue (t_code … ih1) lfalse (t_code … ih2)) (fresh … ih1) 398 (〈a_non_functional_label … lfalse,(a_non_functional_label … lfalse :: (gen_labels … ih2))〉 :: 399 〈a_non_functional_label … ltrue,(a_non_functional_label … ltrue :: (gen_labels … ih1))〉 :: 512 400 ((lab_map … ih1) @ (lab_map … ih2))) 513 401 ((lab_to_keep … ih1) @ (lab_to_keep … ih2)) … … 515 403 let ih ≝ call_post_trans … i1 n abs in 516 404 match r_lb with 517 [ None ⇒ let 〈l',f''〉 ≝ fresh_rc_label (fresh … ih) in518 mk_call_post_info ? ((a_return_postl')::(gen_labels … ih))405 [ None ⇒ let 〈l',f''〉 ≝ fresh_rc_label l_p (fresh … ih) in 406 mk_call_post_info p l_p ((a_return_post … l')::(gen_labels … ih)) 519 407 (CALL … f act_p (Some ? l') (t_code … ih)) f'' (lab_map … ih) (lab_to_keep … ih) 520 408 | Some lbl ⇒ 521 mk_call_post_info ? (nil ?) (CALL ?f act_p (Some ? lbl) (t_code … ih)) (fresh … ih)522 (〈a_return_post lbl,(a_return_postlbl :: (gen_labels … ih))〉 :: (lab_map … ih))409 mk_call_post_info p l_p (nil ?) (CALL … f act_p (Some ? lbl) (t_code … ih)) (fresh … ih) 410 (〈a_return_post … lbl,(a_return_post … lbl :: (gen_labels … ih))〉 :: (lab_map … ih)) 523 411 (lbl :: lab_to_keep … ih) 524 412 ] 525 413 | IO lin io lout i1 ⇒ 526 414 let ih ≝ call_post_trans … i1 n abs in 527 mk_call_post_info ? (nil ?) (IO ?lin io lout (t_code … ih)) (fresh … ih)528 (〈a_non_functional_label lout,(a_non_functional_labellout :: (gen_labels … ih))〉 ::529 〈a_non_functional_label lin,[a_non_functional_labellin]〉 :: (lab_map … ih)) (lab_to_keep … ih)415 mk_call_post_info p l_p (nil ?) (IO … lin io lout (t_code … ih)) (fresh … ih) 416 (〈a_non_functional_label … lout,(a_non_functional_label … lout :: (gen_labels … ih))〉 :: 417 〈a_non_functional_label … lin,[a_non_functional_label … lin]〉 :: (lab_map … ih)) (lab_to_keep … ih) 530 418 ]. 531 419 532 420 533 let rec call_post_clean (p : instr_params) ( i : Instructionsp) on i :534 associative_list DEQCostLabel CostLabel → list ReturnPostCostLabel → list CostLabel→535 option ((list CostLabel) × (Instructionsp)) ≝421 let rec call_post_clean (p : instr_params) (l_p : label_params) (i : Instructions p l_p) on i : 422 associative_list (DEQCostLabel l_p) (CostLabel l_p) → list (ReturnPostCostLabel l_p) → list (CostLabel l_p) → 423 option ((list (CostLabel l_p)) × (Instructions p l_p)) ≝ 536 424 λm,keep,abs. 537 425 match i with … … 569 457 then return 〈nil ?,CALL … f act_p (Some ? lbl) instr1〉 570 458 else None ? 571 else return 〈(a_return_post l bl) :: l1,CALL … f act_p (None ?) instr1〉459 else return 〈(a_return_post l_p lbl) :: l1,CALL … f act_p (None ?) instr1〉 572 460 ] 573 461 | IO lin io lout i1 ⇒ … … 578 466 ]. 579 467 580 let rec foldr2 (A : Type[0]) (B : Type[0]) (C : Type[0]) (a : A) (l1 : list B) 581 (l2 : list C) (f : A → B → C → A) on l1 : option A≝ 582 match l1 with 583 [ nil ⇒ match l2 with [ nil ⇒ return a | cons y ys ⇒ None ? ] 584 | cons x xs ⇒ 585 match l2 with 586 [ nil ⇒ None ? 587 | cons y ys ⇒ ! ih ← (foldr2 … a xs ys f); 588 return f ih x y 589 ] 590 ]. 591 592 definition is_silent_cost_act_b : ActionLabel → bool ≝ 593 λact. match act with 594 [ cost_act x ⇒ match x with [None ⇒ true | _ ⇒ false ] | _ ⇒ false]. 595 596 definition eq_ActionLabel : ActionLabel → ActionLabel → bool ≝ 597 λc1,c2. 598 match c1 with 599 [ call_act f l ⇒ match c2 with [ call_act f' l' ⇒ f == f' ∧ l == l' | _ ⇒ false] 600 | ret_act x ⇒ match c2 with [ret_act y ⇒ match x with [ None ⇒ match y with [ None ⇒ true | _ ⇒ false] 601 | Some l ⇒ match y with [ Some l' ⇒ l == l' | _ ⇒ false] 602 ] 603 | _ ⇒ false 604 ] 605 | cost_act x ⇒ match c2 with [cost_act y ⇒ match x with [ None ⇒ match y with [ None ⇒ true | _ ⇒ false] 606 | Some l ⇒ match y with [ Some l' ⇒ l == l' | _ ⇒ false] 607 ] 608 | _ ⇒ false 609 ] 610 ]. 611 612 definition ret_costed_abs : list ReturnPostCostLabel → option ReturnPostCostLabel → option CostLabel ≝ 613 λkeep,x. 468 469 definition ret_costed_abs : ∀p.list (ReturnPostCostLabel p) → option (ReturnPostCostLabel p) → 470 option (CostLabel p) ≝ 471 λp,keep,x. 614 472 match x with 615 [ Some lbl ⇒ if lbl ∈ keep then return (a_return_post lbl)473 [ Some lbl ⇒ if lbl ∈ keep then return (a_return_post … lbl) 616 474 else None ? 617 475 | None ⇒ None ? … … 619 477 620 478 621 definition check_continuations : ∀p : instr_params. 622 ∀l1,l2 : list ( ActionLabel × (Instructionsp)).623 associative_list DEQCostLabel CostLabel→624 list ReturnPostCostLabel → option (Prop × (list CostLabel) × (list CostLabel)) ≝625 λp, cont1,cont2,m,keep.479 definition check_continuations : ∀p : instr_params.∀l_p : label_params. 480 ∀l1,l2 : list ((ActionLabel l_p) × (Instructions p l_p)). 481 associative_list (DEQCostLabel l_p) (CostLabel l_p) → 482 list (ReturnPostCostLabel l_p) → option (Prop × (list (CostLabel l_p)) × (list (CostLabel l_p))) ≝ 483 λp,l_p,cont1,cont2,m,keep. 626 484 foldr2 ??? 〈True,nil ?,nil ?〉 cont1 cont2 627 485 (λx,y,z. 628 486 let 〈cond,abs_top',abs_tail'〉 ≝ x in 629 match call_post_clean p (\snd z) m keep abs_top' with487 match call_post_clean p l_p (\snd z) m keep abs_top' with 630 488 [ None ⇒ 〈False,nil ?,nil ?〉 631 489 | Some w ⇒ 632 490 match \fst z with 633 491 [ ret_act opt_x ⇒ 634 match ret_costed_abs keep opt_x with492 match ret_costed_abs … keep opt_x with 635 493 [ Some lbl ⇒ 〈\fst y = \fst z ∧ cond ∧ \snd w = \snd y ∧ 636 494 get_element … m lbl = lbl :: (\fst w),(nil ?),abs_tail'〉 637 495 | None ⇒ 638 〈\fst y = ret_act (None ?) ∧ cond ∧ \snd w = \snd y,(nil ?),(\fst w) @ abs_tail'〉496 〈\fst y = ret_act … (None ?) ∧ cond ∧ \snd w = \snd y,(nil ?),(\fst w) @ abs_tail'〉 639 497 ] 640 498 | cost_act opt_x ⇒ … … 662 520 663 521 664 definition state_rel : ∀p. 665 associative_list DEQCostLabel CostLabel → list ReturnPostCostLabel → list CostLabel →(list CostLabel) → 522 definition state_rel : ∀p : state_params. 523 associative_list (DEQCostLabel p) (CostLabel p) → list (ReturnPostCostLabel p) → 524 list (CostLabel p) → list (CostLabel p) → 666 525 relation (state p) ≝ λp,m,keep,abs_top,abs_tail,st1,st2. 667 match check_continuations ?(cont ? st1) (cont … st2) m keep with526 match check_continuations … (cont ? st1) (cont … st2) m keep with 668 527 [ Some x ⇒ let 〈prf1,abs_top',abs_tail'〉 ≝ x in 669 528 prf1 ∧ call_post_clean … (code … st2) m keep abs_top' = return 〈abs_top,(code … st1)〉 … … 672 531 ]. 673 532 674 include "Simulation.ma". 675 676 let rec len (s : abstract_status) (st1 : s) (st2 : s) (t : raw_trace s st1 st2) 677 on t : ℕ ≝ 678 match t with 679 [ t_base s ⇒ O 680 | t_ind s1 s2 s3 l prf lt ⇒ S (len … lt) 681 ]. 682 (* 683 lemma associative_max : ∀n1,n2,n3.max (max n1 n2) n3 = max n1 (max n2 n3). 684 #n1 #n2 #n3 normalize @leb_elim normalize 685 [ @leb_elim normalize 686 [ #H1 #H2 @leb_elim normalize 687 [ @leb_elim normalize // * #H @⊥ @H assumption 688 | @leb_elim normalize 689 [ #H3 * #H @⊥ @H @(transitive_le … H1) assumption 690 | * #H @⊥ @H assumption 691 ] 692 ] 693 | #H1 #H2 @leb_elim normalize 694 [ @leb_elim normalize // #_ #H cases H1 #H1 @⊥ @H1 assumption 695 | @leb_elim normalize 696 [ #_ * #H @⊥ @H assumption 697 | * #H @⊥ @H @(transitive_le … H2) 698 *) 699 let rec compute_max_n (p : instr_params) (i : Instructions p) on i : ℕ ≝ 533 let rec compute_max_n (p : instr_params) (l_p : label_params) (i : Instructions p l_p) on i : l_p ≝ 700 534 match i with 701 [ EMPTY ⇒ O702 | RETURN x ⇒ O535 [ EMPTY ⇒ e … l_p 536 | RETURN x ⇒ e … l_p 703 537 | SEQ x lab instr ⇒ let n ≝ compute_max_n … instr in 704 538 match lab with … … 706 540 | Some l ⇒ 707 541 match l with 708 [ a_non_functional_label n' ⇒ S (max n n')]542 [ a_non_functional_label n' ⇒ op … l_p n n' ] 709 543 ] 710 544 | COND x ltrue i1 lfalse i2 i3 ⇒ … … 712 546 let n2 ≝ compute_max_n … i2 in 713 547 let n3 ≝ compute_max_n … i3 in 714 let mx ≝ max (maxn1 n2) n3 in548 let mx ≝ op … l_p (op … l_p n1 n2) n3 in 715 549 match ltrue with 716 550 [ a_non_functional_label lt ⇒ 717 551 match lfalse with 718 [a_non_functional_label lf ⇒ S (max (max mx lt) lf)] ]552 [a_non_functional_label lf ⇒ op … l_p (op … l_p mx lt) lf ] ] 719 553 | LOOP x ltrue i1 lfalse i2 ⇒ 720 554 let n1 ≝ compute_max_n … i1 in 721 555 let n2 ≝ compute_max_n … i2 in 722 let mx ≝ maxn1 n2 in556 let mx ≝ op … l_p n1 n2 in 723 557 match ltrue with 724 558 [ a_non_functional_label lt ⇒ 725 559 match lfalse with 726 [a_non_functional_label lf ⇒ S (max (max mx lt) lf)] ]560 [a_non_functional_label lf ⇒ op … l_p (op … l_p mx lt) lf ] ] 727 561 | CALL f act_p r_lb i1 ⇒ 728 562 let n ≝ compute_max_n … i1 in 729 563 match r_lb with 730 564 [ None ⇒ n 731 | Some lbl ⇒ match lbl with [a_return_cost_label l ⇒ S(max l n)]565 | Some lbl ⇒ match lbl with [a_return_cost_label l ⇒ op … l_p l n ] 732 566 ] 733 567 | IO lin io lout i1 ⇒ … … 736 570 [a_non_functional_label l1 ⇒ 737 571 match lout with 738 [a_non_functional_label l2 ⇒ S(max (max n l1) l2)] ]572 [a_non_functional_label l2 ⇒ op … l_p (op … l_p n l1) l2 ] ] 739 573 ]. 740 574 741 575 742 definition same_fresh_map_on : list CostLabel → 743 relation (associative_list DEQCostLabel CostLabel) ≝ 744 λdom,m1,m2.∀x.bool_to_Prop (x ∈ dom) → get_element … m1 x = get_element … m2 x. 745 746 definition same_to_keep_on : list CostLabel → relation (list ReturnPostCostLabel) ≝ 747 λdom,keep1,keep2.∀x. bool_to_Prop (a_return_post x ∈ dom) → (x ∈ keep1) = (x ∈ keep2). 748 749 lemma memb_not_append : ∀D : DeqSet.∀l1,l2 : list D.∀x : D. 750 x ∈ l1 = false → x ∈ l2 = false → x ∈ (l1 @ l2) = false. 751 #D #l1 elim l1 752 [ #l2 #x #_ #H @H ] 753 #x #xs #IH #l2 #x1 whd in match (memb ???); inversion (x1 == x) normalize nodelta 754 [ #_ #EQ destruct] #EQx1 #EQxs #EQl2 whd in match (memb ???); >EQx1 755 normalize nodelta @IH // 756 qed. 757 758 lemma memb_no_duplicates_append : ∀A : DeqSet.∀x.∀l1,l2 : list A . 759 no_duplicates … (l1 @ l2) → x ∈ l1 → x ∈ l2 → False. 760 #A #x #l1 elim l1 // #x1 #xs #IH #l2 * #H1 #H2 whd in match (memb ???); 761 inversion (x == x1) normalize nodelta 762 [ #H3 #_ #H4 >memb_append_l2 in H1; [2: <(\P H3) @H4 ] * #H @H % 763 | #_ @IH // 764 ] 765 qed. 766 767 768 lemma same_to_keep_on_append : ∀dom1,dom2,dom3 : list CostLabel. 769 ∀l1,l2,l3,l : list ReturnPostCostLabel. 770 no_duplicates … (dom1@dom2@dom3) → (∀x.x ∈ l1 → a_return_post x ∈ dom1) → 771 (∀x.x ∈ l3 → a_return_post x ∈ dom3) → 772 same_to_keep_on (dom1@dom2@dom3) (l1@l2@l3) l → 773 same_to_keep_on dom2 l2 l. 774 #dom1 #dom2 #dom3 #l1 #l2 #l3 #l #no_dup #sub_set1 #sub_set3 #H2 576 definition same_fresh_map_on : ∀p.list (CostLabel p) → 577 relation (associative_list (DEQCostLabel p) (CostLabel p)) ≝ 578 λp,dom,m1,m2.∀x.bool_to_Prop (x ∈ dom) → get_element … m1 x = get_element … m2 x. 579 580 definition same_to_keep_on : ∀p. list (CostLabel p) → relation (list (ReturnPostCostLabel p)) ≝ 581 λp,dom,keep1,keep2.∀x. bool_to_Prop (a_return_post … x ∈ dom) → (x ∈ keep1) = (x ∈ keep2). 582 583 584 lemma same_to_keep_on_append : ∀p.∀dom1,dom2,dom3 : list (CostLabel p). 585 ∀l1,l2,l3,l : list (ReturnPostCostLabel p). 586 no_duplicates … (dom1@dom2@dom3) → (∀x.x ∈ l1 → a_return_post … x ∈ dom1) → 587 (∀x.x ∈ l3 → a_return_post … x ∈ dom3) → 588 same_to_keep_on … (dom1@dom2@dom3) (l1@l2@l3) l → 589 same_to_keep_on … dom2 l2 l. 590 #p #dom1 #dom2 #dom3 #l1 #l2 #l3 #l #no_dup #sub_set1 #sub_set3 #H2 775 591 #x #Hx inversion (x ∈ l2) 776 592 [ #EQkeep <H2 [> memb_append_l2 // >memb_append_l1 // ] … … 780 596 | @sym_eq @memb_not_append [2: @memb_not_append //] 781 597 [ <associative_append in no_dup; #no_dup ] 782 lapply(memb_no_duplicates_append … (a_return_post x) … no_dup) #H598 lapply(memb_no_duplicates_append … (a_return_post … x) … no_dup) #H 783 599 inversion(memb ???) // #H1 cases H 784 600 [1,4: [>memb_append_l2 | >memb_append_l1] // >Hx // … … 790 606 qed. 791 607 792 lemma same_fresh_map_on_append : ∀ dom1,dom2,dom3,l1,l2,l3,l.608 lemma same_fresh_map_on_append : ∀p.∀dom1,dom2,dom3,l1,l2,l3,l. 793 609 no_duplicates … (dom1 @dom2 @ dom3) → (∀x.x ∈ domain_of_associative_list … l1 → x ∈ dom1) → 794 610 (∀x.x ∈ domain_of_associative_list … l3 → x ∈ dom3) → 795 same_fresh_map_on … (dom1 @dom2 @dom3) (l1 @l2 @ l3) l →796 same_fresh_map_on … dom2 l2 l.797 # dom1 #dom2 #dom3 #l1 #l2 #l3 #l #no_dup #subset1 #subset3 whd in ⊢ (% → ?); #H1611 same_fresh_map_on p … (dom1 @dom2 @dom3) (l1 @l2 @ l3) l → 612 same_fresh_map_on p … dom2 l2 l. 613 #p #dom1 #dom2 #dom3 #l1 #l2 #l3 #l #no_dup #subset1 #subset3 whd in ⊢ (% → ?); #H1 798 614 whd #x #Hx <H1 799 615 [2: >memb_append_l2 // >memb_append_l1 // >Hx //] … … 805 621 806 622 807 lemma lab_to_keep_in_domain : ∀p .∀i : Instructionsp.623 lemma lab_to_keep_in_domain : ∀p,l_p.∀i : Instructions p l_p. 808 624 ∀x,n,l. 809 x ∈ lab_to_keep … (call_post_trans … i n l) → a_return_post x ∈ get_labels_of_code … i.810 #p # i elim i //625 x ∈ lab_to_keep … (call_post_trans … i n l) → a_return_post … x ∈ get_labels_of_code … i. 626 #p #l_p #i elim i // 811 627 [ #seq #opt_l #instr #IH #x #n #l whd in match (call_post_trans ????); 812 628 cases opt_l -opt_l normalize nodelta [|#lbl] … … 836 652 qed. 837 653 838 lemma domain_of_associative_list_append : ∀A,B.∀l1,l2 : associative_list A B. 839 domain_of_associative_list ?? (l1 @ l2) = 840 (domain_of_associative_list ?? l1) @ (domain_of_associative_list ?? l2). 841 #A #B #l1 elim l1 // #x #xs #IH #l2 normalize // 842 qed. 843 844 lemma lab_map_in_domain: ∀p.∀i: Instructions p. 654 lemma lab_map_in_domain: ∀p,l_p.∀i: Instructions p l_p. 845 655 ∀x,n,l. 846 x ∈ domain_of_associative_list … (lab_map p(call_post_trans … i n l)) →656 x ∈ domain_of_associative_list … (lab_map … (call_post_trans … i n l)) → 847 657 x ∈ get_labels_of_code … i. 848 #p # i elim i //658 #p #l_p #i elim i // 849 659 [ #seq * [|#lbl] #i1 #IH #x #n #l whd in match(call_post_trans ????); 850 660 whd in match (get_labels_of_code ??); /2/ whd in match (memb ???); … … 879 689 qed. 880 690 881 lemma eq_a_non_functional_label : 882 ∀c1,c2.c1 == c2 → a_non_functional_label c1 == a_non_functional_label c2. 883 #c1 #c2 #EQ cases(eqb_true DEQNonFunctionalLabel c1 c2) #H1 #_ lapply(H1 (eq_true_to_b … EQ)) -EQ 884 #EQ destruct >(\b (refl …)) @I 885 qed. 886 887 let rec is_fresh_for_return (keep : list CostLabel) (n : ℕ) on keep : Prop ≝ 691 let rec is_fresh_for_return (p : label_params) (keep : list (CostLabel p)) (n : p) on keep : Prop ≝ 888 692 match keep with 889 693 [ nil ⇒ True 890 | cons x xs ⇒ let ih ≝ is_fresh_for_return xs n in694 | cons x xs ⇒ let ih ≝ is_fresh_for_return … xs n in 891 695 match x with 892 [ a_return_post y ⇒ match y with [a_return_cost_label m ⇒ m ≤n ∧ ih ]696 [ a_return_post y ⇒ match y with [a_return_cost_label m ⇒ m ⊑^{p} n ∧ ih ] 893 697 | _ ⇒ ih 894 698 ] 895 699 ]. 896 700 897 lemma fresh_ok_call_post_trans : ∀p : instr_params.∀ i1 : Instructions p.∀n : ℕ.∀l.898 n ≤ fresh … (call_post_trans pi1 n l).899 #p # i1 elim i1 normalize /2 by transitive_le, le_n/701 lemma fresh_ok_call_post_trans : ∀p : instr_params.∀l_p.∀i1 : Instructions p l_p.∀n : l_p.∀l. 702 n ⊑^{l_p} fresh … (call_post_trans … i1 n l). 703 #p #l_p #i1 elim i1 normalize /2 by trans_po_rel, refl_po_rel/ 900 704 [ #seq * [|#lbl] #i2 #IH #n #l normalize /2 by / 901 705 | #cond #ltrue #i_true #lfalse #i_false #i2 #IH1 #IH2 #IH3 #n #l 902 @(trans itive_le … (IH3 …)) [2: @(transitive_le … (IH2 …)) [2: @(transitive_le… (IH1 …)) ]] //903 | #f #act_p * [|#lbl] normalize #i2 #IH /2 by le_S/706 @(trans_po_rel … (IH3 …)) [2: @(trans_po_rel … (IH2 …)) [2: @(trans_po_rel … (IH1 …)) ]] // 707 | #f #act_p * [|#lbl] normalize #i2 #IH /2 by / #n #l @(trans_po_rel … (IH … l …)) @lab_po_rel_succ 904 708 ] 905 709 qed. 906 710 907 lemma fresh_keep_n_ok : ∀ n,m,l.908 is_fresh_for_return l n → n ≤ m → is_fresh_for_returnl m.909 # n #m #l lapply n -n lapply m -m elim l // *711 lemma fresh_keep_n_ok : ∀p : label_params.∀n,m.∀l. 712 is_fresh_for_return p l n → n ⊑^{p} m → is_fresh_for_return p l m. 713 #p #n #m #l lapply n -n lapply m -m elim l // * 910 714 [1,2,3: * #x] #xs #IH #n #m 911 normalize [2: * #H1] #H2 #H3 [ %] /2 by trans itive_le/912 qed. 913 914 definition cast_return_to_cost_labels ≝ map … (a_return_post…).715 normalize [2: * #H1] #H2 #H3 [ %] /2 by trans_po_rel/ @(IH … H2) assumption 716 qed. 717 718 definition cast_return_to_cost_labels ≝ λp.map … (a_return_post p …). 915 719 coercion cast_return_to_cost_labels. 916 720 917 lemma fresh_false : ∀n.∀l: list ReturnPostCostLabel.is_fresh_for_return l n → 918 a_return_cost_label (S n) ∈ l = false. 919 #n #l lapply n -n elim l // * #x #xs #IH #n whd in ⊢ (% → ??%?); * #H1 721 lemma succ_label_leq_absurd : ∀p : label_params.∀x : p.succ_label … p … x ⊑^{p} x → False. 722 #p #x #ABS @(absurd ?? (no_maps_in_id … p x)) @(antisym_po_rel … (po_label … p)) // 723 qed. 724 725 lemma fresh_false : ∀p.∀n.∀l: list (ReturnPostCostLabel p).is_fresh_for_return … l n → 726 a_return_cost_label p (succ_label … n) ∈ l = false. 727 #p #n #l lapply n -n elim l // * #x #xs #IH #n whd in ⊢ (% → ??%?); * #H1 920 728 #H2 @eq_return_cost_lab_elim 921 [ #EQ destruct @⊥ /2 by absurd/729 [ #EQ destruct @⊥ @succ_label_leq_absurd // 922 730 | #_ >IH // 923 731 ] 924 732 qed. 925 733 926 lemma inverse_call_post_trans : ∀p : instr_params.∀ i1 : Instructions p.∀n : ℕ.734 lemma inverse_call_post_trans : ∀p : instr_params.∀l_p : label_params.∀i1 : Instructions p l_p.∀n : l_p. 927 735 let dom ≝ get_labels_of_code … i1 in 928 736 no_duplicates … dom → 929 737 ∀m,keep. 930 ∀info,l.call_post_trans pi1 n l = info →931 same_fresh_map_on dom (lab_map … info) m →932 same_to_keep_on dom (lab_to_keep … info) keep →933 is_fresh_for_return keep n →934 call_post_clean ? (t_code ?info) m keep l738 ∀info,l.call_post_trans … i1 n l = info → 739 same_fresh_map_on … dom (lab_map … info) m → 740 same_to_keep_on … dom (lab_to_keep … info) keep → 741 is_fresh_for_return … keep n → 742 call_post_clean … (t_code … info) m keep l 935 743 = return 〈gen_labels … info,i1〉. 936 #p # i1 elim i1744 #p #l_p #i1 elim i1 937 745 [ #n #no_dup #m #keep * #gen_lab #t_code #fresh #lab_map #lab_keep #l whd in ⊢ (??%? → ?); 938 746 #EQ destruct(EQ) // … … 950 758 ] 951 759 normalize nodelta <(H1 lbl) 952 [2: whd in match (get_labels_of_code ??); @orb_Prop_l cases(eqb_true … (a_non_functional_label lbl) lbl)760 [2: whd in match (get_labels_of_code ??); @orb_Prop_l cases(eqb_true … (a_non_functional_label … lbl) lbl) 953 761 #H3 #H4 >H4 % ] 954 762 whd in match (get_element ????); >(\b (refl …)) normalize nodelta … … 960 768 change with ([?;?]@?) in match (?::?) in H2; 961 769 <associative_append in H2; <associative_append 962 <(append_nil … (?@?)) <associative_append in ⊢ (?? %? → ?);963 <(append_nil … (?@?)) in ⊢ (?? %? → ?); >associative_append964 >associative_append in ⊢ (?? %? → ?); #H2770 <(append_nil … (?@?)) <associative_append in ⊢ (???%? → ?); 771 <(append_nil … (?@?)) in ⊢ (???%? → ?); >associative_append 772 >associative_append in ⊢ (???%? → ?); #H2 965 773 @(same_to_keep_on_append … H2) // [ >append_nil 966 774 whd in ⊢ (??%); whd in no_dup:(??%); >associative_append // ] … … 972 780 <associative_append in H1; <associative_append 973 781 <(append_nil … (?@?)) >associative_append 974 change with ([?;?]@?) in match (?::?::?) in ⊢ (?? %? → ?);975 <associative_append in ⊢ (?? %? → ?);976 <associative_append in ⊢ (?? %? → ?);977 <(append_nil … (?@?)) in ⊢ (?? %? → ?);978 >associative_append in ⊢ (?? %? → ?); #H1782 change with ([?;?]@?) in match (?::?::?) in ⊢ (???%? → ?); 783 <associative_append in ⊢ (???%? → ?); 784 <associative_append in ⊢ (???%? → ?); 785 <(append_nil … (?@?)) in ⊢ (???%? → ?); 786 >associative_append in ⊢ (???%? → ?); #H1 979 787 @(same_fresh_map_on_append … H1) // 980 788 [ >append_nil >associative_append // ] … … 1001 809 |4: whd in match (get_labels_of_code ??) in H1; 1002 810 change with ([?;?]@?) in match (?::?) in H1; 1003 change with ([?;?]@?) in match (?::?::?) in H1 : (?? %?);1004 <associative_append in H1; <associative_append in ⊢ (?? %? → ?); #H1811 change with ([?;?]@?) in match (?::?::?) in H1 : (???%?); 812 <associative_append in H1; <associative_append in ⊢ (???%? → ?); #H1 1005 813 @(same_fresh_map_on_append … H1) // [2: /2 by lab_map_in_domain/ ] 1006 814 #x >domain_of_associative_list_append #H cases(memb_append … H) … … 1019 827 |3: whd in match (get_labels_of_code ??) in H2; 1020 828 change with ([?;?]@?) in match (?::?) in H2; 1021 change with ([ ] @ ?) in match (lab_to_keep ?? ) in H2;1022 >associative_append in H2 : (?? %?); #H2829 change with ([ ] @ ?) in match (lab_to_keep ???) in H2; 830 >associative_append in H2 : (???%?); #H2 1023 831 @(same_to_keep_on_append … H2) // #x #Hx cases(memb_append … Hx) 1024 832 -Hx #Hx [ >memb_append_l1 | >memb_append_l2] // … … 1026 834 |4: whd in match (get_labels_of_code ??) in H1; 1027 835 change with ([?;?]@?) in match (?::?) in H1; 1028 change with ([?;?]@?) in match (?::?::?) in H1 : (?? %?);836 change with ([?;?]@?) in match (?::?::?) in H1 : (???%?); 1029 837 @(same_fresh_map_on_append … H1) // #x >domain_of_associative_list_append 1030 838 #Hx cases(memb_append … Hx) -Hx #Hx [ >memb_append_l1 | >memb_append_l2 ] … … 1041 849 whd in match (get_element ????); @eq_costlabel_elim normalize nodelta 1042 850 [ #ABS @⊥ cases no_dup >ABS * #H #_ @H @orb_Prop_l 1043 >(\b (refl ? (a_non_functional_label ltrue))) % ] #_851 >(\b (refl ? (a_non_functional_label … ltrue))) % ] #_ 1044 852 whd in match (get_element ????); >(\b (refl …)) normalize nodelta 1045 853 >(\b (refl …)) % … … 1050 858 [ >m_return_bind <fresh_map [2: @orb_Prop_l >(\b (refl …)) % ] 1051 859 whd in match (get_element ????); 1052 inversion(a_non_functional_label ltrue == a_non_functional_labellfalse)860 inversion(a_non_functional_label … ltrue == a_non_functional_label … lfalse) 1053 861 #Hltrue normalize nodelta 1054 862 [ cases no_dup whd in match (memb ???); 1055 cases(eqb_true … (a_non_functional_label ltrue) (a_non_functional_labellfalse))863 cases(eqb_true … (a_non_functional_label … ltrue) (a_non_functional_label … lfalse)) 1056 864 #H1 #_ lapply(H1 Hltrue) #EQ destruct(EQ) >(\b (refl …)) * #ABS @⊥ @ABS % ] 1057 865 whd in match (get_element ????); >(\b (refl …)) normalize nodelta >(\b (refl …)) … … 1060 868 >(\b (refl …)) % 1061 869 | /2 by fresh_keep_n_ok/ 1062 | change with ([?;?]@?@?) in keep_on : (? %??); change with ((nil ?) @ ? @ ?) in keep_on : (??%?);870 | change with ([?;?]@?@?) in keep_on : (??%??); change with ((nil ?) @ ? @ ?) in keep_on : (???%?); 1063 871 @(same_to_keep_on_append … keep_on) // #x /2 by lab_to_keep_in_domain/ 1064 | change with ([?;?]@?@?) in fresh_map : (? %%?); @(same_fresh_map_on_append … fresh_map)872 | change with ([?;?]@?@?) in fresh_map : (??%%?); @(same_fresh_map_on_append … fresh_map) 1065 873 /2 by lab_map_in_domain/ #x whd in match (memb ???); inversion(x==lfalse) #Hlfalse 1066 874 normalize nodelta … … 1072 880 ] 1073 881 | // 1074 | change with ([?;?]@?@?) in keep_on : (? %??); <associative_append in keep_on;1075 <(append_nil … (?@?)) <(append_nil … (?@?)) in ⊢ (?? %? → ?);1076 >associative_append in ⊢ (? %%? → ?); >associative_append in ⊢ (??%? → ?);882 | change with ([?;?]@?@?) in keep_on : (??%??); <associative_append in keep_on; 883 <(append_nil … (?@?)) <(append_nil … (?@?)) in ⊢ (???%? → ?); 884 >associative_append in ⊢ (??%%? → ?); >associative_append in ⊢ (???%? → ?); 1077 885 #keep_on @(same_to_keep_on_append … keep_on) // 1078 886 [ >associative_append >append_nil // 1079 887 | #x #Hx @orb_Prop_r @orb_Prop_r /2 by lab_to_keep_in_domain/ 1080 888 ] 1081 | change with ([?;?]@?@?) in fresh_map : (? %??); <associative_append in fresh_map;1082 <(append_nil … (?@?)) change with ([?;?]@?@?) in ⊢ (?? %? → ?);1083 <associative_append in ⊢ (?? %? → ?); <(append_nil … (?@?)) in ⊢ (??%? → ?);1084 >associative_append in ⊢ (? %%? → ?); >associative_append in ⊢ (??%? → ?);889 | change with ([?;?]@?@?) in fresh_map : (??%??); <associative_append in fresh_map; 890 <(append_nil … (?@?)) change with ([?;?]@?@?) in ⊢ (???%? → ?); 891 <associative_append in ⊢ (???%? → ?); <(append_nil … (?@?)) in ⊢ (???%? → ?); 892 >associative_append in ⊢ (??%%? → ?); >associative_append in ⊢ (???%? → ?); 1085 893 #fresh_map @(same_fresh_map_on_append … fresh_map) // 1086 894 [ >append_nil // … … 1113 921 ] 1114 922 |2,7: // 1115 |3,8: whd in match (get_labels_of_code ?? ) in same_keep; // #x #Hx <same_keep923 |3,8: whd in match (get_labels_of_code ???) in same_keep; // #x #Hx <same_keep 1116 924 [2: >memb_cons // >Hx // ] cases no_dup * #ABS #_ whd in ⊢ (???%); 1117 925 inversion(x==?) [2: #_ //] #ABS1 @⊥ @ABS <(\P ABS1) >Hx // 1118 |4,9: whd in match (get_labels_of_code ?? ) in fresh_map; // #x #Hx <fresh_map926 |4,9: whd in match (get_labels_of_code ???) in fresh_map; // #x #Hx <fresh_map 1119 927 [2: >memb_cons // >Hx //] cases no_dup * #ABS #_ whd in ⊢ (???%); 1120 928 inversion(x==?) [2: #_ //] #ABS1 @⊥ @ABS <(\P ABS1) // 1121 929 |5,10: [ @no_dup | cases no_dup // ] 1122 930 ] 1123 | #lin #io #lout #i #IH #n whd in match (get_labels_of_code ?? ); #no_dup931 | #lin #io #lout #i #IH #n whd in match (get_labels_of_code ???); #no_dup 1124 932 #m #keep #info #l whd in ⊢ (??%? → ?); #EQ destruct(EQ) #fresh_map #same_keep 1125 933 #f_k whd in ⊢ (??%?); >(IH … (refl …)) … … 1130 938 [ #ABS @⊥ cases no_dup * #ABS1 #_ @ABS1 whd in match (memb ???); >(\P ABS) 1131 939 >(\b (refl …)) // 1132 | #H inversion (a_non_functional_label lin== ? lout)940 | #H inversion (a_non_functional_label … lin== ? lout) 1133 941 [ #ABS lapply(\P ABS) #EQ destruct >(\b (refl …)) in H; #EQ destruct 1134 942 | #_ normalize nodelta whd in match (get_element ????); >(\b (refl …)) … … 1150 958 qed. 1151 959 1152 definition fresh_for_prog_aux : ∀p,p'.Program p p' → ℕ → ℕ ≝ 1153 λp,p',prog,n.foldl … (λn,i.max n (compute_max_n … (f_body … i))) n (env … prog). 1154 1155 include alias "arithmetics/nat.ma". 1156 1157 lemma max_1 : ∀n,m,k.k ≤ m → k ≤ max m n. 1158 #m #n #k #H normalize @leb_elim // normalize nodelta #H1 1159 /2 by transitive_le/ 1160 qed. 1161 1162 lemma max_2 : ∀n,m,k.k≤ n → k ≤ max m n. 1163 #m #n #k #H normalize @leb_elim // #H1 normalize nodelta 1164 @(transitive_le … H) @(transitive_le … (not_le_to_lt … H1)) // 1165 qed. 1166 1167 lemma fresh_aux_ok : ∀p,p'.∀prog : Program p p'.∀n,m.n ≤ m → 1168 fresh_for_prog_aux … prog n ≤ fresh_for_prog_aux … prog m. 1169 #p #p' * #env #main elim env // #hd #tl #IH #n #m #H whd in ⊢ (?%%); 1170 @IH whd in ⊢ (?%?); @leb_elim normalize nodelta 1171 [ #H1 @max_2 // | #_ @max_1 // ] 1172 qed. 1173 1174 definition fresh_for_prog : ∀p,p'.Program p p' → ℕ ≝ 1175 λp,p',prog.fresh_for_prog_aux … prog 960 definition fresh_for_prog_aux : ∀p,p',l_p.Program p p' l_p → l_p → l_p ≝ 961 λp,p',l_p,prog,n.foldl … (λn,i.op … l_p … n (compute_max_n … (f_body … i))) n (env … prog). 962 963 964 lemma fresh_aux_ok : ∀p,p',l_p.∀prog : Program p p' l_p.∀n,m.n ⊑^{l_p} m → 965 fresh_for_prog_aux … prog n ⊑^{l_p} fresh_for_prog_aux … prog m. 966 #p #p' #l_p * #env #main elim env // #hd #tl #IH #n #m #H whd in ⊢ (???%%); 967 @IH whd in ⊢ (???%?); @(monotonic_magg … l_p … H) 968 qed. 969 970 definition fresh_for_prog : ∀p,p',l_p.Program p p' l_p → l_p ≝ 971 λp,p',l_p,prog.fresh_for_prog_aux … prog 1176 972 (compute_max_n … (main … prog)). 1177 973 1178 974 definition translate_env ≝ 1179 λp,p' .λenv : list (env_item p p').λmax_all.(foldr ??975 λp,p',l_p.λenv : list (env_item p p' l_p).λmax_all.(foldr ?? 1180 976 (λi,x.let 〈t_env,n,m,keep〉 ≝ x in 1181 977 let info ≝ call_post_trans … (f_body … i) n (nil ?) in 1182 〈(mk_env_item ?? 978 〈(mk_env_item ??? 1183 979 (mk_signature ??(f_name ?? i) (f_pars … i) (f_ret … i)) 1184 980 (f_lab … i) (t_code … info)) :: t_env, 1185 fresh … info, 〈a_call (f_lab … i),(a_call (f_lab … i)) :: (gen_labels? info)〉 ::981 fresh … info, 〈a_call … (f_lab … i),(a_call … (f_lab … i)) :: (gen_labels ?? info)〉 :: 1186 982 ((lab_map … info) @ m),(lab_to_keep … info) @ keep〉) 1187 983 (〈nil ?,max_all,nil ?,nil ?〉) env). 1188 984 1189 definition trans_prog : ∀p,p' .Program p p'→1190 ((Program p p' ) × (associative_list DEQCostLabel CostLabel)) × ((list ReturnPostCostLabel))≝1191 λp,p', prog.985 definition trans_prog : ∀p,p',l_p.Program p p' l_p → 986 ((Program p p' l_p) × (associative_list (DEQCostLabel l_p) (CostLabel l_p)) × ((list (ReturnPostCostLabel l_p))))≝ 987 λp,p',l_p,prog. 1192 988 let max_all ≝ fresh_for_prog … prog in 1193 989 let info_main ≝ (call_post_trans … (main … prog) max_all (nil ?)) in 1194 990 let 〈t_env,n,m,keep〉 ≝ translate_env … (env … prog) (fresh … info_main) in 1195 〈mk_Program ?? t_env (t_code … info_main),m @ (lab_map … info_main),keep @ (lab_to_keep … info_main)〉.1196 1197 definition map_labels_on_trace : 1198 (associative_list DEQCostLabel CostLabel) → list CostLabel → list CostLabel≝1199 λ m,l.foldr … (λlab,t.(get_element … m lab) @ t) (nil ?) l.991 〈mk_Program ??? t_env (t_code … info_main),m @ (lab_map … info_main),keep @ (lab_to_keep … info_main)〉. 992 993 definition map_labels_on_trace : ∀p : label_params. 994 (associative_list (DEQCostLabel p) (CostLabel p)) → list (CostLabel p) → list (CostLabel p) ≝ 995 λp,m,l.foldr … (λlab,t.(get_element … m lab) @ t) (nil ?) l. 1200 996 1201 997 lemma map_labels_on_trace_append: 1202 ∀ m,l1,l2. map_labels_on_tracem (l1@l2) =1203 map_labels_on_trace m l1 @ map_labels_on_tracem l2.1204 #m #l1 elim l1 // #hd #tl #IH #l2 >associative_append <IH //998 ∀p,m,l1,l2. map_labels_on_trace p m (l1@l2) = 999 map_labels_on_trace p m l1 @ map_labels_on_trace p m l2. 1000 #p #m #l1 elim l1 // #hd #tl #IH #l2 >associative_append <IH // 1205 1001 qed. 1206 1002 … … 1241 1037 *) 1242 1038 1243 lemma memb_append_l22 : ∀A : DeqSet.∀x : A.∀l1,l2 : list A. 1244 ¬ (x ∈ l1) → x∈ l1 @ l2 = (x ∈ l2). 1245 #A #x #l1 elim l1 normalize // #y #ys #IH #l2 cases(x==y) 1246 normalize [*] @IH 1247 qed. 1248 1249 lemma memb_append_l12 : ∀A : DeqSet.∀x : A.∀l1,l2 : list A. 1250 ¬ (x ∈ l2) → x∈ l1 @ l2 = (x ∈ l1). 1251 #A #x #l1 elim l1 1252 [ #l2 #H whd in match (append ???); @not_b_to_eq_false @Prop_notb >H % ] 1253 #y #ys #IH #l2 #H1 whd in match (memb ???); >IH // 1254 qed. 1255 1256 1257 lemma lookup_ok_append : ∀p,p',l,f,env_it. 1258 lookup p p' l f = return env_it → ∃l1,l2. l = l1 @ [env_it] @ l2 ∧ 1039 1040 1041 lemma lookup_ok_append : ∀p,p',l_p,l,f,env_it. 1042 lookup p p' l_p l f = return env_it → ∃l1,l2. l = l1 @ [env_it] @ l2 ∧ 1259 1043 f_name … env_it = f. 1260 #p #p' #l elim l [ #f #env_it normalize #EQ destruct]1044 #p #p' #l_p #l elim l [ #f #env_it normalize #EQ destruct] 1261 1045 #x #xs #IH #f #env_it whd in ⊢ (??%? → ?); @eq_function_name_elim 1262 1046 [ #EQ destruct(EQ) normalize nodelta whd in ⊢ (???% → ?); #EQ destruct … … 1274 1058 *) 1275 1059 1276 lemma foldr_map_append : 1277 ∀A,B:Type[0]. ∀l1, l2 : list A. 1278 ∀f:A → list B. ∀seed. 1279 foldr ?? (λx,acc. (f x) @ acc) seed (l1 @ l2) = 1280 append ? (foldr ?? (λx,acc. (f x) @ acc) (nil ?) l1) 1281 (foldr ?? (λx,acc. (f x) @ acc) seed l2). 1282 #A #B #l1 elim l1 normalize // /3 by eq_f, trans_eq/ 1283 qed. 1284 1285 lemma cons_append : ∀A.∀x : A.∀l.x::l = ([x]@l). 1286 // 1287 qed. 1288 1289 lemma eq_a_call_label : 1290 ∀c1,c2.c1 == c2 → a_call c1 == a_call c2. 1291 #c1 #c2 #EQ cases(eqb_true ? c1 c2) #H1 #_ lapply(H1 (eq_true_to_b … EQ)) -EQ 1292 #EQ destruct >(\b (refl …)) @I 1293 qed. 1294 1295 1296 lemma no_duplicates_append_commute : ∀ A : DeqSet.∀l1,l2 : list A. 1297 no_duplicates … (l1 @ l2) → 1298 no_duplicates … (l2 @ l1). 1299 #A #l1 elim l1 1300 [ #l2 >append_nil //] 1301 #x #xs #IH #l2 * #H1 #H2 lapply(IH … H2) lapply H1 -H1 -IH -H2 1302 elim l2 -l1 1303 [ >append_nil #H1 #H2 % // ] 1304 #y #ys #IH * #H1 * #H2 #H3 % 1305 [2: @IH 1306 [ % #H4 @H1 cases(memb_append … H4) 1307 [ #H5 >memb_append_l1 // 1308 | #H5 >memb_append_l2 // @orb_Prop_r >H5 // 1309 ] 1310 | // 1311 ] 1312 | % #H4 cases(memb_append … H4) 1313 [ #H5 @(absurd ?? H2) >memb_append_l1 // 1314 | whd in match (memb ???); inversion(y==x) 1315 [ #H5 #_ <(\P H5) in H1; #H1 @H1 >memb_append_l2 // 1316 | #_ normalize nodelta #H5 @(absurd ?? H2) >memb_append_l2 // 1317 ] 1318 ] 1319 ] 1320 qed. 1060 1321 1061 1322 1062 (* aggiungere fresh_to_keep al lemma seguente??*) 1323 1063 1324 let rec subset (A : Type[0]) (l1,l2 : list A) on l1 : Prop ≝ 1325 match l1 with 1326 [ nil ⇒ True 1327 | cons x xs ⇒ mem … x l2 ∧ subset … xs l2 1328 ]. 1329 1330 interpretation "subset" 'subseteq a b = (subset ? a b). 1331 1332 lemma subset_append_l2 : ∀A,l2,l1.subset A l2 (l1 @ l2). 1333 #A #l2 elim l2 // normalize #x #xs #IH #l1 % // @mem_append_l2 whd /2/ 1334 qed. 1335 1336 lemma refl_subset : ∀A.reflexive … (subset A). 1337 #A #l1 elim l1 // #x #xs #IH normalize % /2/ 1338 qed. 1339 1340 lemma fresh_for_subset : ∀l1,l2,n.l1 ⊆ l2 → is_fresh_for_return … l2 n → 1341 is_fresh_for_return … l1 n. 1342 #l1 elim l1 // * [1,2,3: * #x] #xs #IH #l2 #n * #H1 #H2 #H3 whd 1064 lemma fresh_for_subset : ∀p : label_params.∀l1,l2,n.l1 ⊆ l2 → is_fresh_for_return p … l2 n → 1065 is_fresh_for_return p … l1 n. 1066 #p #l1 elim l1 // * [1,2,3: * #x] #xs #IH #l2 #n * #H1 #H2 #H3 whd 1343 1067 try(@IH) // % [2: @IH //] elim l2 in H1 H3; normalize [*] 1344 1068 * [1,2,3: * #y] #ys #IH normalize … … 1350 1074 qed. 1351 1075 1352 lemma fresh_append : ∀ n,l1,l2.is_fresh_for_return l1 n → is_fresh_for_returnl2 n →1353 is_fresh_for_return (l1@l2) n.1354 # n #l1 lapply n -n elim l1 // * [1,2,3: * #x] #xs #IH #n #l2 [2: * #H1 ] #H2 #H31076 lemma fresh_append : ∀p.∀n,l1,l2.is_fresh_for_return p l1 n → is_fresh_for_return p l2 n → 1077 is_fresh_for_return p (l1@l2) n. 1078 #p #n #l1 lapply n -n elim l1 // * [1,2,3: * #x] #xs #IH #n #l2 [2: * #H1 ] #H2 #H3 1355 1079 [ % // @IH //] @IH // 1356 1080 qed. 1357 1081 1358 definition labels_of_prog : ∀p,p' .Program p p'→ ? ≝1359 λp,p', prog.foldr … (λx,l.l @ (get_labels_of_code … (f_body … x)))1082 definition labels_of_prog : ∀p,p',l_p.Program p p' l_p → ? ≝ 1083 λp,p',l_p,prog.foldr … (λx,l.l @ (get_labels_of_code … (f_body … x))) 1360 1084 (get_labels_of_code … (main … prog)) (env … prog). 1361 1085 1362 lemma cast_return_append : ∀ l1,l2.cast_return_to_cost_labels… (l1 @ l2) =1363 (cast_return_to_cost_labels … l1) @ (cast_return_to_cost_labels… l2).1364 # l1 #l2 @(sym_eq … (map_append …)) qed.1086 lemma cast_return_append : ∀p.∀l1,l2.cast_return_to_cost_labels p … (l1 @ l2) = 1087 (cast_return_to_cost_labels p … l1) @ (cast_return_to_cost_labels p … l2). 1088 #p #l1 #l2 @(sym_eq … (map_append …)) qed. 1365 1089 1366 1090 include alias "arithmetics/nat.ma". 1367 1091 1368 1092 1369 lemma is_fresh_code : ∀p.∀i : Instructions p. 1370 is_fresh_for_return (get_labels_of_code … i) (compute_max_n … i). 1371 #p #main elim main // 1372 [ #seq * [| * #lbl] #i #IH normalize // @(fresh_keep_n_ok … IH) 1373 inversion(leb ? lbl) // @leb_elim #Hlbl #EQ destruct normalize nodelta 1374 /2 by le_S/ 1375 | #cond * #ltrue #i1 * #lfalse #i2 #i3 #IH1 #IH2 #IH3 whd in ⊢ (?%%); 1093 lemma is_fresh_code : ∀p,l_p.∀i : Instructions p l_p. 1094 is_fresh_for_return l_p (get_labels_of_code … i) (compute_max_n … i). 1095 #p #l_p #main elim main // 1096 [ #seq * [| * #lbl] #i #IH normalize // @(fresh_keep_n_ok … IH) // 1097 | #cond * #ltrue #i1 * #lfalse #i2 #i3 #IH1 #IH2 #IH3 whd in ⊢ (??%%); 1376 1098 @fresh_append 1377 [ @(fresh_keep_n_ok … IH1) @ le_S @max_1 @max_1 @max_1 @max_1 //1099 [ @(fresh_keep_n_ok … IH1) @max_1 @max_1 @max_1 @max_1 // 1378 1100 | @fresh_append 1379 [ @(fresh_keep_n_ok … IH2) @ le_S @max_1 @max_1 @max_1 @max_2 //1380 | @(fresh_keep_n_ok … IH3) @ le_S @max_1 @max_1 @max_2 //1101 [ @(fresh_keep_n_ok … IH2) @max_1 @max_1 @max_1 @max_2 // 1102 | @(fresh_keep_n_ok … IH3) @max_1 @max_1 @max_2 // 1381 1103 ] 1382 1104 ] 1383 | #cond * #ltrue #i1 * #lfalse #i2 #IH1 #IH2 whd in ⊢ (? %%); @fresh_append1384 [ @(fresh_keep_n_ok … IH1) @ le_S @max_1 @max_1 @max_1 //1385 | @(fresh_keep_n_ok … IH2) @ le_S @max_1 @max_1 @max_2 //1386 ] 1387 | #f #act_p * [| * #lbl] #i #IH whd in ⊢ (? %%); //1388 change with ([?]@?) in ⊢ (? %?); @fresh_append1389 [ whd % // @le_S @max_1 //1390 | @(fresh_keep_n_ok … IH) @ le_S @max_2 //1391 ] 1392 | * #lin #io * #lout #i #IH whd in ⊢ (? %%); @(fresh_keep_n_ok … IH)1393 @ le_S @max_1 @max_1 //1105 | #cond * #ltrue #i1 * #lfalse #i2 #IH1 #IH2 whd in ⊢ (??%%); @fresh_append 1106 [ @(fresh_keep_n_ok … IH1) @max_1 @max_1 @max_1 // 1107 | @(fresh_keep_n_ok … IH2) @max_1 @max_1 @max_2 // 1108 ] 1109 | #f #act_p * [| * #lbl] #i #IH whd in ⊢ (??%%); // 1110 change with ([?]@?) in ⊢ (??%?); @fresh_append 1111 [ whd % // 1112 | @(fresh_keep_n_ok … IH) @max_2 // 1113 ] 1114 | * #lin #io * #lout #i #IH whd in ⊢ (??%%); @(fresh_keep_n_ok … IH) 1115 @max_1 @max_1 // 1394 1116 ] 1395 1117 qed. 1396 1118 1397 lemma is_fresh_fresh_for_prog : ∀p,p' .∀prog : Program p p'.1398 is_fresh_for_return (labels_of_prog … prog) (fresh_for_prog … prog).1399 #p #p' * #env #main whd in match fresh_for_prog; normalize nodelta whd in ⊢ (?%?);1400 elim env // * #sig #cost #i #tail #IH whd in ⊢ (? %?); @fresh_append1119 lemma is_fresh_fresh_for_prog : ∀p,p',l_p.∀prog : Program p p' l_p. 1120 is_fresh_for_return … (labels_of_prog … prog) (fresh_for_prog … prog). 1121 #p #p' #l_p * #env #main whd in match fresh_for_prog; normalize nodelta whd in ⊢ (??%?); 1122 elim env // * #sig #cost #i #tail #IH whd in ⊢ (??%?); @fresh_append 1401 1123 [ @(fresh_keep_n_ok … IH) @fresh_aux_ok @max_1 // 1402 1124 | @(fresh_keep_n_ok … (is_fresh_code … i)) whd in match fresh_for_prog_aux; normalize nodelta 1403 whd in ⊢ (?? %); elim tail [ @max_2 // ] #hd1 #tl1 #IH1 @(transitive_le… IH1)1404 whd in ⊢ (?? %); change with (fresh_for_prog_aux ?? (mk_Program ?? tl1 main) ?) in ⊢ (?%%);1125 whd in ⊢ (????%); elim tail [ @max_2 // ] #hd1 #tl1 #IH1 @(trans_po_rel … IH1) 1126 whd in ⊢ (????%); change with (fresh_for_prog_aux ??? (mk_Program ??? tl1 main) ?) in ⊢ (???%%); 1405 1127 @fresh_aux_ok @max_1 // 1406 1128 ] 1407 1129 qed. 1408 1130 1409 lemma subset_def : ∀A : DeqSet.∀l1,l2 : list A.(∀x.x ∈l1 → x ∈ l2) → l1 ⊆ l2. 1410 #A #l1 elim l1 // #x #xs #IH #l2 #H % 1411 [ @memb_to_mem >H // >memb_hd // 1412 | @IH #y #H1 @H >memb_cons // >H1 // 1413 ] 1414 qed. 1415 1416 lemma memb_cast_return : ∀keep,x.x ∈ cast_return_to_cost_labels keep → 1417 ∃ y.x = a_return_post y ∧ bool_to_Prop (y ∈ keep). 1418 #keep elim keep 1131 lemma memb_cast_return : ∀p.∀keep,x.x ∈ cast_return_to_cost_labels p keep → 1132 ∃ y.x = a_return_post … y ∧ bool_to_Prop (y ∈ keep). 1133 #p #keep elim keep 1419 1134 [ #x *] #x #xs #IH #y whd in match cast_return_to_cost_labels; 1420 1135 whd in match (map ????); whd in match (memb ???); inversion(y==x) … … 1424 1139 qed. 1425 1140 1426 lemma subset_append : ∀A.∀l1,l2,l3 : list A.l1 ⊆ l3 → l2 ⊆ l3 → (l1 @ l2) ⊆ l3. 1427 #A #l1 elim l1 // #x #xs #IH #l2 #l3 * #H1 #H2 #H3 % /2/ 1428 qed. 1429 1430 lemma subset_def_inv : ∀A.∀l1,l2 : list A. l1 ⊆ l2 → ∀x.mem … x l1 → mem … x l2. 1431 #A #l1 elim l1 [ #l2 * #x * ] #x #xs #IH #l2 * #H1 #H2 #y * 1432 [ #EQ destruct // | #H3 @IH // ] 1433 qed. 1434 1435 lemma transitive_subset : ∀A.transitive … (subset A). 1436 #A #l1 elim l1 // #x #xs #IH #l2 #l3 * #H1 #H2 #H3 % 1437 [ @(subset_def_inv … H3) // | @IH // ] 1438 qed. 1439 1440 lemma subset_append_h1 : ∀A.∀l1,l2,l3 : list A.l1 ⊆ l3 → l1 ⊆ (l3 @ l2). 1441 #A #l1 elim l1 // #x #x2 #IH #l2 #l3 * #H1 #H2 % [ @mem_append_l1 // | @IH // ] 1442 qed. 1443 1444 lemma subset_append_h2 : ∀A.∀l1,l2,l3 : list A.l2 ⊆ l3 → l2 ⊆ (l1 @ l3). 1445 #A #l1 elim l1 // #x #xs #IH #l2 elim l2 // #y #ys #IH #l3 * #H1 #H2 % 1446 [ @mem_append_l2 // | @IH // ] 1447 qed. 1448 1449 lemma lab_to_keep_in_prog : ∀p,p'.∀prog : Program p p'. 1141 lemma lab_to_keep_in_prog : ∀p,p',l_p.∀prog : Program p p' l_p. 1450 1142 ∀t_prog,m,keep.trans_prog … prog = 〈t_prog,m,keep〉 → 1451 (cast_return_to_cost_labels keep) ⊆ (labels_of_prog p p'prog).1452 #p #p' * #env #main #t_prog #m #keep whd in match trans_prog; normalize nodelta1143 (cast_return_to_cost_labels l_p keep) ⊆ (labels_of_prog p p' l_p prog). 1144 #p #p' #l_p * #env #main #t_prog #m #keep whd in match trans_prog; normalize nodelta 1453 1145 @pair_elim * #env1 #fresh * #m1 #keep1 #EQenv1 normalize nodelta #EQ destruct 1454 1146 lapply EQenv1 -EQenv1 lapply keep1 -keep1 lapply m1 -m1 lapply fresh -fresh 1455 lapply env1 -env1 generalize in match (fresh_for_prog ??? ); elim env1147 lapply env1 -env1 generalize in match (fresh_for_prog ????); elim env 1456 1148 [ #n #t_env #n1 #m #keep whd in ⊢ (??%? → ?); #EQ destruct whd in match (append ???); 1457 1149 @subset_def #x #H whd in match (labels_of_prog); normalize nodelta … … 1460 1152 | #x #xs #IH #n #t_env #n1 #m #keep whd in ⊢ (??%? → ?); @pair_elim 1461 1153 * #t_env_tail #fresh_tail * #t_m_tail #t_keep_tail 1462 change with (translate_env ???? ) in match (foldr ?????); #EQt_env_tail1154 change with (translate_env ?????) in match (foldr ?????); #EQt_env_tail 1463 1155 normalize nodelta #EQ1 destruct >cast_return_append @subset_append 1464 1156 [ >cast_return_append @subset_append … … 1467 1159 >memb_append_l2 // @(lab_to_keep_in_domain … H) 1468 1160 | whd in match labels_of_prog; normalize nodelta whd in match (foldr ?????); 1469 change with (labels_of_prog ?? (mk_Program?? xs ?)) in match (foldr ?????);1161 change with (labels_of_prog ??? (mk_Program ??? xs ?)) in match (foldr ?????); 1470 1162 @subset_append_h1 @(transitive_subset … (IH … EQt_env_tail)) 1471 1163 >cast_return_append @subset_append_h1 // 1472 1164 ] 1473 1165 | whd in match labels_of_prog; normalize nodelta whd in match (foldr ?????); 1474 change with (labels_of_prog ?? (mk_Program?? xs ?)) in match (foldr ?????);1166 change with (labels_of_prog ??? (mk_Program ??? xs ?)) in match (foldr ?????); 1475 1167 @subset_append_h1 @(transitive_subset … (IH … EQt_env_tail)) 1476 1168 >cast_return_append @subset_append_h2 // … … 1479 1171 qed. 1480 1172 1481 lemma fresh_call_post_trans_ok : ∀p .∀i : Instructionsp.∀n,l.1482 n ≤fresh … (call_post_trans … i n l).1483 #p # i elim i //1484 qed. 1485 1486 lemma fresh_translate_env_ok : ∀p,p' .∀env,t_env : list (env_item p p').∀n,n1,m,keep.1487 translate_env … env n = 〈t_env,n1,m,keep〉 → n ≤n1.1488 #p #p' # env elim env1173 lemma fresh_call_post_trans_ok : ∀p,l_p.∀i : Instructions p l_p.∀n,l. 1174 n ⊑^{l_p} fresh … (call_post_trans … i n l). 1175 #p #l_p #i elim i // 1176 qed. 1177 1178 lemma fresh_translate_env_ok : ∀p,p',l_p.∀env,t_env : list (env_item p p' l_p).∀n,n1,m,keep. 1179 translate_env … env n = 〈t_env,n1,m,keep〉 → n ⊑^{l_p} n1. 1180 #p #p' #l_p #env elim env 1489 1181 [ #t_env #n #n1 #m #keep whd in ⊢ (??%? → ?); #EQ destruct // ] 1490 1182 #x #xs #IH #t_env #n #n1 #m #keep whd in ⊢ (??%? → ?); 1491 change with (translate_env ???? ) in match (foldr ?????); @pair_elim1183 change with (translate_env ?????) in match (foldr ?????); @pair_elim 1492 1184 * #t_env_tail #fresh_tail * #t_m_tail #t_keep_tail #EQt_env_tail normalize nodelta 1493 #EQ destruct @(trans itive_le… (IH … EQt_env_tail)) @fresh_call_post_trans_ok1185 #EQ destruct @(trans_po_rel … (IH … EQt_env_tail)) @fresh_call_post_trans_ok 1494 1186 qed. 1495 1187 … … 1498 1190 no_duplicates_labels … prog → 1499 1191 let 〈t_prog,m,keep〉 ≝ trans_prog … prog in 1500 ∀f,env_it.lookup p p (env … prog) f = return env_it →1192 ∀f,env_it.lookup p p p (env … prog) f = return env_it → 1501 1193 let dom ≝ get_labels_of_code … (f_body … env_it) in 1502 ∃env_it',n.is_fresh_for_return keep n ∧lookup p p (env … t_prog) f = return env_it' ∧1194 ∃env_it',n.is_fresh_for_return p keep n ∧lookup p p p (env … t_prog) f = return env_it' ∧ 1503 1195 let info ≝ call_post_trans … (f_body … env_it) n (nil ?) in 1504 1196 t_code … info = f_body … env_it' ∧ 1505 get_element … m (a_call (f_lab … env_it')) = (a_call(f_lab … env_it')) :: gen_labels … info ∧1197 get_element … m (a_call … (f_lab … env_it')) = (a_call … (f_lab … env_it')) :: gen_labels … info ∧ 1506 1198 f_sig … env_it = f_sig … env_it' ∧ f_lab … env_it = f_lab … env_it' ∧ 1507 same_fresh_map_on dom m (lab_map … info) ∧ same_to_keep_ondom keep (lab_to_keep … info).1199 same_fresh_map_on … dom m (lab_map … info) ∧ same_to_keep_on … dom keep (lab_to_keep … info). 1508 1200 #p #prog inversion(trans_prog … prog) * #t_prog0 #m0 #keep0 #EQt_prog 1509 1201 lapply EQt_prog normalize nodelta … … 1512 1204 whd in match trans_prog; normalize nodelta 1513 1205 @pair_elim 1514 cut(fresh_for_prog ?? prog ≤ fresh_for_prog?? (mk_Program … env main)) [ >EQprog //]1515 generalize in match (fresh_for_prog ??? ) in ⊢ (??% → %);1206 cut(fresh_for_prog ??? prog ⊑^{p} fresh_for_prog ??? (mk_Program … env main)) [ >EQprog //] 1207 generalize in match (fresh_for_prog ????) in ⊢ (????% → %); 1516 1208 lapply t_prog0 lapply m0 lapply keep0 1517 1209 elim env in ⊢ (?→ ? → ? → ? → ? → %); … … 1519 1211 * #hd_sig #hd_lab #hd_code #tail #IH #keep #m #t_prog #fresh1 #Hfresh1 * #env' #fresh * #m' #keep' 1520 1212 normalize in ⊢ (% → ?); normalize nodelta @pair_elim * #env_tail #fresh_tail 1521 * #m_tail #keep_tail change with (translate_env ???? ) in ⊢ (??%? → ?); #EQtail normalize nodelta #EQ1 destruct(EQ1) #EQ2 destruct(EQ2)1213 * #m_tail #keep_tail change with (translate_env ?????) in ⊢ (??%? → ?); #EQtail normalize nodelta #EQ1 destruct(EQ1) #EQ2 destruct(EQ2) 1522 1214 whd in ⊢ (% → ?); whd in match (foldr ?????); * #Hhd_lab #H lapply(no_duplicates_append_r … H) 1523 change with (no_duplicates_labels p p (mk_Programp p tail main)) in match1524 (no_duplicates_labels p p (mk_Programp p tail main)); #no_dup_tail1215 change with (no_duplicates_labels p p p (mk_Program p p p tail main)) in match 1216 (no_duplicates_labels p p p (mk_Program p p p tail main)); #no_dup_tail 1525 1217 lapply(no_duplicates_append_l … H) #no_dup_head normalize nodelta 1526 1218 #f #env_it whd in ⊢ (??%? → ?); @eq_function_name_elim normalize nodelta … … 1534 1226 @(fresh_for_subset … (labels_of_prog … prog)) 1535 1227 [ @(lab_to_keep_in_prog … EQkeep1) | @is_fresh_fresh_for_prog ] 1536 | @(trans itive_le… (fresh_translate_env_ok … EQtail)) //1228 | @(trans_po_rel … (fresh_translate_env_ok … EQtail)) // 1537 1229 ] 1538 1230 | whd in ⊢ (??%?); @eq_function_name_elim [2: * #H @⊥ @H %] #_ normalize nodelta … … 1583 1275 >memb_append_l2 // >IH % 1584 1276 ] 1585 >memb_append_l12 // inversion(memb ???) // #ABS @(memb_no_duplicates_append … (a_return_post x) … H)1277 >memb_append_l12 // inversion(memb ???) // #ABS @(memb_no_duplicates_append … (a_return_post … x) … H) 1586 1278 // @⊥ 1587 1279 (* subproof with no nice statement *) … … 1599 1291 #H1 #H2 #H3 cases(memb_append … H3) -H3 1600 1292 [ #H3 change with ([?]@?) in match (?::?) in H2; 1601 lapply(no_duplicates_append_commute … H2) -H2 * #_ #H4 @(memb_no_duplicates_append … (a_return_post x) … H4)1293 lapply(no_duplicates_append_commute … H2) -H2 * #_ #H4 @(memb_no_duplicates_append … (a_return_post … x) … H4) 1602 1294 [ whd in match (append ???); >memb_append_l1 // >(lab_to_keep_in_domain … (eq_true_to_b … H3)) % 1603 1295 | // … … 1630 1322 % [2: #x #Hx <same_to_keep // >associative_append @memb_append_l22 1631 1323 inversion(memb ???) // #ABS lapply(lab_to_keep_in_domain … (eq_true_to_b … ABS)) 1632 #ABS1 @(memb_no_duplicates_append … (a_return_post x) … H) //1324 #ABS1 @(memb_no_duplicates_append … (a_return_post … x) … H) // 1633 1325 cases(lookup_ok_append … Henv_it) #l1 * #l2 * #EQ1 #EQ2 destruct(EQ1 EQ2) 1634 1326 >foldr_map_append >memb_append_l2 // >foldr_map_append >memb_append_l1 // … … 1658 1350 >associative_append 1659 1351 @get_element_append_r1 % >domain_of_associative_list_append #ABS cases(memb_append … ABS) 1660 [ whd in match (memb ???); inversion(a_call (f_lab … new_env_it)== a_callhd_lab)1352 [ whd in match (memb ???); inversion(a_call … (f_lab … new_env_it)== a_call … hd_lab) 1661 1353 #EQ_hdlab normalize nodelta 1662 1354 [2: whd in ⊢ (??%? → ?); #EQ destruct ] … … 1665 1357 * #ABS1 @ABS1 >memb_append_l2 // >memb_append_l2 // 1666 1358 >memb_append_l1 // whd in ⊢ (??%?); >EQ_f_lab >(\b (refl …)) // 1667 | #ABS1 @(memb_no_duplicates_append … (a_call (f_lab … new_env_it)) … H)1359 | #ABS1 @(memb_no_duplicates_append … (a_call … (f_lab … new_env_it)) … H) 1668 1360 [ @(lab_map_in_domain … (eq_true_to_b … ABS1)) 1669 1361 | cases(lookup_ok_append … Henv_it) … … 1676 1368 qed. 1677 1369 1678 lemma len_append : ∀S : abstract_status.∀st1,st2,st3 : S.1679 ∀t1 : raw_trace … st1 st2.∀t2 : raw_trace … st2 st3.1680 len … (t1 @ t2) = len … t1 + len … t2.1681 #S #st1 #st2 #st3 #t1 elim t1 normalize //1682 qed.1683 1370 (* 1684 1371 axiom permute_ok : ∀A.∀l1,l2,l3,l4,l5,l6,l7,l8,l9,x,y. … … 1689 1376 (((x ::l4 @y ::l3) @l8) @l9)). 1690 1377 *) 1691 lemma append_premeasurable : ∀S : abstract_status. 1692 ∀st1,st2,st3 : S.∀t1 : raw_trace … st1 st2.∀t2: raw_trace … st2 st3. 1693 pre_measurable_trace … t1 → pre_measurable_trace … t2 → 1694 pre_measurable_trace … (t1 @ t2). 1695 #S #st1 #st2 #st3 #t1 #t2 #Hpre lapply t2 -t2 elim Hpre -Hpre // 1696 [ #s1 #s2 #s3 #l #prf #tl #Hclass #Hcost #pre_tl #IH #t2 #pr3_t2 1697 %2 // @IH // 1698 | #s1 #s2 #s3 #l #Hclass #prf #tl #ret_l #pre_tl #IH #t2 #pre_t2 %3 // @IH // 1699 | #s1 #s2 #s3 #l #prf #tl #Hclass #call #callp #pre_tl #IH #t2 #pre_t2 %4 // @IH // 1700 | #s1 #s2 #s3 #s4 #s5 #l1 #l2 #prf1 #tl1 #tl2 #prf2 #Hclass1 #Hclass3 #call_l1 1701 #nopost #pre_tl1 #pre_tl2 #succ #unlab #IH1 #IH2 #t2 #pre_t2 1702 normalize >append_associative in ⊢ (????(???????%)); %5 // @IH2 // 1703 ] 1704 qed. 1705 1706 lemma correctness_lemma : ∀p,p',prog. 1707 no_duplicates_labels … prog → 1708 let 〈t_prog,m,keep〉 ≝ trans_prog … prog in 1709 ∀s1,s2,s1' : state p.∀abs_top,abs_tail. 1710 ∀t : raw_trace (operational_semantics … p' prog) s1 s2. 1711 pre_measurable_trace … t → 1712 state_rel … m keep abs_top abs_tail s1 s1' → 1713 ∃abs_top'.∃s2'.∃t' : raw_trace (operational_semantics … p' t_prog) s1' s2'. 1714 state_rel … m keep abs_top' abs_tail s2 s2' ∧ 1715 is_permutation ? ((abs_top @ (map_labels_on_trace m (get_costlabels_of_trace … t))) @ abs_tail) 1716 (((get_costlabels_of_trace … t') @ abs_top' ) @ abs_tail) ∧ 1717 len … t = len … t' ∧ 1718 (∀k.∀i.option_hd … (cont … s2) = Some ? 〈ret_act (None ?),i〉 → 1719 cont … s1 = k @ cont … s2 → 1720 ∃k'.cont … s1' = k' @ cont … s2' ∧ |k| = |k'|) 1721 ∧ pre_measurable_trace … t' ∧ 1722 (pre_silent_trace … t → pre_silent_trace … t'). 1723 #p #p' #prog #no_dup @pair_elim * #t_prog #m #keep #EQtrans 1724 #s1 #s2 #s1' #abs_top #abs_tail #t #Hpre lapply abs_top -abs_top lapply abs_tail 1725 -abs_tail lapply s1' -s1' elim Hpre 1726 [ #st #Hst #s1' #abs_tail #abs_top #H %{abs_top} %{s1'} %{(t_base …)} 1727 cut(? : Prop) 1728 [3: #Hpre' % 1729 [ % 1730 [ % 1731 [ % [2: %] %{H} whd in match (get_costlabels_of_trace ????); >append_nil 1732 @is_permutation_eq 1733 | #k #i #_ #EQcont_st %{[ ]} % // @sym_eq @eq_f @append_l1_injective_r // 1734 ] 1735 | @Hpre' 1736 ] 1737 | #_ %1 /2 by head_of_premeasurable_is_not_io/ 1738 ] 1739 | skip 1740 ] 1741 %1 whd in H; cases(check_continuations ?????) in H; [*] ** #H1 #abs_top_cont 1742 #abs_tail_cont normalize nodelta **** #_ #rel #_ #H2 #_ normalize 1743 inversion(code ??) normalize nodelta 1744 [|#r_t| #seq #lbl #i #_ | #cond #ltrue #i1 #lfalse #i2 #i3 #_ #_ #_ 1745 | #cond #ltrue #i1 #lfalse #i2 #_ #_ | #f #act_p #rlbl #i #_ 1746 | #lin #io #lout #i #_ ] 1747 #EQcode_s1' normalize nodelta [|*: % #EQ destruct] <H2 1748 >EQcode_s1' in rel; whd in ⊢ (??%% → ?); #EQ destruct(EQ) 1749 normalize in Hst; <e1 in Hst; normalize nodelta // 1750 | #st1 #st2 #st3 #lab whd in ⊢ (% → ?); #H inversion H in ⊢ ?; #st11 #st12 1751 [ * #lab1 #new_code #new_cont #EQcode11 #EQcont11 #EQcode12 #EQcont12 #EQstore 1752 #Hio11 #EQio12 #Hl1 #EQ1 #EQ2 #EQ3 #EQ4 destruct #tl #class_st11 * #opt_l #EQ destruct(EQ) 1753 #pre_tl #IH #s1' #abs_tail #abs_top whd in match state_rel in ⊢ (% → ?); normalize nodelta >EQcont11 in ⊢ (% → ?); 1754 whd in match check_continuations; normalize nodelta whd in match (foldr2 ???????); 1755 inversion (cont ? s1') [ #_ *] * #l1' #new_code' #new_cont' #_ #EQconts1' 1756 normalize nodelta change with (check_continuations ?????) in match (foldr2 ???????); 1757 inversion(check_continuations ?????) [ #_ *] ** #H1 #l2 #ll2 #EQHl2 1758 >m_return_bind normalize nodelta inversion (call_post_clean ?????) [ #_ ***** ] 1759 * #l3 #code_c_st12 #EQcode_c_st12 normalize nodelta 1760 cases l1' in EQconts1'; (*in Hio11 Hl1 EQcont11 H;*) normalize nodelta 1761 [ #x #y #_ (*#_ #_ #_ #H*) ***** 1762 | #x #_ cases (ret_costed_abs ??) normalize nodelta [|#c] ******[|*] #EQ @⊥ >EQ in Hl1; 1763 normalize * /2/ ] * 1764 [ #EQconts1' normalize nodelta ****** #EQ destruct(EQ) 1765 #HH1 #EQ destruct(EQ) >EQcode11 in ⊢ (% → ?); inversion(code … s1') 1766 [ 1767 | #x 1768 | #seq #lbl #i #_ 1769 | #cond #ltrue #i1 #lfalse #i2 #i3 #_ #_ #_ 1770 | #cond #ltrue #i1 #lfalse #i2 #_ #_ 1771 | #f #act_p #ret_post #i #_ 1772 | #l_in #io #l_out #i #_ 1773 ] 1774 [|*: #_ whd in ⊢ (??%% → ?); [ #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ)] 1775 cases(call_post_clean ?????) normalize nodelta 1776 [1,3,5,7,9: #EQ destruct(EQ)] 1777 [ * #z1 #z2 cases lbl normalize nodelta [|#z3 cases(?==?) normalize nodelta] 1778 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1779 | * #z1 #z2 cases(call_post_clean ?????) [| * #z3 #z4 ] whd in ⊢ (??%% → ?); 1780 [2: cases(call_post_clean ?????) normalize nodelta 1781 [| * #z5 #z6 cases(?∧?) normalize nodelta]] whd in ⊢ (??%% → ?); 1782 #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1783 | * #z1 #z2 cases(call_post_clean ?????) 1784 [| * #z3 #z4 >m_return_bind cases(?∧?) normalize nodelta] 1785 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1786 | * #z1 #z2 cases ret_post normalize nodelta 1787 [| #z3 cases(memb ???) normalize nodelta [ cases(?==?) normalize nodelta] ] 1788 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1789 | * #z1 #z2 cases(?∧?) normalize nodelta 1790 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1791 ] 1792 ] 1793 #EQcodes1' whd in ⊢ (??%% → ?); #EQ destruct(EQ) #EQstore11 #EQio11 #EQ destruct(EQ) 1794 cases(IH (mk_state ? new_code' new_cont' (store … s1') false) ll2 abs_top …) 1795 [2: whd whd in match check_continuations; normalize nodelta 1796 change with (check_continuations ?????) in match (foldr2 ???????); >EQHl2 1797 normalize nodelta % // % // % // % // @EQcode_c_st12 ] 1798 #abs_top' * #st3' * #t' ***** 1799 #Hst3st3' #EQcosts #EQlen #stack_safety #pre_t' #Hsil 1800 %{abs_top'} %{st3'} %{(t_ind … (cost_act (None ?)) … t')} 1801 [ @hide_prf whd @(empty ????? 〈(cost_act (None ?)),?〉) 1802 [3: assumption |4: assumption |*:] /3 by nmk/ ] 1803 % [ % 1804 [ % [% [2: whd in ⊢ (??%%); >EQlen % ] %{Hst3st3'} @EQcosts] 1805 * [| #hd #tl] #i #EQcont_st3 >EQcont11 #EQ whd in EQ : (???%); 1806 [ <EQ in EQcont_st3; whd in ⊢ (??%% → ?); #EQ1 destruct ] 1807 destruct cases(stack_safety … EQcont_st3 … e0) #k1 * 1808 #EQk1 #EQlength %{(〈cost_act (None ?),new_code'〉::k1)} 1809 % [ >EQk1 % | whd in ⊢ (??%%); @eq_f // ] ] 1810 %2 // [2: % //] normalize >EQcodes1' normalize nodelta <EQio11 1811 normalize in class_st11; >EQcode11 in class_st11; // ] 1812 #H inversion H in ⊢ ?; 1813 [ #H1 #H2 #H3 #H4 #H5 #H6 destruct ] #H8 #H9 #H10 #H11 #H12 #H13 #H14 #H15 #H16 #H17 #H18 #H19 1814 destruct %2 1815 [ % whd in ⊢ (??%% → ?); >EQcodes1' normalize nodelta <EQio11 inversion(io_info … H8) 1816 normalize nodelta [2: #_ #EQ destruct] #ABS lapply class_st11 whd in match (as_classify ??); 1817 >EQcode11 normalize nodelta >ABS normalize nodelta * /2/ 1818 ] 1819 @Hsil // 1820 | #lbl #EQconts1' normalize nodelta ******* #EQ destruct(EQ) 1821 #HH1 #EQ destruct(EQ) #EQget_el >EQcode11 in ⊢ (% → ?); 1822 inversion(code … s1') 1823 [ 1824 | #x 1825 | #seq #lbl #i #_ 1826 | #cond #ltrue #i1 #lfalse #i2 #i3 #_ #_ #_ 1827 | #cond #ltrue #i1 #lfalse #i2 #_ #_ 1828 | #f #act_p #ret_post #i #_ 1829 | #l_in #io #l_out #i #_ 1830 ] 1831 [|*: #_ whd in ⊢ (??%% → ?); 1832 [ #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1833 | cases(call_post_clean ?????) normalize nodelta 1834 [| * #z2 #z3 cases lbl normalize nodelta 1835 [| #z4 cases(?==?) normalize nodelta] ] 1836 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1837 | cases(call_post_clean ?????) normalize nodelta 1838 [| * #z2 #z3 cases(call_post_clean ?????) 1839 [| * #z4 #z5 >m_return_bind cases(call_post_clean ?????) 1840 [| * #z6 #z7 >m_return_bind cases(?∧?) normalize nodelta ]]] 1841 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1842 | cases(call_post_clean ?????) normalize nodelta 1843 [| * #z2 #z3 cases(call_post_clean ?????) 1844 [| * #z4 #z5 >m_return_bind cases(?∧?) normalize nodelta ]] 1845 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1846 | cases(call_post_clean ?????) normalize nodelta 1847 [| * #z1 #z2 cases ret_post normalize nodelta 1848 [| #z3 cases(memb ???) normalize nodelta 1849 [ cases (?==?) normalize nodelta]]] 1850 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1851 | cases(call_post_clean ?????) normalize nodelta 1852 [| * #z1 #z2 cases(?∧?) normalize nodelta ] 1853 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1854 ] 1855 ] 1856 #EQcodes1' whd in ⊢ (??%% → ?); #EQ destruct(EQ) #EQstore1 #EQio #EQ destruct(EQ) 1857 cases(IH (mk_state ? new_code' new_cont' (store … s1') false) ll2 l3 …) 1858 [2: whd whd in match check_continuations; normalize nodelta 1859 change with (check_continuations ?????) in match (foldr2 ???????); 1860 >EQHl2 in ⊢ (match % with [ _⇒ ? | _ ⇒ ?]); 1861 normalize nodelta % // % // % // % // @EQcode_c_st12 ] 1862 #abs_top' * #st3' * #t' ***** #Hst3st3' #EQcost #EQlen #stack_safety 1863 #pre_t' #Hsil 1864 %{abs_top'} %{st3'} 1865 %{(t_ind … (cost_act (Some ? lbl)) … t')} 1866 [ @hide_prf whd @(empty ????? 〈(cost_act (Some ? lbl)),?〉) 1867 [3: assumption |4: assumption |*:] /3 by nmk/ ] 1868 cut(? : Prop) 1869 [3: #Hpre' % 1870 [ % 1871 [ % [% [2: whd in ⊢ (??%%); >EQlen % ] %{Hst3st3'} 1872 >(map_labels_on_trace_append … [lbl]) (*CSC: funzionava prima*) 1873 whd in match (map_labels_on_trace ? [lbl]); >append_nil >EQget_el 1874 @is_permutation_cons @EQcost 1875 | * [| #hd #tl] #i #EQcont_st3 >EQcont11 #EQ whd in EQ : (???%); 1876 [ <EQ in EQcont_st3; whd in ⊢ (??%% → ?); #EQ1 destruct ] 1877 destruct cases(stack_safety … EQcont_st3 … e0) #k1 * 1878 #EQk1 #EQlength %{(〈cost_act (Some ? lbl),new_code'〉::k1)} 1879 % [ >EQk1 % | whd in ⊢ (??%%); @eq_f // ] 1880 ]] 1881 @Hpre' ] 1882 #h inversion h in ⊢ ?; [ #H21 #H22 #H23 #H24 #H25 #H26 destruct] 1883 #H28 #H29 #H30 #H31 #H32 #H33 #H34 #H35 #H36 #H37 #H38 #H39 destruct 1884 |skip | 1885 %2 // [2: % //] normalize >EQcodes1' normalize nodelta <EQio 1886 normalize in class_st11; >EQcode11 in class_st11; // 1887 ]] 1888 | #seq #i #mem * [|#lbl] #EQcode_st11 #EQeval_seq #EQ destruct(EQ) #EQcont 1889 #EQ destruct(EQ) #EQio11 #EQio12 #EQ1 #EQ2 #EQ3 #EQ4 destruct #t #Hst11 1890 #_ #Hpre_t #IH #s1' #abs_tail #abs_top whd in ⊢ (% → ?); 1891 inversion(check_continuations ?????) normalize nodelta [1,3: #_ *] 1892 ** #H1 #abs_top_cont #abs_tail_cont #EQcheck normalize nodelta **** 1893 #HH1 [ >EQcode_st11 in ⊢ (% → ?); | >EQcode_st11 in ⊢ (% → ?); ] 1894 inversion(code ??) 1895 [1,2,4,5,6,7,8,9,11,12,13,14: 1896 [1,7: #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1897 |2,8: #r_t #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1898 |3,9: #cond #ltrue #i1 #lfalse #i2 #i3 #_ #_ #_ #_ 1899 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 1900 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 1901 [2,4: * #z3 #z4 >m_return_bind cases(call_post_clean ?????) 1902 [2,4: * #z5 #z6 >m_return_bind cases(?∧?) normalize nodelta ]]] 1903 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1904 |4,10: #cond #ltrue #i1 #lfalse #i2 #_ #_ #_ 1905 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 1906 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 1907 [2,4: * #z3 #z4 >m_return_bind cases(?∧?) normalize nodelta ]] 1908 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1909 |5,11: #f #act_p #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 1910 normalize nodelta 1911 [2,4: * #z1 #z2 cases lbl normalize nodelta 1912 [2,4: #l1 cases(memb ???) normalize nodelta 1913 [1,3: cases(?==?) normalize nodelta ]]] 1914 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1915 |6,12: #lin #io #lout #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 1916 normalize nodelta 1917 [2,4: * #z1 #z2 cases(?∧?) normalize nodelta] 1918 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1919 ] 1920 ] 1921 #seq1 #opt_l #i' #_ #EQcode_s1' 1922 change with (m_bind ?????) in ⊢ (??%? → ?); 1923 inversion(call_post_clean ????) [1,3: #_ whd in ⊢ (??%% → ?); #EQ destruct] 1924 * #gen_labs #i'' #EQclean >m_return_bind inversion(opt_l) normalize nodelta 1925 [2,4: #lab1 ] #EQ destruct(EQ) 1926 [1,2: inversion(?==?) normalize nodelta #EQget_el ] 1927 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1928 #EQstore #EQinfo #EQ destruct(EQ) 1929 cases(IH … (mk_state ? i' (cont … s1') (store … st12) (io_info … s1')) abs_tail_cont ?) 1930 [3,6: whd 1931 [ <EQcont in ⊢ (match % with [_ ⇒ ? | _⇒ ?]); >EQcheck in ⊢ (match % with [_ ⇒ ? | _⇒ ?]); 1932 | <EQcont in ⊢ (match % with [_ ⇒ ? | _⇒ ?]); >EQcheck in ⊢ (match % with [_ ⇒ ? | _⇒ ?]); 1933 ] 1934 normalize nodelta % // % // % // % // assumption 1935 |2,5: 1936 ] 1937 #abs_top1 * #s2' * #t' ***** #Hst3_s2' #EQcosts #EQlen #stack_safety 1938 #pre_t' #Hsil %{abs_top1} %{s2'} %{(t_ind … t')} 1939 [1,4: @hide_prf @(seq_sil … EQcode_s1') // 1940 |2,5: 1941 ] 1942 cut(? : Prop) 1943 [3,6: #Hpre' 1944 % 1945 1946 [1,3: % [1,3: % [1,3: % [2,4: whd in ⊢ (??%%); @eq_f @EQlen ] %{Hst3_s2'} [2: @EQcosts] 1947 whd in ⊢ (??%%); whd in match (get_costlabels_of_trace ????); 1948 whd in match (map_labels_on_trace ??); >(\P EQget_el) @is_permutation_cons 1949 @EQcosts 1950 |*: #k #i #EQcont_st3 >EQcont #EQ 1951 cases(stack_safety … EQcont_st3 … EQ) #k1 * #EQk1 #EQlength %{k1} % // 1952 ]] 1953 @Hpre' 1954 |*: #h inversion h in ⊢ ?; 1955 [1,3: #H41 #H42 #H43 #H44 #H45 #H46 destruct 1956 |*: #H55 #H56 #H57 #H58 #H59 #H60 #H61 #H62 #H63 #H64 #H65 #H66 destruct 1957 %2 [ /2 by head_of_premeasurable_is_not_io/ ] @Hsil // 1958 ] 1959 ] |1,4:] 1960 %2 // [2,4: % //] normalize >EQcode_s1' normalize nodelta % #EQ destruct 1961 | 1962 #cond #ltrue #i_true #lfalse #i_false #i #mem #EQcode_st11 #EQcond #EQcont_st12 1963 #EQ1 #EQ2 destruct(EQ1 EQ2) #EQio_st11 #EQio_st12 #EQ1 #EQ2 #EQ3 #EQ4 destruct #t 1964 #Hclass #_ #Hpre #IH #s1' #abs_tail #abs_top whd in ⊢ (% → ?); inversion(check_continuations ?????) 1965 [ #_ *] ** #H1 #abs_top_cont #abs_tail_cont #EQcheck normalize nodelta **** #HH1 1966 >EQcode_st11 in ⊢ (% → ?); inversion(code ??) 1967 [1,2,3,5,6,7: 1968 [ #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1969 | #r_t #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1970 | #seq #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 1971 normalize nodelta [2: * #z1 #z2 cases lbl normalize nodelta 1972 [2: #l1 cases(?==?) normalize nodelta]] 1973 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1974 | #cond #ltrue #i1 #lfalse #i2 #_ #_ #_ 1975 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 1976 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 1977 [2,4: * #z3 #z4 >m_return_bind cases(?∧?) normalize nodelta ]] 1978 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1979 | #f #act_p #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 1980 normalize nodelta 1981 [2,4: * #z1 #z2 cases lbl normalize nodelta 1982 [2,4: #l1 cases(memb ???) normalize nodelta 1983 [1,3: cases(?==?) normalize nodelta ]]] 1984 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1985 | #lin #io #lout #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 1986 normalize nodelta 1987 [2,4: * #z1 #z2 cases(?∧?) normalize nodelta] 1988 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 1989 ] 1990 ] 1991 #cond1 #ltrue1 #i_true1 #lfalse1 #ifalse1 #i' #_ #_ #_ 1992 #EQcode_s1' whd in ⊢ (??%% → ?); inversion(call_post_clean ?????) normalize nodelta 1993 [ #_ #EQ destruct ] * #gen_lab_i' #i'' #EQi'' inversion(call_post_clean ????) 1994 [ #_ whd in ⊢ (??%% → ?); #EQ destruct] * #gen_lab_ifalse1 #i_false' #EQi_false' 1995 >m_return_bind inversion(call_post_clean ????) [ #_ whd in ⊢ (??%% → ?); #EQ destruct] 1996 * #gen_lab_i_true' #i_true' #EQi_true' >m_return_bind inversion(?==?) normalize nodelta 1997 [2: #_ #EQ destruct] #EQget_ltrue1 inversion(?==?) normalize nodelta #EQget_lfalse1 1998 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) #EQstore11 #EQio11 1999 #EQ destruct(EQ) 2000 cases(IH … 2001 (mk_state ? i_true1 (〈cost_act (None ?),i'〉 :: cont … s1') (store … st12) (io_info … st12)) 2002 abs_tail_cont gen_lab_i_true') 2003 [2: whd whd in match (check_continuations ?????); 2004 >EQcont_st12 in ⊢ (match % with [_ ⇒ ? | _ ⇒ ?]); whd in match (foldr2 ???????); 2005 change with (check_continuations ?????) in match (foldr2 ???????); 2006 >EQcheck in ⊢ (match % with [_ ⇒ ? | _ ⇒ ?]); normalize nodelta 2007 >EQi'' normalize nodelta % // % // % // % [2: assumption ] % // % // 2008 ] #abs_top1 * #s2' * #t' ***** #Hst3_s2' #EQcost #EQlen 2009 #stack_safety #pre_t' #Hsil %{abs_top1} 2010 %{s2'} %{(t_ind … t')} 2011 [ @hide_prf @(cond_true … EQcode_s1') // 2012 | 2013 ] % [ 2014 % [ % 2015 [ % [2: whd in ⊢ (??%%); @eq_f // ] %{Hst3_s2'} 2016 whd in match (get_costlabels_of_trace ????); whd in match(map_labels_on_trace ??); 2017 >(\P EQget_ltrue1) whd in match (get_costlabels_of_trace ????) in ⊢ (???%); @is_permutation_cons 2018 assumption 2019 | #k #i1 #EQcont_st3 #EQcont_st11 2020 cases(stack_safety … (〈cost_act (None NonFunctionalLabel),i〉::k) … EQcont_st3 …) 2021 [2: >EQcont_st12 >EQcont_st11 % ] 2022 * [|#hd1 #tl1] * #EQk1 whd in ⊢ (??%% → ?); #EQlength destruct 2023 %{tl1} % // whd in EQk1 : (??%%); destruct // 2024 ]] 2025 %2 // [2: % //] normalize >EQcode_s1' normalize nodelta % #EQ destruct 2026 ] 2027 #h inversion h in ⊢ ?; 2028 [#H1 #H2 #H3 #H4 #H5 #H6 destruct 2029 | #H8 #H9 #H10 #H11 #H12 #H13 #H14 #H15 #H16 #H17 #H18 #H19 destruct 2030 ] 2031 | #cond #ltrue #i_true #lfalse #i_false #i #mem #EQcode_st11 #EQcond #EQcont_st12 2032 #EQ1 #EQ2 destruct(EQ1 EQ2) #EQio_st11 #EQio_st12 #EQ1 #EQ2 #EQ3 #EQ4 destruct #t 2033 #Hclass #_ #Hpre #IH #s1' #abs_tail #abs_top whd in ⊢ (% → ?); inversion(check_continuations ?????) 2034 [ #_ *] ** #H1 #abs_top_cont #abs_tail_cont #EQcheck normalize nodelta **** #HH1 2035 >EQcode_st11 in ⊢ (% → ?); inversion(code ??) 2036 [1,2,3,5,6,7: 2037 [ #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2038 | #r_t #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2039 | #seq #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2040 normalize nodelta [2: * #z1 #z2 cases lbl normalize nodelta 2041 [2: #l1 cases(?==?) normalize nodelta]] 2042 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2043 | #cond #ltrue #i1 #lfalse #i2 #_ #_ #_ 2044 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2045 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2046 [2,4: * #z3 #z4 >m_return_bind cases(?∧?) normalize nodelta ]] 2047 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2048 | #f #act_p #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2049 normalize nodelta 2050 [2,4: * #z1 #z2 cases lbl normalize nodelta 2051 [2,4: #l1 cases(memb ???) normalize nodelta 2052 [1,3: cases(?==?) normalize nodelta ]]] 2053 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2054 | #lin #io #lout #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2055 normalize nodelta 2056 [2,4: * #z1 #z2 cases(?∧?) normalize nodelta] 2057 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2058 ] 2059 ] #cond1 #ltrue1 #i_true1 #lfalse1 #ifalse1 #i' #_ #_ #_ 2060 #EQcode_s1' whd in ⊢ (??%% → ?); inversion(call_post_clean ?????) normalize nodelta 2061 [ #_ #EQ destruct ] * #gen_lab_i' #i'' #EQi'' inversion(call_post_clean ????) 2062 [ #_ whd in ⊢ (??%% → ?); #EQ destruct] * #gen_lab_ifalse1 #i_false' #EQi_false' 2063 >m_return_bind inversion(call_post_clean ????) [ #_ whd in ⊢ (??%% → ?); #EQ destruct] 2064 * #gen_lab_i_true' #i_true' #EQi_true' >m_return_bind inversion(?==?) normalize nodelta 2065 [2: #_ #EQ destruct] #EQget_ltrue1 inversion(?==?) normalize nodelta #EQget_lfalse1 2066 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) #EQstore11 #EQio11 2067 #EQ destruct(EQ) 2068 cases(IH … 2069 (mk_state ? ifalse1 (〈cost_act (None ?),i'〉 :: cont … s1') (store … st12) (io_info … st12)) 2070 abs_tail_cont gen_lab_ifalse1) 2071 [2: whd whd in match (check_continuations ?????); 2072 >EQcont_st12 in ⊢ (match % with [_ ⇒ ? | _ ⇒ ?]); whd in match (foldr2 ???????); 2073 change with (check_continuations ?????) in match (foldr2 ???????); 2074 >EQcheck in ⊢ (match % with [_ ⇒ ? | _ ⇒ ?]); normalize nodelta 2075 >EQi'' normalize nodelta % // % // % // % [2: assumption ] % // % // 2076 ] #abs_top1 * #s2' * #t' ***** #Hst3_s2' #EQcost #EQlen #stack_safety 2077 #pre_t' #Hsil 2078 %{abs_top1} %{s2'} %{(t_ind … t')} 2079 [ @hide_prf @(cond_false … EQcode_s1') // 2080 | 2081 ] % [ 2082 % [ % 2083 [ % [2: whd in ⊢ (??%%); @eq_f // ] %{Hst3_s2'} 2084 whd in match (get_costlabels_of_trace ????); whd in match(map_labels_on_trace ??); 2085 >(\P EQget_lfalse1) whd in match (get_costlabels_of_trace ????) in ⊢ (???%); @is_permutation_cons 2086 assumption 2087 | #k #i1 #EQcont_st3 #EQcont_st11 2088 cases(stack_safety … (〈cost_act (None NonFunctionalLabel),i〉::k) … EQcont_st3 …) 2089 [2: >EQcont_st12 >EQcont_st11 % ] 2090 * [|#hd1 #tl1] * #EQk1 whd in ⊢ (??%% → ?); #EQlength destruct 2091 %{tl1} % // whd in EQk1 : (??%%); destruct // 2092 ] ] 2093 %2 // [2,4: % //] normalize >EQcode_s1' normalize nodelta % #EQ destruct ] 2094 #h inversion h in ⊢ ?; 2095 [ #H21 #H22 #H23 #H24 #H25 #H26 destruct 2096 | #H28 #H29 #H30 #H31 #H32 #H33 #H34 #H35 #H36 #H37 #H38 #H39 destruct 2097 ] 2098 | #cond #ltrue #i_true #lfalse #i_false #store #EQcode_st11 #EQev_loop #EQcont_st12 #EQ1 #EQ2 destruct 2099 #EQio_11 #EQio_12 #EQ1 #EQ2 #EQ3 #EQ4 destruct #t #Hclass11 #_ #Hpre #IH #s1' #abs_top #abs_tail 2100 whd in ⊢ (% → ?); inversion(check_continuations ?????) [ #_ * ] ** #H1 #abs_top_cont #abs_tail_cont 2101 #EQcheck normalize nodelta **** #HH1 >EQcode_st11 in ⊢ (% → ?); inversion(code ??) 2102 [1,2,3,4,6,7: 2103 [ #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2104 | #r_t #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2105 | #seq #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2106 normalize nodelta [2: * #z1 #z2 cases lbl normalize nodelta 2107 [2: #l1 cases(?==?) normalize nodelta]] 2108 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2109 | #cond #ltrue #i1 #lfalse #i2 #i3 #_ #_ #_ #_ 2110 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2111 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2112 [2,4: * #z3 #z4 >m_return_bind cases(call_post_clean ?????) 2113 [2,4: * #z5 #z6 >m_return_bind cases(?∧?) normalize nodelta ]]] 2114 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2115 | #f #act_p #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2116 normalize nodelta 2117 [2,4: * #z1 #z2 cases lbl normalize nodelta 2118 [2,4: #l1 cases(memb ???) normalize nodelta 2119 [1,3: cases(?==?) normalize nodelta ]]] 2120 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2121 | #lin #io #lout #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2122 normalize nodelta 2123 [2,4: * #z1 #z2 cases(?∧?) normalize nodelta] 2124 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2125 ] 2126 ] #cond1 #ltrue1 #i_true1 #lfalse1 #i_false1 #_ #_ 2127 #EQcode_s1' whd in ⊢ (??%% → ?); inversion(call_post_clean ?????) normalize nodelta 2128 [ #_ #EQ destruct] * #gen_ifalse1 #i_false2 #EQi_false2 inversion(call_post_clean ?????) 2129 [ #_ whd in ⊢ (??%% → ?); #EQ destruct ] * #gen_itrue1 #i_true2 #EQi_true2 >m_return_bind 2130 inversion (?==?) normalize nodelta [2: #_ #EQ destruct] #EQget_ltrue1 2131 inversion(?==?) #EQget_lfalse1 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) 2132 -EQ #EQ destruct(EQ) #EQstore_11 #EQ_info11 #EQ destruct 2133 cases(IH … 2134 (mk_state ? i_true1 2135 (〈cost_act (None ?),LOOP … cond ltrue i_true1 lfalse i_false1〉 :: 2136 cont … s1') (store … st12) (io_info … st12)) abs_tail_cont gen_itrue1) 2137 [2: whd >EQcont_st12 whd in match (check_continuations ?????); 2138 whd in match (foldr2 ???????); 2139 change with (check_continuations ?????) in match (foldr2 ???????); 2140 >EQcheck normalize nodelta whd in match (call_post_clean ?????); 2141 >EQi_false2 normalize nodelta >EQi_true2 >m_return_bind >EQget_ltrue1 2142 >EQget_lfalse1 normalize nodelta % // % // % // % [2: assumption] % // 2143 % // 2144 ] 2145 #abs_cont1 * #s2' * #t' ***** #Hst3_s2' #EQcosts #EQlen #stack_safety 2146 #pre_t' #Hsil %{abs_cont1} %{s2'} %{(t_ind … t')} 2147 [ @hide_prf @(loop_true … EQcode_s1') // |] % [ 2148 % [ % 2149 [ % [2: whd in ⊢ (??%%); @eq_f //] %{Hst3_s2'} 2150 whd in match (get_costlabels_of_trace ????); whd in match(map_labels_on_trace ??); 2151 >(\P EQget_ltrue1) whd in match (get_costlabels_of_trace ????) in ⊢ (???%); 2152 @is_permutation_cons assumption 2153 | #k #i1 #EQcont_st3 #EQcont_st11 2154 cases(stack_safety … (〈cost_act (None NonFunctionalLabel), LOOP p cond ltrue (code p st12) lfalse i_false〉::k) … EQcont_st3 …) 2155 [2: >EQcont_st12 >EQcont_st11 % ] 2156 * [|#hd1 #tl1] * #EQk1 whd in ⊢ (??%% → ?); #EQlength destruct 2157 %{tl1} % // whd in EQk1 : (??%%); destruct // 2158 ] ] 2159 %2 // [2,4: % //] normalize >EQcode_s1' normalize nodelta % #EQ destruct ] 2160 #h inversion h in ⊢ ?; 2161 [ #H41 #H42 #H43 #H44 #H45 #H46 destruct 2162 | #H48 #H49 #H50 #H51 #H52 #H53 #H54 #H55 #H56 #H57 #H58 #H59 destruct 2163 ] 2164 | #cond #ltrue #i_true #lfalse #i_false #mem #EQcode_st11 #EQev_loop #EQcont_st12 #EQ1 #EQstore11 destruct 2165 #EQio_11 #EQio_12 #EQ1 #EQ2 #EQ3 #EQ4 destruct #t #Hclass11 #_ #Hpre #IH #s1' #abs_top #abs_tail 2166 whd in ⊢ (% → ?); inversion(check_continuations ?????) [ #_ * ] ** #H1 #abs_top_cont #abs_tail_cont 2167 #EQcheck normalize nodelta **** #HH1 >EQcode_st11 in ⊢ (% → ?); inversion(code ??) 2168 [1,2,3,4,6,7: 2169 [ #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2170 | #r_t #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2171 | #seq #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2172 normalize nodelta [2: * #z1 #z2 cases lbl normalize nodelta 2173 [2: #l1 cases(?==?) normalize nodelta]] 2174 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2175 | #cond #ltrue #i1 #lfalse #i2 #i3 #_ #_ #_ #_ 2176 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2177 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2178 [2,4: * #z3 #z4 >m_return_bind cases(call_post_clean ?????) 2179 [2,4: * #z5 #z6 >m_return_bind cases(?∧?) normalize nodelta ]]] 2180 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2181 | #f #act_p #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2182 normalize nodelta 2183 [2,4: * #z1 #z2 cases lbl normalize nodelta 2184 [2,4: #l1 cases(memb ???) normalize nodelta 2185 [1,3: cases(?==?) normalize nodelta ]]] 2186 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2187 | #lin #io #lout #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2188 normalize nodelta 2189 [2,4: * #z1 #z2 cases(?∧?) normalize nodelta] 2190 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2191 ] 2192 ] #cond1 #ltrue1 #i_true1 #lfalse1 #i_false1 #_ #_ 2193 #EQcode_s1' whd in ⊢ (??%% → ?); inversion(call_post_clean ?????) normalize nodelta 2194 [ #_ #EQ destruct] * #gen_ifalse1 #i_false2 #EQi_false2 inversion(call_post_clean ?????) 2195 [ #_ whd in ⊢ (??%% → ?); #EQ destruct ] * #gen_itrue1 #i_true2 #EQi_true2 >m_return_bind 2196 inversion (?==?) normalize nodelta [2: #_ #EQ destruct] #EQget_ltrue1 2197 inversion(?==?) #EQget_lfalse1 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) 2198 -EQ #EQ destruct(EQ) #EQstore_11 #EQ_info11 #EQ destruct 2199 cases(IH … 2200 (mk_state ? i_false1 2201 (cont … s1') (store … st12) (io_info … st12)) abs_tail_cont gen_ifalse1) 2202 [2: whd >EQcont_st12 whd in match (check_continuations ?????); 2203 whd in match (foldr2 ???????); 2204 change with (check_continuations ?????) in match (foldr2 ???????); 2205 >EQcheck normalize nodelta whd in match (call_post_clean ?????); 2206 >EQi_false2 normalize nodelta >EQi_true2 2207 % // % // % // % [2: %] // 2208 ] 2209 #abs_cont1 * #s2' * #t' ***** #Hst3_s2' #EQcosts #EQlen #stack_safety 2210 #pre_t' #Hsil %{abs_cont1} %{s2'} %{(t_ind … t')} 2211 [ @hide_prf @(loop_false … EQcode_s1') // |] % [ 2212 % [ % 2213 [ % [2: whd in ⊢ (??%%); @eq_f //] %{Hst3_s2'} 2214 whd in match (get_costlabels_of_trace ????); whd in match(map_labels_on_trace ??); 2215 >(\P EQget_lfalse1) whd in match (get_costlabels_of_trace ????) in ⊢ (???%); 2216 @is_permutation_cons assumption 2217 | #k #i #EQcont_st3 <EQcont_st12 #EQ 2218 cases(stack_safety … EQcont_st3 … EQ) #k1 * #EQk1 #EQlength %{k1} % // 2219 ]] 2220 %2 // [2,4: % //] normalize >EQcode_s1' normalize nodelta % #EQ destruct] 2221 #h inversion h in ⊢ ?; 2222 [ #H61 #H62 #H63 #H64 #H65 #H66 destruct 2223 | #H68 #H69 #H70 #H71 #H72 #H73 #H74 #H75 #H76 #H77 #H78 #H79 destruct 2224 ] 2225 | #lin #io #lout #i #mem #EQcode_st11 #EQev_io #EQcost_st12 #EQcont_st12 2226 #EQ destruct #EQio_12 #EQ1 #EQ2 #EQ3 #EQ4 destruct #t #Hclass_11 #_ #Hpre 2227 #IH #s1' #abs_top #abs_tail whd in ⊢ (% → ?); inversion(check_continuations ?????) 2228 [#_ *] ** #H1 #abs_top_cont #abs_tail_cont #EQcheck normalize nodelta **** 2229 #HH1 >EQcode_st11 in ⊢ (% → ?); inversion(code ??) 2230 [1,2,3,4,5,6: 2231 [ #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2232 | #r_t #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2233 | #seq #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2234 normalize nodelta [2: * #z1 #z2 cases lbl normalize nodelta 2235 [2: #l1 cases(?==?) normalize nodelta]] 2236 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2237 | #cond #ltrue #i1 #lfalse #i2 #i3 #_ #_ #_ #_ 2238 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2239 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2240 [2,4: * #z3 #z4 >m_return_bind cases(call_post_clean ?????) 2241 [2,4: * #z5 #z6 >m_return_bind cases(?∧?) normalize nodelta ]]] 2242 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2243 | #cond #ltrue #i1 #lfalse #i2 #_ #_ #_ 2244 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2245 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2246 [2,4: * #z3 #z4 >m_return_bind cases(?∧?) normalize nodelta ]] 2247 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2248 | #f #act_p #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2249 normalize nodelta 2250 [2,4: * #z1 #z2 cases lbl normalize nodelta 2251 [2,4: #l1 cases(memb ???) normalize nodelta 2252 [1,3: cases(?==?) normalize nodelta ]]] 2253 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2254 ] 2255 ] #lin1 #io1 #lout1 #i1 #_ #EQcode_s1' 2256 whd in ⊢ (??%% → ?); inversion(call_post_clean ?????) normalize nodelta 2257 [ #_ #EQ destruct] * #gen_lab #i2 #EQ_i2 inversion(?==?) normalize nodelta 2258 [2: #_ #EQ destruct] #EQget_lout1 inversion(?==?) #EQget_lin1 normalize nodelta 2259 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2260 #EQstore11 #EQinfo11 #EQ destruct 2261 cases(IH … 2262 (mk_state ? (EMPTY p) (〈cost_act (Some ? lout),i1〉::cont … s1') (store … st12) true) 2263 abs_tail_cont [ ]) 2264 [2: whd >EQcont_st12 whd in match (check_continuations ?????); 2265 whd in match (foldr2 ???????); 2266 change with (check_continuations ?????) in match (foldr2 ???????); 2267 >EQcheck normalize nodelta >EQ_i2 normalize nodelta % // % // % // % 2268 [2: >EQcost_st12 % ] % [ % // % //] >(\P EQget_lout1) % 2269 ] 2270 #abs_cont1 * #s2' * #t' ***** #Hst3_s2' #EQcosts #EQlen #stack_safety 2271 #pre_t' #Hsil %{abs_cont1} %{s2'} %{(t_ind … t')} 2272 [ @hide_prf @(io_in … EQcode_s1') // |] % [ 2273 % [ % 2274 [ % [2: whd in ⊢ (??%%); @eq_f //] %{Hst3_s2'} 2275 whd in match (get_costlabels_of_trace ????); whd in match(map_labels_on_trace ??); 2276 >(\P EQget_lin1) whd in match (get_costlabels_of_trace ????) in ⊢ (???%); 2277 @is_permutation_cons assumption 2278 | #k #i1 #EQcont_st3 #EQcont_st11 2279 cases(stack_safety … (〈cost_act (Some ? lout),i〉::k) … EQcont_st3 …) 2280 [2: >EQcont_st12 >EQcont_st11 %] 2281 * [|#hd1 #tl1] * #EQk1 whd in ⊢ (??%% → ?); #EQlength destruct 2282 %{tl1} % // whd in EQk1 : (??%%); destruct // 2283 ]] 2284 %2 // [2,4: % //] normalize >EQcode_s1' normalize nodelta % #EQ destruct ] 2285 #h inversion h in ⊢ ?; 2286 [ #H81 #H82 #H83 #H84 #H85 #H86 destruct 2287 | #H88 #H89 #H90 #H91 #H92 #H93 #H94 #H95 #H96 #H97 #H98 #H99 destruct 2288 ] 2289 | #f #act_p #r_lb #cd #mem #env_it #EQ1 #EQ2 #EQ3 #EQ4 #EQ5 #EQ6 #EQ7 #EQ8 #EQ9 #EQ10 2290 #EQ11 #EQ12 destruct #tl #_ * #x #EQ destruct(EQ) 2291 | #r_t #mem #con #rb #i #EQ1 #EQ2 #EQ3 #EQ4 #EQ5 #EQ6 #EQ7 #EQ8 #EQ9 #EQ10 #EQ11 #EQ12 2292 destruct #t #_ * #nf #EQ destruct 2293 ] 2294 | #st1 #st2 #st3 #lab #st1_noio #H inversion H in ⊢ ?; 2295 [1,2,3,4,5,6,7,8: 2296 [ #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 2297 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 2298 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 #x20 2299 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 #x20 2300 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 2301 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 2302 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x18 2303 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x18 #x19 #x20 #x21 2304 ] 2305 destruct #tl * #y #EQ destruct(EQ) @⊥ >EQ in x12; * #ABS @ABS % 2306 ] 2307 #st11 #st12 #r_t #mem #cont * [|#x] #i #EQcode_st11 #EQcont_st11 #EQ destruct(EQ) 2308 #EQio_11 #EQio_12 #EQev_ret #EQ1 #EQ2 #EQ3 #EQ4 #EQ5 #EQ6 destruct(EQ1 EQ2 EQ3 EQ4 EQ5 EQ6) 2309 #t * #lbl #EQ destruct(EQ) #pre_t #IH #s1' #abs_tail #abs_top whd in ⊢ (% → ?); 2310 inversion(check_continuations …) [#_ *] ** #H1 #abs_top_cont #abs_tail_cont 2311 >EQcont_st11 in ⊢ (% → ?); inversion(cont … s1') [ #_ whd in ⊢ (??%% → ?); #EQ destruct(EQ)] 2312 * #act_lab #i #cont_s1'' #_ #EQcont_s1' whd in ⊢ (??%% → ?); 2313 change with (check_continuations ?????) in match (foldr2 ???????); 2314 inversion(check_continuations ?????) normalize nodelta [ #_ whd in ⊢ (??%% → ?); #EQ destruct] 2315 ** #H2 #abs_top_cont' #abs_tail_cont' #EQcheck normalize nodelta 2316 inversion(call_post_clean ?????) normalize nodelta [ #_ whd in ⊢ (??%% → ?); #EQ destruct *****] 2317 * #gen_labs #i' #EQi' inversion act_lab normalize nodelta 2318 [ #f #lab | * [| #lbl1 ]| * [| #nf_l]] #EQ destruct(EQ) 2319 [1,6: whd in ⊢ (??%% → ?); #EQ destruct ***** 2320 |4,5: normalize nodelta whd in ⊢ (??%% → ?); #EQ destruct(EQ) ****** [2: *] #EQ destruct 2321 ] 2322 whd in match ret_costed_abs; normalize nodelta 2323 [ whd in ⊢ (??%% → ?); #EQ destruct(EQ) ****** #EQ destruct(EQ) ] 2324 inversion (memb ???) normalize nodelta #Hlbl1 whd in ⊢ (??%% → ?); #EQ destruct(EQ) 2325 ****** [ * ] #EQ destruct(EQ) #HH2 #EQ destruct #EQget_el >EQcode_st11 in ⊢ (% → ?); 2326 inversion(code … s1') 2327 [1,3,4,5,6,7: 2328 [ #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2329 | #seq #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2330 normalize nodelta [2: * #z1 #z2 cases lbl normalize nodelta 2331 [2: #l1 cases(?==?) normalize nodelta]] 2332 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2333 | #cond #ltrue #i1 #lfalse #i2 #i3 #_ #_ #_ #_ 2334 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2335 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2336 [2,4: * #z3 #z4 >m_return_bind cases(call_post_clean ?????) 2337 [2,4: * #z5 #z6 >m_return_bind cases(?∧?) normalize nodelta ]]] 2338 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2339 | #cond #ltrue #i1 #lfalse #i2 #_ #_ #_ 2340 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2341 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2342 [2,4: * #z3 #z4 >m_return_bind cases(?∧?) normalize nodelta ]] 2343 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2344 | #f #act_p #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2345 normalize nodelta 2346 [2,4: * #z1 #z2 cases lbl normalize nodelta 2347 [2,4: #l1 cases(memb ???) normalize nodelta 2348 [1,3: cases(?==?) normalize nodelta ]]] 2349 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2350 | #lin #io #lout #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2351 normalize nodelta 2352 [2,4: * #z1 #z2 cases(?∧?) normalize nodelta] 2353 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2354 ] 2355 ] 2356 #r_t' #EQcode_s1' 2357 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2358 #EQstore11 #EQio_111 #EQ destruct 2359 cases(IH … (mk_state ? i cont_s1'' (store … st12) (io_info … st12)) abs_tail_cont gen_labs) 2360 [2: whd >EQcheck normalize nodelta % // % // % // % // @EQi' ] 2361 #abs_top1 * #s2' * #t' ***** #Hst3_s2' #EQcosts #EQlen #stack_safety 2362 #pre_t' #Hsil %{abs_top1} %{s2'} %{(t_ind … t')} 2363 [ @hide_prf @(ret_instr … EQcode_s1' … EQcont_s1') // 2364 | ] % [ 2365 % [ % 2366 [ % [2: whd in ⊢ (??%%); @eq_f // ] %{Hst3_s2'} 2367 whd in match (get_costlabels_of_trace ????); 2368 whd in match (get_costlabels_of_trace ????) in ⊢ (???%); 2369 whd in match (map_labels_on_trace ??); >EQget_el @is_permutation_cons 2370 assumption 2371 | * [| #hd #tl] #i1 #EQcont_st3 >EQcont_st11 #EQ whd in EQ : (???%); 2372 [ <EQ in EQcont_st3; whd in ⊢ (??%% → ?); #EQ1 destruct ] 2373 destruct cases(stack_safety … EQcont_st3 … e0) #k1 * 2374 #EQk1 #EQlength %{(〈ret_act (Some ? lbl1),i〉::k1)} 2375 % [ >EQk1 % | whd in ⊢ (??%%); @eq_f // ] 2376 ]] 2377 %3 // [2,4: % //] normalize >EQcode_s1' normalize nodelta % #EQ destruct ] 2378 #h inversion h in ⊢ ?; 2379 [ #H101 #H102 #H103 #H104 #H105 #H106 destruct 2380 | #H108 #H109 #H110 #H111 #H112 #H113 #H114 #H115 #H116 #H117 #H118 #H119 destruct 2381 ] 2382 | #st1 #st2 #st3 #lab #H inversion H in ⊢ ?; #st11 #st12 2383 [1,2,3,4,5,6,7,9: 2384 [ #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 2385 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 2386 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 #x20 2387 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 #x20 2388 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 2389 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 2390 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x18 2391 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x18 #x19 #x20 2392 ] 2393 destruct * #z1 * #z2 #EQ destruct(EQ) whd in ⊢ (?% → ?); >x3 in ⊢ (% → ?); * 2394 | #f #act_p #opt_l #i #mem #env_it #EQcode11 #EQenv_it #EQmem 2395 #EQ destruct(EQ) #EQcode12 #EQcont12 #EQio11 #EQio12 #EQ1 #EQ2 #EQ3 2396 destruct(EQ1 EQ2 EQ3) #EQ destruct(EQ) #tl #_ #_ whd in ⊢ (?% → ?); >EQcode11 in ⊢ (% → ?); 2397 normalize nodelta cases opt_l in EQcode11 EQcont12; normalize nodelta [ #x #y *] 2398 #lbl #EQcode11 #EQcont12 * #pre_tl #IH #st1' #abs_top #abs_tail 2399 whd in ⊢ (% → ?); inversion(check_continuations ????) [ #_ *] ** #H1 2400 #abs_top_cont #abs_tail_cont #EQcheck 2401 normalize nodelta **** #HH1 >EQcode11 in ⊢ (% → ?); inversion(code … st1') 2402 [1,2,3,4,5,7: 2403 [ #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2404 | #r_t #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2405 | #seq #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2406 normalize nodelta [2: * #z1 #z2 cases lbl normalize nodelta 2407 [2: #l1 cases(?==?) normalize nodelta]] 2408 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2409 | #cond #ltrue #i1 #lfalse #i2 #i3 #_ #_ #_ #_ 2410 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2411 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2412 [2,4: * #z3 #z4 >m_return_bind cases(call_post_clean ?????) 2413 [2,4: * #z5 #z6 >m_return_bind cases(?∧?) normalize nodelta ]]] 2414 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2415 | #cond #ltrue #i1 #lfalse #i2 #_ #_ #_ 2416 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2417 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2418 [2,4: * #z3 #z4 >m_return_bind cases(?∧?) normalize nodelta ]] 2419 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2420 | #lin #io #lout #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2421 normalize nodelta 2422 [2,4: * #z1 #z2 cases(?∧?) normalize nodelta] 2423 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2424 ] 2425 ] #f' #act_p' #opt_l' #i' #_ 2426 #EQcode_st1' #EQclean #EQstore #EQio #EQ destruct(EQ) 2427 lapply(trans_env_ok … no_dup) >EQtrans normalize nodelta #H 2428 cases(H … EQenv_it) -H #env_it' * #fresh' ** #is_fresh_fresh' 2429 #EQenv_it' ***** #EQtrans 2430 #EQgen_labels #EQsignature #EQlab_env_it #same_map #same_keep 2431 change with (m_bind ?????) in EQclean : (??%?); inversion(call_post_clean ?????) in EQclean; 2432 [ #_ whd in ⊢ (??%% → ?); #EQ destruct] * #abs_top'' #i'' #EQi' >m_return_bind 2433 inversion opt_l' in EQcode_st1'; [| #lbl'] #EQopt_l' #EQcode_st1' normalize nodelta 2434 [2: inversion(memb ???) normalize nodelta #Hlbl_keep 2435 [ inversion (get_element ????) normalize nodelta [ #_ whd in ⊢ (??%% → ?); #EQ destruct] 2436 #lbl'' #l3' #_ #EQget_el whd in match (eqb ???); inversion (eqb ???) normalize nodelta 2437 [2: #_ whd in ⊢ (??%% → ?); #EQ destruct] #H cases(eqb_true ? lbl'' lbl') #H1 #_ 2438 lapply(H1 H) -H -H1 #EQ destruct(EQ) inversion(eqb_list ???) normalize nodelta 2439 [2: #_ whd in ⊢ (??%% → ?); #EQ destruct] #H cases(eqb_true (DeqSet_List ?) l3' abs_top'') 2440 #H1 #_ lapply(H1 H) -H -H1 #EQ destruct(EQ) whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ 2441 #EQ destruct(EQ) 2442 cases(IH (mk_state ? (f_body … env_it') (〈ret_act opt_l',i'〉 :: (cont … st1')) (store ? st12) false) (abs_tail_cont) (gen_labels … (call_post_trans … (f_body … env_it) fresh' (nil ?)))) 2443 [2: whd >EQcont12 2444 change with (m_bind ??? (check_continuations ?????) ?) in match (check_continuations ?????); 2445 >EQcheck >m_return_bind normalize nodelta >EQi' normalize nodelta >EQopt_l' 2446 whd in match ret_costed_abs; normalize nodelta >Hlbl_keep normalize nodelta 2447 % // % // % // % [/5 by conj/] >EQgen_labels >EQcode12 <EQtrans 2448 @(inverse_call_post_trans … is_fresh_fresh') 2449 [2: % |*: [2,3: /2 by / ] 2450 cases(lookup_ok_append … EQenv_it) #env1 * #env2 * #EQ1 #EQ2 2451 whd in no_dup; destruct >EQ1 in no_dup; >foldr_map_append >foldr_map_append 2452 #no_dup lapply(no_duplicates_append_r … no_dup) #H1 2453 lapply(no_duplicates_append_l … H1) whd in match (foldr ?????); -H1 2454 change with ([?]@?) in match (?::?); #H1 2455 lapply(no_duplicates_append_r … H1) >append_nil // 2456 ] 2457 ] 2458 #abs_top''' * #st3' * #t' ***** #Hst3st3' #EQcosts #EQlen #stack_safety 2459 #pre_t' #Hsil %{abs_top'''} 2460 %{st3'} %{(t_ind … (call_act f (f_lab … env_it')) … t')} 2461 [ @hide_prf @call /2 width=10 by jmeq_to_eq/ ] % [ 2462 % [ % 2463 [ % [2: whd in ⊢ (??%%); >EQlen %] 2464 %{Hst3st3'} >map_labels_on_trace_append whd in match (get_costlabels_of_trace ????) in ⊢ (???%); 2465 >EQlab_env_it >associative_append whd in match (append ???); >associative_append 2466 >associative_append in ⊢ (???%); >(associative_append … [?]) in ⊢ (???%); 2467 whd in match (map_labels_on_trace ??); >EQgen_labels @is_permutation_cons 2468 >append_nil whd in match (append ???) in ⊢ (???%); // 2469 | #k #i1 #EQcont_st3 #EQcont_st11 2470 cases(stack_safety … (〈ret_act (Some ? lbl),i〉::k) … EQcont_st3 …) 2471 [2: >EQcont12 >EQcont_st11 % ] 2472 * [|#hd1 #tl1] * #EQk1 whd in ⊢ (??%% → ?); #EQlength destruct 2473 %{tl1} % // whd in EQk1 : (??%%); destruct // 2474 ]] 2475 %4 // [2,4: % [2: % // |]] normalize >EQcode_st1' normalize nodelta % #EQ destruct 2476 ] 2477 #h inversion h in ⊢ ?; 2478 [ #H121 #H122 #H123 #H124 #H125 #H126 destruct 2479 | #H128 #H129 #H130 #H131 #H132 #H133 #H134 #H135 #H136 #H137 #H138 #H139 destruct 2480 ] 2481 | whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2482 ] 2483 | whd in ⊢ (??%% → ?); #EQ destruct(EQ) 2484 ] 2485 ] 2486 | #st1 #st2 #st3 #st4 #st5 #l1 #l2 #H inversion H in ⊢ ?; 2487 [1,2,3,4,5,6,7,9: 2488 [ #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 2489 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 2490 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 #x20 2491 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 #x20 2492 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 2493 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 2494 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x18 2495 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x18 #x19 #x20 2496 ] 2497 destruct #t1 #t2 #prf #_ #_ * #y1 * #y2 #EQ destruct(EQ) @⊥ >EQ in x12; * #H @H % 2498 ] 2499 #st11 #st12 #f #act_p #opt_l #i #mem #env_it #EQcode11 #EQenv_it #EQmem 2500 #EQ destruct(EQ) #EQcode12 #EQcont12 #EQio11 #EQio12 #EQ1 #EQ2 #EQ3 2501 destruct(EQ1 EQ2 EQ3) #EQ destruct(EQ) #t1 #t2 #H inversion H in ⊢ ?; 2502 [1,2,3,4,5,6,7,8: 2503 [ #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 2504 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 2505 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 #x20 2506 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 #x20 2507 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 2508 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x17 #x18 #x19 2509 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x18 2510 | #x1 #x2 #x3 #x4 #x5 #x6 #x7 #x8 #x9 #x10 #x11 #x12 #x13 #x14 #x15 #x16 #x18 #x19 #x20 #x21 2511 ] 2512 destruct #_ #_ #_ #_ #_ #_ #_ whd in ⊢ (% → ?); #EQ destruct(EQ) 2513 @⊥ >EQ in x12; * #H @H % 2514 ] 2515 #st41 #st42 #r_t #mem1 #new_cont #opt_l1 #cd #EQcode41 #EQcont41 #EQ destruct(EQ) 2516 #EQio41 #EQio42 #EQmem1 #EQ1 #EQ2 #EQ3 #EQ4 #EQ5 #EQ6 destruct #st11_noio 2517 #st41_noio #_ whd in ⊢ (?(?%) → ?); >EQcode11 in ⊢ (% → ?); normalize nodelta 2518 inversion(opt_l) in ⊢(% → ?); normalize nodelta [2: #lbl #_ * ] #EQopt_l destruct(EQopt_l) 2519 #_ #Hpre_t1 #Hpre_t2 whd in ⊢ (% → ?); * #EQcont11_42 >EQcode11 in ⊢ (% → ?); 2520 normalize nodelta #EQ destruct(EQ) whd in ⊢ (% → ?); 2521 #EQ destruct #IH1 #IH2 #st1' #abs_top #abs_tail 2522 whd in ⊢ (% → ?); inversion(check_continuations ????) [ #_ *] ** #H1 2523 #abs_top_cont #abs_tail_cont #EQcheck 2524 normalize nodelta **** #HH1 >EQcode11 in ⊢ (% → ?); inversion(code … st1') 2525 [1,2,3,4,5,7: 2526 [ #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2527 | #r_t #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2528 | #seq #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2529 normalize nodelta [2: * #z1 #z2 cases lbl normalize nodelta 2530 [2: #l1 cases(?==?) normalize nodelta]] 2531 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2532 | #cond #ltrue #i1 #lfalse #i2 #i3 #_ #_ #_ #_ 2533 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2534 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2535 [2,4: * #z3 #z4 >m_return_bind cases(call_post_clean ?????) 2536 [2,4: * #z5 #z6 >m_return_bind cases(?∧?) normalize nodelta ]]] 2537 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2538 | #cond #ltrue #i1 #lfalse #i2 #_ #_ #_ 2539 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2540 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2541 [2,4: * #z3 #z4 >m_return_bind cases(?∧?) normalize nodelta ]] 2542 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2543 | #lin #io #lout #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2544 normalize nodelta 2545 [2,4: * #z1 #z2 cases(?∧?) normalize nodelta] 2546 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2547 ] 2548 ] #f' #act_p' #opt_l' #i' #_ 2549 #EQcode_st1' #EQclean #EQstore #EQio #EQ destruct(EQ) 2550 lapply(trans_env_ok … no_dup) >EQtrans normalize nodelta #H 2551 cases(H … EQenv_it) -H #env_it' * #fresh' ** #is_fresh_fresh' #EQenv_it' ***** #EQtrans 2552 #EQgen_labels #EQsignature #EQlab_env_it #same_map #same_keep 2553 change with (m_bind ?????) in EQclean : (??%?); inversion(call_post_clean ?????) in EQclean; 2554 [ #_ whd in ⊢ (??%% → ?); #EQ destruct] * #abs_top'' #i'' #EQi' >m_return_bind 2555 inversion opt_l' in EQcode_st1'; [| #lbl'] #EQopt_l' #EQcode_st1' normalize nodelta 2556 [ whd in ⊢ (??%% → ?); #EQ destruct(EQ) ] destruct(EQopt_l') 2557 inversion(memb ???) normalize nodelta #Hlbl_keep' 2558 [ cases(?==?) normalize nodelta ] 2559 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct 2560 cases(IH1 2561 (mk_state ? 2562 (f_body … env_it') 2563 (〈ret_act (Some ? lbl'),i'〉 :: (cont … st1')) 2564 (store ? st12) false) 2565 (abs_top''@abs_tail_cont) 2566 (gen_labels … (call_post_trans … (f_body … env_it) fresh' (nil ?)))) 2567 [2: whd >EQcont12 2568 change with (m_bind ??? (check_continuations ?????) ?) in match (check_continuations ?????); 2569 >EQcheck >m_return_bind normalize nodelta >EQi' normalize nodelta 2570 whd in match ret_costed_abs; normalize nodelta >Hlbl_keep' normalize nodelta 2571 % // % // % // % [/5 by conj/] >EQcode12 <EQtrans 2572 @(inverse_call_post_trans … fresh') 2573 [2: % |*: [2,3: /2 by / ] 2574 cases(lookup_ok_append … EQenv_it) #env1 * #env2 * #EQ1 #EQ2 2575 whd in no_dup; destruct >EQ1 in no_dup; >foldr_map_append >foldr_map_append 2576 #no_dup lapply(no_duplicates_append_r … no_dup) #H1 2577 lapply(no_duplicates_append_l … H1) whd in match (foldr ?????); -H1 2578 change with ([?]@?) in match (?::?); #H1 2579 lapply(no_duplicates_append_r … H1) >append_nil // 2580 ] 2581 ] 2582 #abs_top''' * #st41' * #t1' ***** #Hst41st41' #EQcosts #EQlen #stack_safety1 2583 #pre_t1' #Hsil1 2584 whd in Hst41st41'; inversion(check_continuations …) in Hst41st41'; [ #_ * ] 2585 ** #H2 #abs_top_cont' #abs_tail_cont' >EQcont41 in ⊢ (% → ?); 2586 whd in ⊢ (??%? → ?); inversion(cont … st41') normalize nodelta 2587 [ #_ #EQ destruct(EQ) ] * #act_lbl #i'' #cont_st42' #_ #EQcont_st41' 2588 change with (check_continuations ?????) in match (foldr2 ???????); 2589 inversion(check_continuations ?????) [ #_ whd in ⊢ (??%% → ?); #EQ destruct] 2590 ** #H3 #abs_top_cont'' #abs_tail_cont'' #EQ_contst42 >m_return_bind 2591 normalize nodelta inversion(call_post_clean ?????) normalize nodelta 2592 [ #_ whd in ⊢ (??%% → ?); #EQ destruct ***** ] * #abs_top'''' #i''' 2593 #EQ_clean_i'' inversion(act_lbl) normalize nodelta 2594 [1,3,4: 2595 [ #x1 #x2 #EQ whd in ⊢ (??%% → ?); #EQ1 destruct ***** 2596 | * [|#x1] #EQ destruct normalize nodelta whd in ⊢ (??%% → ?); 2597 #EQ destruct ****** [2: *] #EQ destruct 2598 ] 2599 ] 2600 cut(act_lbl = ret_act (Some ? lbl') ∧ cont_st42' = cont … st1' ∧ i'' = i') 2601 [ cases(stack_safety1 [ ] …) 2602 [3: >EQcont41 in ⊢ (??%?); % |4: normalize // |5: % |2:] * [|#x #xs] * 2603 whd in ⊢ (??%% → ??%% → ?); #EQ1 #EQ2 destruct > EQcont_st41' in EQ1; 2604 #EQ destruct(EQ) /3 by conj/ 2605 ] 2606 ** #EQ1 #EQ2 #EQ3 destruct #x #EQ destruct whd in match ret_costed_abs; normalize nodelta 2607 >Hlbl_keep' normalize nodelta 2608 whd in ⊢ (??%% → ?); #EQ destruct ****** #_ #HH3 #EQ destruct(EQ) 2609 >EQcode41 in ⊢ (% → ?); inversion(code … st41') 2610 [1,3,4,5,6,7: 2611 [ #_ whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2612 | #seq #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2613 normalize nodelta [2: * #z1 #z2 cases lbl normalize nodelta 2614 [2: #l1 cases(?==?) normalize nodelta]] 2615 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2616 | #cond #ltrue #i1 #lfalse #i2 #i3 #_ #_ #_ #_ 2617 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2618 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2619 [2,4: * #z3 #z4 >m_return_bind cases(call_post_clean ?????) 2620 [2,4: * #z5 #z6 >m_return_bind cases(?∧?) normalize nodelta ]]] 2621 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2622 | #cond #ltrue #i1 #lfalse #i2 #_ #_ #_ 2623 whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2624 [2,4: * #z1 #z2 normalize nodelta cases (call_post_clean ?????) 2625 [2,4: * #z3 #z4 >m_return_bind cases(?∧?) normalize nodelta ]] 2626 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2627 | #f #act_p #lbl #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2628 normalize nodelta 2629 [2,4: * #z1 #z2 cases lbl normalize nodelta 2630 [2,4: #l1 cases(memb ???) normalize nodelta 2631 [1,3: cases(?==?) normalize nodelta ]]] 2632 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2633 | #lin #io #lout #i #_ #_ whd in ⊢ (??%% → ?); cases(call_post_clean ?????) 2634 normalize nodelta 2635 [2,4: * #z1 #z2 cases(?∧?) normalize nodelta] 2636 whd in ⊢ (??%% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) 2637 ] 2638 ] #r_t' #EQcode_st41' whd in ⊢ (??%% → ?); 2639 #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct(EQ) #EQstore_st41' #EQinfo_st41' 2640 #EQ1 >EQcont11_42 in EQcheck; >EQ_contst42 in ⊢ (% → ?); 2641 #EQ destruct(EQ) >EQi' in EQ_clean_i''; #EQ destruct(EQ) 2642 cases(IH2 2643 (mk_state ? i' (cont … st1') (store … st42) (io_info … st42)) 2644 abs_tail_cont abs_top'''') 2645 [2: whd >EQ_contst42 normalize nodelta % // % // % // % // @EQi' ] 2646 #abs_top_1 * #st5' * #t2' ***** #Hst5_st5' #EQcosts' 2647 #EQlen' #stack_safety2 #pre_t2' #Hsil2 %{abs_top_1} %{st5'} %{(t_ind … (t1' @ (t_ind … t2')))} 2648 [3: @hide_prf @(call … EQcode_st1' … EQenv_it') // 2649 |1: @hide_prf @(ret_instr … EQcode_st41' … EQcont_st41') // 2650 |2,4: skip 2651 ] % [ 2652 % [ % 2653 [ % [2: whd in ⊢ (??%%); @eq_f >len_append >len_append @eq_f2 // 2654 whd in ⊢ (??%%); @eq_f // ] 2655 %{Hst5_st5'} whd in match (map_labels_on_trace ??); 2656 change with (map_labels_on_trace ??) in match (foldr ?????); 2657 >map_labels_on_trace_append >get_cost_label_append 2658 whd in match (get_costlabels_of_trace ??? (t_ind …)); 2659 whd in match (get_costlabels_of_trace ??? (t_ind …)); 2660 >get_cost_label_append 2661 whd in match (get_costlabels_of_trace ??? (t_ind …)); 2662 >map_labels_on_trace_append >EQlab_env_it >EQgen_labels 2663 whd in match (map_labels_on_trace ? [ ]); 2664 whd in match (append ? (nil ?) ?); 2665 cut (∀A.∀l: list A. [ ] @ l = l) [//] #nil_append 2666 >nil_append >nil_append >nil_append @(permute_ok … EQcosts EQcosts') 2667 | #k #i #EQcont_st3 #EQ 2668 cases(stack_safety2 … EQcont_st3 … EQ) #k1 * #EQk1 #EQlength %{k1} % // 2669 ]] 2670 %4 2671 [ normalize >EQcode_st1' normalize nodelta % #EQ destruct 2672 | %{f} % // 2673 | whd in ⊢ (?%); >EQcode_st1' normalize nodelta @I 2674 | @append_premeasurable // %3 // 2675 [ whd in match (as_classify ??); >EQcode_st41' normalize nodelta % #EQ destruct 2676 | whd %{lbl'} % 2677 ] 2678 ] 2679 ] #h inversion h in ⊢ ?; 2680 [ #H141 #H142 #H143 #H144 #H145 #H146 destruct 2681 |#H148 #H149 #H150 #H151 #H152 #H153 #H154 #H155 #H156 #H157 #H158 #H159 destruct 2682 ] 2683 ] 2684 qed. 2685 2686 2687 lemma silent_in_silent : ∀p,p',prog. 2688 no_duplicates_labels … prog → 2689 let 〈t_prog,m,keep〉 ≝ trans_prog … prog in 2690 ∀s1,s2,s1' : state p.∀abs_top,abs_tail. 2691 ∀t : raw_trace (operational_semantics … p' prog) s1 s2. 2692 silent_trace … t → 2693 state_rel … m keep abs_top abs_tail s1 s1' → 2694 ∃abs_top'.∃s2'.∃t' : raw_trace (operational_semantics … p' t_prog) s1' s2'. 2695 state_rel … m keep abs_top' abs_tail s2 s2' ∧ 2696 silent_trace … t'. 2697 #p #p' #prog #no_dup @pair_elim * #t_prog #m #keep #EQt_prog normalize nodelta 2698 #s1 #s2 #s1' #abs_top #abs_tail #t * 2699 [2: cases t -s1 -s2 [2: #H81 #H82 #H83 #H84 #H85 #H86 #H87 #H88 cases H87 #H cases(H I)] 2700 #st #_ #H %{abs_top} %{s1'} %{(t_base …)} % // %2 % * ] 2701 #H #H1 lapply(correctness_lemma p p' prog no_dup) >EQt_prog normalize nodelta #H2 2702 cases(H2 … (pre_silent_is_premeasurable … H) … H1) -H2 #abs_top' * #s2' * #t' ***** #REL' 2703 #_ #_ #_ #_ #H2 %{abs_top'} %{s2'} %{t'} %{REL'} % @H2 // 2704 qed. 1378 -
LTS/Simulation.ma
r3531 r3549 4 4 discriminator option. 5 5 6 record relations ( S1,S2 : abstract_status) : Type[0] ≝6 record relations (p : label_params) (S1,S2 : abstract_status p) : Type[0] ≝ 7 7 { Srel: S1 → S2 → Prop 8 8 ; Crel: S1 → S2 → Prop 9 9 }. 10 10 11 definition Rrel ≝ λS1,S2 : abstract_status.λrel : relations S1 S2.λs,s'. 12 ∀s1,s1'. as_syntactical_succ S1 s1 s → Crel … rel s1 s1' → as_syntactical_succ S2 s1' s'. 13 14 record simulation_conditions (S1,S2 : abstract_status) (rel : relations S1 S2) : Prop ≝ 11 definition Rrel ≝ λp.λS1,S2 : abstract_status p.λrel : relations … S1 S2.λs,s'. 12 ∀s1,s1'. as_syntactical_succ … S1 s1 s → Crel … rel s1 s1' → as_syntactical_succ … S2 s1' s'. 13 14 record simulation_conditions (p : label_params) (S1,S2 : abstract_status p) 15 (rel : relations …S1 S2) : Prop ≝ 15 16 { initial_is_initial : 16 17 ∀s1,s2.Srel … rel s1 s2 → (bool_to_Prop (as_initial … s1) → bool_to_Prop (as_initial … s2)) … … 21 22 ; simulate_tau: 22 23 ∀s1:S1.∀s2:S2.∀s1':S1. 23 as_execute … (cost_act (None ?)) s1 s1'→24 as_execute … (cost_act … (None ?)) s1 s1'→ 24 25 Srel … rel s1 s2 → as_classify … s1 ≠ cl_io → as_classify … s1' ≠ cl_io → 25 ∃s2'. ∃t: raw_trace S2 s2 s2'.26 ∃s2'. ∃t: raw_trace … S2 s2 s2'. 26 27 Srel … rel s1' s2' ∧ silent_trace … t 27 28 ; simulate_label: 28 29 ∀s1:S1.∀s2:S2.∀l.∀s1':S1. 29 as_execute S1 (cost_act(Some ? l)) s1 s1' →30 as_execute … S1 (cost_act … (Some ? l)) s1 s1' → 30 31 as_classify … s1 ≠ cl_io → as_classify … s1' ≠ cl_io → 31 32 Srel … rel s1 s2 → 32 33 ∃s2',s2'',s2'''. 33 ∃t1: raw_trace S2 s2 s2'.34 as_execute … (cost_act (Some ? l)) s2' s2'' ∧34 ∃t1: raw_trace … S2 s2 s2'. 35 as_execute … (cost_act … (Some ? l)) s2' s2'' ∧ 35 36 ∃t3: raw_trace … s2'' s2'''. 36 37 Srel … rel s1' s2''' ∧ silent_trace … t1 ∧ pre_silent_trace … t3 … … 38 39 ∀s1:S1.∀s2:S2.∀s1':S1.∀f,l. 39 40 is_call_post_label … s1 → 40 as_execute … (call_act f l) s1 s1' →41 as_execute … (call_act … f l) s1 s1' → 41 42 as_classify … s1 ≠ cl_io → as_classify … s1' ≠ cl_io → 42 43 Srel … rel s1 s2 → 43 44 ∃s2',s2'',s2'''. 44 ∃t1: raw_trace S2 s2 s2'.45 as_execute … (call_act f l) s2' s2'' ∧45 ∃t1: raw_trace … S2 s2 s2'. 46 as_execute … (call_act … f l) s2' s2'' ∧ 46 47 bool_to_Prop(is_call_post_label … s2') ∧ 47 ∃t3: raw_trace S2 s2'' s2'''.48 ∃t3: raw_trace … S2 s2'' s2'''. 48 49 Srel … rel s1' s2''' ∧ silent_trace … t1 ∧ pre_silent_trace … t3 49 50 ; simulate_ret_l: 50 51 ∀s1:S1.∀s2:S2.∀s1':S1.∀l. 51 as_execute S1 (ret_act(Some ? l)) s1 s1' →52 as_execute … S1 (ret_act … (Some ? l)) s1 s1' → 52 53 as_classify … s1 ≠ cl_io → as_classify … s1' ≠ cl_io → 53 54 Srel … rel s1 s2 → 54 55 ∃s2',s2'',s2'''. 55 ∃t1: raw_trace S2 s2 s2'.56 as_execute … (ret_act(Some ? l)) s2' s2'' ∧57 ∃t3: raw_trace S2 s2'' s2'''.56 ∃t1: raw_trace … S2 s2 s2'. 57 as_execute p S2 (ret_act p (Some ? l)) s2' s2'' ∧ 58 ∃t3: raw_trace … S2 s2'' s2'''. 58 59 Srel … rel s1' s2''' ∧ silent_trace … t1 ∧ pre_silent_trace … t3 59 60 ; simulate_call_n: 60 61 ∀s1:S1.∀s2:S2.∀s1':S1.∀f,l. 61 as_execute … (call_act f l) s1 s1' →62 as_execute … (call_act … f l) s1 s1' → 62 63 ¬ is_call_post_label … s1 → 63 64 as_classify … s1 ≠ cl_io → as_classify … s1' ≠ cl_io → 64 65 Srel … rel s1 s2 → 65 66 ∃s2',s2'',s2'''. 66 ∃t1: raw_trace S2 s2 s2'.67 ∃t1: raw_trace … S2 s2 s2'. 67 68 bool_to_Prop (¬ is_call_post_label … s2') ∧ 68 as_execute … (call_act f l) s2' s2'' ∧69 ∃t3: raw_trace S2 s2'' s2'''.69 as_execute … (call_act … f l) s2' s2'' ∧ 70 ∃t3: raw_trace … S2 s2'' s2'''. 70 71 Srel … rel s1' s2''' ∧ silent_trace … t1 ∧ pre_silent_trace … t3 71 72 ∧ Crel … rel s1 s2' 72 73 ; simulate_ret_n: 73 74 ∀s1:S1.∀s2:S2.∀s1':S1. 74 as_execute … (ret_act (None ?)) s1 s1' →75 as_execute … (ret_act … (None ?)) s1 s1' → 75 76 as_classify … s1 ≠ cl_io → as_classify … s1' ≠ cl_io → 76 77 Srel … rel s1 s2 → 77 78 ∃s2',s2'',s2''': S2. 78 ∃t1: raw_trace S2 s2 s2'.79 as_execute … (ret_act (None ?)) s2' s2'' ∧80 ∃t3: raw_trace S2 s2'' s2'''.79 ∃t1: raw_trace … S2 s2 s2'. 80 as_execute … (ret_act … (None ?)) s2' s2'' ∧ 81 ∃t3: raw_trace … S2 s2'' s2'''. 81 82 Srel … rel s1' s2''' ∧ pre_silent_trace … t1 ∧ pre_silent_trace … t3 82 83 ∧ Rrel … rel s1' s2'' … … 104 105 qed. 105 106 106 lemma simulate_labelled_action : ∀ S1,S2 : abstract_status.107 ∀rel : relations S1 S2.107 lemma simulate_labelled_action : ∀p.∀S1,S2 : abstract_status p. 108 ∀rel : relations … S1 S2. 108 109 ∀s1,s2 : S1.∀s1' : S2. 109 ∀l : ActionLabel .110 ∀l : ActionLabel p. 110 111 simulation_conditions … rel → 111 is_costlabelled_act l →112 (is_call_act l → is_call_post_label … s1) →112 is_costlabelled_act … l → 113 (is_call_act … l → is_call_post_label … s1) → 113 114 as_execute … l s1 s2 → 114 115 Srel … rel s1 s1' → … … 119 120 silent_trace … t_pre ∧ silent_trace … t_post ∧ 120 121 as_execute … l pre_em post_em ∧ 121 (is_call_act l → is_call_post_label … pre_em) ∧122 (is_call_act … l → is_call_post_label … pre_em) ∧ 122 123 Srel … rel s2 s2'. 123 # S1 #S2 #rel #s1 #s2 #s1' *124 #p #S1 #S2 #rel #s1 #s2 #s1' * 124 125 [ #f #c | * [|#lbl] | * [|#lbl]] 125 126 #sim * #Hcall #prf #REL * #Hio … … 162 163 163 164 164 theorem simulates_pre_mesurable: 165 ∀S1,S2 : abstract_status .∀rel : relationsS1 S2.166 ∀s1,s1' : S1. ∀t1: raw_trace S1 s1 s1'.165 theorem simulates_pre_mesurable: ∀p. 166 ∀S1,S2 : abstract_status p.∀rel : relations … S1 S2. 167 ∀s1,s1' : S1. ∀t1: raw_trace … S1 s1 s1'. 167 168 simulation_conditions … rel → 168 169 pre_measurable_trace … t1 → … … 175 176 (measurable_suffix … t1 → measurable_suffix … t2) ∧ 176 177 (silent_trace … t1 → silent_trace … t2). 177 # S1 #S2 #rel #s1 #s1' #t1 #sim #H elim H -t1 -s1 -s1'178 #p #S1 #S2 #rel #s1 #s1' #t1 #sim #H elim H -t1 -s1 -s1' 178 179 [ #st #Hst #s2 #Hs2 #Rsts2 %{s2} %{(t_base …)} % 179 180 [ % [ % // % // % //] * #s_em ** [ #sx | #x1 #x2 #x3 #l1 #prf1 #tl1] * #t_pre * #s_post * #t_post * #sil_tpost … … 243 244 | #Hsuff cases(measurable_suffix_tind … Hsuff) 244 245 [ * #_ #sil_tl lapply(Hs sil_tl) #sil_ts3 %{s2''} %{t_s2''} %{s2'''} %{(t_s2'''' @ t_s3)} 245 % [ /3 by or_introl, append_silent/ ] %{(cost_act (Some ? lbl))} %{(exe_s2''')} % // % // * #f * #c #EQ destruct246 | #H @measurable_suffix_append change with ((t_ind ?????? t_s2'''')@t_s3) in ⊢ (????%);246 % [ /3 by or_introl, append_silent/ ] %{(cost_act … (Some ? lbl))} %{(exe_s2''')} % // % // * #f * #c #EQ destruct 247 | #H @measurable_suffix_append change with ((t_ind ??????? t_s2'''')@t_s3) in ⊢ (?????%); 247 248 @measurable_suffix_append /2/ 248 249 ] … … 279 280 | #Hsuff cases(measurable_suffix_tind … Hsuff) 280 281 [ * #_ #sil_tl lapply(Hs sil_tl) #sil_ts3 %{st1} %{t1} %{st2} %{(t2 @ t_fin)} 281 % [ /3 by or_introl, append_silent/ ] %{(ret_act (Some ? ret_lab))} %{(exe_s2'')} % // % // * #f * #c #EQ destruct282 | #H @measurable_suffix_append change with ((t_ind ?????? t2)@t_fin) in ⊢ (????%);282 % [ /3 by or_introl, append_silent/ ] %{(ret_act … (Some ? ret_lab))} %{(exe_s2'')} % // % // * #f * #c #EQ destruct 283 | #H @measurable_suffix_append change with ((t_ind ??????? t2)@t_fin) in ⊢ (?????%); 283 284 @measurable_suffix_append /2/ 284 285 ] … … 317 318 | #Hsuff cases(measurable_suffix_tind … Hsuff) 318 319 [ * #_ #sil_tl lapply(Hs sil_tl) #sil_ts3 %{s'} %{t1} %{s2''} %{(t2 @ t_fin)} 319 % [ /3 by or_introl, append_silent/ ] %{(call_act f lab)} %{(exe_s2'')} % // % //320 | #H @measurable_suffix_append change with ((t_ind ?????? t2)@t_fin) in ⊢ (????%);320 % [ /3 by or_introl, append_silent/ ] %{(call_act … f lab)} %{(exe_s2'')} % // % // 321 | #H @measurable_suffix_append change with ((t_ind ??????? t2)@t_fin) in ⊢ (?????%); 321 322 @measurable_suffix_append /2/ 322 323 ] … … 381 382 #st #EQ1 #EQ2 #EQ3 destruct normalize in ⊢ (% → ?); #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct 382 383 #ABS1 >ABS1 in Hst1; [2: /3 by ex_intro/] * 383 | #Hsuff @measurable_suffix_append change with (t_ind ??????? @ ?) in ⊢ (????%); @measurable_suffix_append384 change with (t_ind ??????? @ ?) in ⊢ (????%); @measurable_suffix_append @Hm2384 | #Hsuff @measurable_suffix_append change with (t_ind ???????? @ ?) in ⊢ (?????%); @measurable_suffix_append 385 change with (t_ind ???????? @ ?) in ⊢ (?????%); @measurable_suffix_append @Hm2 385 386 cases(measurable_suffix_append_case … Hsuff) 386 387 [ * #_ * [2: * #H cases(H I)] #H inversion H in ⊢ ?; … … 403 404 qed. 404 405 405 lemma silent_in_silent : ∀ S1,S2 : abstract_status.∀rel : relationsS1 S2.406 ∀s1,s1' : S1. ∀t1: raw_trace S1 s1 s1'.406 lemma silent_in_silent : ∀p.∀S1,S2 : abstract_status p.∀rel : relations … S1 S2. 407 ∀s1,s1' : S1. ∀t1: raw_trace … S1 s1 s1'. 407 408 simulation_conditions … rel → 408 409 ∀s2 : S2.Srel … rel s1 s2 → … … 411 412 silent_trace … t2 ∧ 412 413 Srel … rel s1' s2'. 414 #p 413 415 #S1 #S2 #rel #s1 #s1' #t1 #sim #s2 #REL * 414 416 [ #pre_sil_ti0 cases(simulates_pre_mesurable … sim … (pre_silent_is_premeasurable … pre_sil_ti0) … REL) … … 423 425 qed. 424 426 425 lemma tail_of_premeasurable_is_not_io : ∀ S : abstract_status.427 lemma tail_of_premeasurable_is_not_io : ∀p.∀S : abstract_status p. 426 428 ∀s1,s2 : S.∀t : raw_trace … s1 s2. 427 429 pre_measurable_trace … t → 428 as_classify … s2 ≠ cl_io. 430 as_classify … s2 ≠ cl_io. #p 429 431 #S #s1 #s2 #t #H elim H // 430 432 qed. 431 433 432 434 theorem simulates_measurable: 433 ∀S1,S2: abstract_status. 434 435 ∀rel: relations S1 S2. simulation_conditions … rel → 435 ∀p. 436 ∀S1,S2: abstract_status p. 437 438 ∀rel: relations … S1 S2. simulation_conditions … rel → 436 439 437 440 ∀si,sn: S1. ∀t: raw_trace … si sn. … … 451 454 452 455 ∃s1',s2'. measurable … s1' s2' … t'. 453 # S1 #S2 #rel #sim_rel #si #sn #t #s1 #s3 #Hmeas #s1' #Hclass #Rsi_si'456 #p #S1 #S2 #rel #sim_rel #si #sn #t #s1 #s3 #Hmeas #s1' #Hclass #Rsi_si' 454 457 cases Hmeas -Hmeas #s0 * #s2 * #ti0 * #t12 * #t3n * #l1 * #l2 * #prf1 * #prf2 ******* 455 458 #pre_t12 #EQ destruct #Hl1 #Hl2 #Hcall_fin #sil_ti0 #sil_t3n #Hcall_init … … 469 472 #sn' * #t_3n' * #sil_t3n' #Rsn_sn' %{sn'} 470 473 %{(ti0'@ t_s0'' @ (t_ind … prf1' (t_s1''@t12' @ t_s2'' @ (t_ind … prf2' (t_s3'' @ t_3n')))))} 471 % [ % // cases daemon (*TODO*) ] 474 % [ % // >(get_cost_label_invariant_for_append_silent_trace_l … sil_ti0) 475 >(get_cost_label_invariant_for_append_silent_trace_l … sil_ti0') 476 >(get_cost_label_invariant_for_append_silent_trace_l … sil_ts0'') 477 whd in ⊢ (??%%); @eq_f2 // 478 >(get_cost_label_invariant_for_append_silent_trace_l … sil_ts1'') 479 >get_cost_label_append >get_cost_label_append in ⊢ (???%); @eq_f2 // 480 >(get_cost_label_invariant_for_append_silent_trace_l … sil_ts2'') 481 whd in ⊢ (??%%); @eq_f2 // 482 >(get_cost_label_invariant_for_append_silent_trace_l … sil_ts3'') 483 >(get_cost_label_silent_is_empty … sil_t3n') >(get_cost_label_silent_is_empty … sil_t3n) % ] 472 484 %{s1''} %{s3''} whd %{s0''} %{s2''} %{(ti0' @t_s0'')} %{(t_s1'' @ t12' @ t_s2'')} 473 485 %{(t_s3'' @ t_3n')} %{l1} %{l2} %{prf1'} %{prf2'} % [2: /2/] 474 486 % [2: /2 by append_silent/] % [2: /2 by append_silent/] % [2: /2/] % // % // % // 475 487 @append_premeasurable_silent /2/ 476 qed. 477 478 488 qed. 489 490 -
LTS/Traces.ma
r3535 r3549 17 17 include "basics/deqsets.ma". 18 18 include "../src/ASM/Util.ma". 19 include "basics/finset.ma". 20 include "../src/utilities/hide.ma". 21 include "utils.ma". 22 23 (*Label parameters*) 24 25 record label_params : Type[1] ≝ 26 { label_type :> monoid 27 ; abelian_magg : is_abelian … label_type 28 ; succ_label : label_type → label_type 29 ; po_label : partial_order label_type 30 ; no_maps_in_id : ∀x.succ_label x ≠ x 31 ; le_lab_ok : ∀x,y.po_label … (op … x (succ_label y)) (succ_label (op … x y)) 32 ; magg_def : ∀x,y.po_label … x (op … x y) 33 ; monotonic_magg : ∀x,y,z.po_label … x y → po_label … (op … x z) (op … y z) 34 }. 35 36 notation "hvbox(a break ⊑ ^ {b} break term 46 c)" 37 non associative with precedence 45 38 for @{ 'lab_leq $a $b $c}. 39 40 interpretation "label 'less or equal to'" 'lab_leq x y z = (po_rel ? (po_label y) x z). 41 42 lemma max_1 : ∀p : label_params.∀n,m,k : p.k ⊑^{p} m → k ⊑^{p} op … p m n. 43 #p #m #n #k #H /2 by magg_def, trans_po_rel/ 44 qed. 45 46 lemma max_2 : ∀p : label_params.∀n,m,k: p.k ⊑^{p} n → k ⊑^{p} op … p m n. 47 #p #m #n #k #H >(abelian_magg … p) /2 by magg_def, trans_po_rel/ 48 qed. 49 50 lemma lab_po_rel_succ : ∀p : label_params.∀x : p.x ⊑^{p} succ_label … x. 51 #p #x @(trans_po_rel … (op … p … x (succ_label … (e … p)))) [ @(magg_def … p)] 52 @(trans_po_rel … (le_lab_ok … p …)) >neutral_r // 53 qed. 54 55 (*flat labels*) 56 57 definition deqnat_monoid ≝ mk_monoid DeqNat max O ???. 58 @hide_prf // [* //] #x #y #z normalize inversion(leb x y) normalize nodelta 59 #leb_xy 60 [ inversion(leb y z) normalize nodelta #leb_yz 61 [ >(le_to_leb_true …) // @(transitive_le … y) @leb_true_to_le // 62 | >leb_xy % 63 ] 64 | inversion(leb x z) normalize nodelta #leb_xz 65 [ >(le_to_leb_true … y z) normalize nodelta [>leb_xz %] @(transitive_le … x) /2 by leb_true_to_le/ 66 lapply(not_le_to_lt … (leb_false_to_not_le … leb_xy)) #H1 /2/ 67 | cases(leb y z) normalize nodelta [ >leb_xz | >leb_xy ] // 68 ] 69 ] 70 qed. 71 72 definition part_order_nat ≝ mk_partial_order ℕ le ???. 73 /2 by le_to_le_to_eq, transitive_le, le_n/ 74 qed. 75 76 definition flat_labels ≝ mk_label_params deqnat_monoid ? S part_order_nat ????. 77 @hide_prf 78 [ normalize #x #y @leb_elim normalize nodelta 79 [2: #H >le_to_leb_true // lapply(not_le_to_lt … H) #H1 /2/ 80 | * -y [ >le_to_leb_true //] #y #H >not_le_to_leb_false // @lt_to_not_le /2 by le_to_lt_to_lt/ 81 ] 82 | normalize #x /2 by sym_not_eq/ 83 | normalize #x #y @(leb_elim … x y) normalize nodelta 84 [ #H >le_to_leb_true /3 by le_to_lt_to_lt, monotonic_pred, le_n/ 85 | #H @leb_elim normalize nodelta [ #H1 @le_S_S lapply(not_le_to_lt … H) #H3 /2/ ] 86 #_ /2/ 87 ] 88 | normalize #x #y @leb_elim // 89 | normalize #x #y #z #H @(leb_elim … y z) 90 [ #H1 >(le_to_leb_true … (transitive_le … H H1)) // 91 | #H1 normalize cases(leb ??) normalize /3 by not_le_to_lt, lt_to_le/ 92 ] 93 ] 94 qed. 19 95 20 96 (*LABELS*) … … 23 99 | a_function_id : ℕ → FunctionName. 24 100 25 inductive CallCostLabel : Type[0] ≝26 | a_call_label : ℕ → CallCostLabel.101 inductive CallCostLabel (p : label_params) : Type[0] ≝ 102 | a_call_label : p → CallCostLabel p. 27 103 28 inductive ReturnPostCostLabel : Type[0] ≝29 | a_return_cost_label : ℕ → ReturnPostCostLabel.104 inductive ReturnPostCostLabel (p : label_params) : Type[0] ≝ 105 | a_return_cost_label : p → ReturnPostCostLabel p. 30 106 31 inductive NonFunctionalLabel : Type[0] ≝32 | a_non_functional_label : ℕ → NonFunctionalLabel.107 inductive NonFunctionalLabel (p : label_params) : Type[0] ≝ 108 | a_non_functional_label : p → NonFunctionalLabel p. 33 109 34 inductive CostLabel : Type[0] ≝35 | a_call : CallCostLabel → CostLabel36 | a_return_post : ReturnPostCostLabel → CostLabel37 | a_non_functional_label : NonFunctionalLabel → CostLabel.110 inductive CostLabel (p : label_params) : Type[0] ≝ 111 | a_call : CallCostLabel p → CostLabel p 112 | a_return_post : ReturnPostCostLabel p → CostLabel p 113 | a_non_functional_label : NonFunctionalLabel p → CostLabel p. 38 114 39 115 coercion a_call. … … 41 117 coercion a_non_functional_label. 42 118 43 definition eq_nf_label : NonFunctionalLabel → NonFunctionalLabel → bool ≝ 44 λx,y.match x with [ a_non_functional_label n ⇒ 45 match y with [a_non_functional_label m ⇒ eqb n m ] ]. 46 47 lemma eq_fn_label_elim : ∀P : bool → Prop.∀l1,l2 : NonFunctionalLabel. 48 (l1 = l2 → P true) → (l1 ≠ l2 → P false) → P (eq_nf_label l1 l2). 49 #P * #l1 * #l2 #H1 #H2 normalize @eqb_elim /3/ * #H3 @H2 % #EQ destruct @H3 % qed. 50 51 definition DEQNonFunctionalLabel ≝ mk_DeqSet NonFunctionalLabel eq_nf_label ?. 119 definition eq_nf_label : ∀p.NonFunctionalLabel p → NonFunctionalLabel p → bool ≝ 120 λp,x,y.match x with [ a_non_functional_label n ⇒ 121 match y with [a_non_functional_label m ⇒ eqb … n m ] ]. 122 123 lemma eq_fn_label_elim : ∀P : bool → Prop.∀p.∀l1,l2 : NonFunctionalLabel p. 124 (l1 = l2 → P true) → (l1 ≠ l2 → P false) → P (eq_nf_label … l1 l2). 125 #P #p * #l1 * #l2 #H1 #H2 normalize inversion(?==?) #H 126 [ @H1 >(\P H) % 127 | @H2 % #EQ destruct lapply(\Pf H) * #H3 @H3 % 128 ] 129 qed. 130 131 definition DEQNonFunctionalLabel ≝ λp.mk_DeqSet (NonFunctionalLabel p) (eq_nf_label p) ?. 52 132 #x #y @eq_fn_label_elim [ #EQ destruct /2/] * #H % [ #EQ destruct] #H1 @⊥ @H 53 133 assumption 54 134 qed. 55 135 56 unification hint 0 ≔ ;57 X ≟ DEQNonFunctionalLabel136 unification hint 0 ≔ p; 137 X ≟ (DEQNonFunctionalLabel p) 58 138 (* ---------------------------------------- *) ⊢ 59 NonFunctionalLabel≡ carr X.60 61 unification hint 0 ≔ p 1,p2;62 X ≟ DEQNonFunctionalLabel139 (NonFunctionalLabel p) ≡ carr X. 140 141 unification hint 0 ≔ p,p1,p2; 142 X ≟ (DEQNonFunctionalLabel p) 63 143 (* ---------------------------------------- *) ⊢ 64 eq_nf_label p 1 p2 ≡ eqb X p1 p2.144 eq_nf_label p p1 p2 ≡ eqb X p1 p2. 65 145 66 146 definition eq_function_name : FunctionName → FunctionName → bool ≝ … … 86 166 eq_function_name p1 p2 ≡ eqb X p1 p2. 87 167 88 definition eq_return_cost_lab : ReturnPostCostLabel → ReturnPostCostLabel → bool ≝ 89 λc1,c2.match c1 with [a_return_cost_label n ⇒ match c2 with [ a_return_cost_label m ⇒ eqb n m]]. 90 91 lemma eq_return_cost_lab_elim : ∀P : bool → Prop.∀c1,c2.(c1 = c2 → P true) → 92 (c1 ≠ c2 → P false) → P (eq_return_cost_lab c1 c2). 93 #P * #c1 * #c2 #H1 #H2 normalize @eqb_elim /2/ * #H @H2 % #EQ destruct @H % qed. 94 95 definition DEQReturnPostCostLabel ≝ mk_DeqSet ReturnPostCostLabel eq_return_cost_lab ?. 168 definition eq_return_cost_lab :∀p.ReturnPostCostLabel p → ReturnPostCostLabel p → bool ≝ 169 λp,c1,c2.match c1 with [a_return_cost_label n ⇒ match c2 with [ a_return_cost_label m ⇒ eqb … n m]]. 170 171 lemma eq_return_cost_lab_elim : ∀P : bool → Prop.∀p.∀c1,c2.(c1 = c2 → P true) → 172 (c1 ≠ c2 → P false) → P (eq_return_cost_lab p c1 c2). 173 #P #p * #l1 * #l2 #H1 #H2 normalize inversion(?==?) #H 174 [ @H1 >(\P H) % 175 | @H2 % #EQ destruct lapply(\Pf H) * #H3 @H3 % 176 ] 177 qed. 178 179 definition DEQReturnPostCostLabel ≝ λp.mk_DeqSet (ReturnPostCostLabel p) (eq_return_cost_lab …) ?. 96 180 #x #y @eq_return_cost_lab_elim [ #EQ destruct /2/] * #H % [ #EQ destruct] #H1 @⊥ @H 97 181 assumption 98 182 qed. 99 183 100 unification hint 0 ≔ ;101 X ≟ DEQReturnPostCostLabel184 unification hint 0 ≔ p ; 185 X ≟ (DEQReturnPostCostLabel p) 102 186 (* ---------------------------------------- *) ⊢ 103 ReturnPostCostLabel≡ carr X.104 105 unification hint 0 ≔ p 1,p2;106 X ≟ DEQReturnPostCostLabel187 (ReturnPostCostLabel p) ≡ carr X. 188 189 unification hint 0 ≔ p,p1,p2; 190 X ≟ (DEQReturnPostCostLabel p) 107 191 (* ---------------------------------------- *) ⊢ 108 eq_return_cost_lab p1 p2 ≡ eqb X p1 p2. 109 110 definition eq_call_cost_lab : CallCostLabel → CallCostLabel → bool ≝ 111 λc1,c2.match c1 with [a_call_label x ⇒ match c2 with [a_call_label y ⇒ eqb x y ]]. 112 113 lemma eq_call_cost_lab_elim : ∀P : bool → Prop.∀c1,c2.(c1 = c2 → P true) → 114 (c1 ≠ c2 → P false) → P (eq_call_cost_lab c1 c2). 115 #P * #c1 * #c2 #H1 #H2 normalize @eqb_elim /2/ * #H @H2 % #EQ destruct @H % qed. 116 117 definition DEQCallCostLabel ≝ mk_DeqSet CallCostLabel eq_call_cost_lab ?. 192 eq_return_cost_lab p p1 p2 ≡ eqb X p1 p2. 193 194 definition eq_call_cost_lab : ∀p.(CallCostLabel p) → (CallCostLabel p) → bool ≝ 195 λp,c1,c2.match c1 with [a_call_label x ⇒ match c2 with [a_call_label y ⇒ eqb … x y ]]. 196 197 lemma eq_call_cost_lab_elim : ∀P : bool → Prop.∀p.∀c1,c2.(c1 = c2 → P true) → 198 (c1 ≠ c2 → P false) → P (eq_call_cost_lab p c1 c2). 199 #P #p * #l1 * #l2 #H1 #H2 normalize inversion(?==?) #H 200 [ @H1 >(\P H) % 201 | @H2 % #EQ destruct lapply(\Pf H) * #H3 @H3 % 202 ] 203 qed. 204 205 definition DEQCallCostLabel ≝ λp.mk_DeqSet (CallCostLabel p) (eq_call_cost_lab p) ?. 118 206 #x #y @eq_call_cost_lab_elim [ #EQ destruct /2/] * #H % [ #EQ destruct] #H1 @⊥ @H 119 207 assumption 120 208 qed. 121 209 122 unification hint 0 ≔ ;123 X ≟ DEQCallCostLabel210 unification hint 0 ≔ p; 211 X ≟ (DEQCallCostLabel p) 124 212 (* ---------------------------------------- *) ⊢ 125 CallCostLabel≡ carr X.126 127 unification hint 0 ≔ p 1,p2;128 X ≟ DEQCallCostLabel213 (CallCostLabel p) ≡ carr X. 214 215 unification hint 0 ≔ p,p1,p2; 216 X ≟ (DEQCallCostLabel p) 129 217 (* ---------------------------------------- *) ⊢ 130 eq_call_cost_lab p 1 p2 ≡ eqb X p1 p2.131 132 definition eq_costlabel : CostLabel → CostLabel→ bool ≝133 λ c1.match c1 with218 eq_call_cost_lab p p1 p2 ≡ eqb X p1 p2. 219 220 definition eq_costlabel :∀p. (CostLabel p) → (CostLabel p) → bool ≝ 221 λp,c1.match c1 with 134 222 [ a_call x1 ⇒ λc2.match c2 with [a_call y1 ⇒ x1 == y1 | _ ⇒ false ] 135 223 | a_return_post x1 ⇒ … … 139 227 ]. 140 228 141 lemma eq_costlabel_elim : ∀P : bool → Prop.∀ c1,c2.(c1 = c2 → P true) →142 (c1 ≠ c2 → P false) → P (eq_costlabel c1 c2).143 #P * #c1 * #c2 #H1 #H2 whd in match (eq_costlabel ??);229 lemma eq_costlabel_elim : ∀P : bool → Prop.∀p.∀c1,c2.(c1 = c2 → P true) → 230 (c1 ≠ c2 → P false) → P (eq_costlabel p c1 c2). 231 #P #p * #c1 * #c2 #H1 #H2 whd in match (eq_costlabel ??); 144 232 try (@H2 % #EQ destruct) 145 [ change with (eq_call_cost_lab ?? ) in ⊢ (?%); @eq_call_cost_lab_elim233 [ change with (eq_call_cost_lab ???) in ⊢ (?%); @eq_call_cost_lab_elim 146 234 [ #EQ destruct @H1 % | * #H @H2 % #EQ destruct @H % ] 147 | change with (eq_return_cost_lab ?? ) in ⊢ (?%);235 | change with (eq_return_cost_lab ???) in ⊢ (?%); 148 236 @eq_return_cost_lab_elim 149 237 [ #EQ destruct @H1 % | * #H @H2 % #EQ destruct @H % ] 150 | change with (eq_nf_label ?? ) in ⊢ (?%); @eq_fn_label_elim238 | change with (eq_nf_label ???) in ⊢ (?%); @eq_fn_label_elim 151 239 [ #EQ destruct @H1 % | * #H @H2 % #EQ destruct @H % ] 152 240 ] … … 154 242 155 243 156 definition DEQCostLabel ≝ mk_DeqSet CostLabel eq_costlabel?.244 definition DEQCostLabel ≝ λp.mk_DeqSet (CostLabel p) (eq_costlabel p) ?. 157 245 #x #y @eq_costlabel_elim [ #EQ destruct /2/] * #H % [ #EQ destruct] #H1 @⊥ @H 158 246 assumption 159 247 qed. 160 248 161 unification hint 0 ≔ ;162 X ≟ DEQCostLabel249 unification hint 0 ≔ p ; 250 X ≟ (DEQCostLabel p) 163 251 (* ---------------------------------------- *) ⊢ 164 CostLabel≡ carr X.165 166 unification hint 0 ≔ p 1,p2;167 X ≟ DEQCostLabel252 (CostLabel p) ≡ carr X. 253 254 unification hint 0 ≔ p,p1,p2; 255 X ≟ (DEQCostLabel p) 168 256 (* ---------------------------------------- *) ⊢ 169 eq_costlabel p 1 p2 ≡ eqb X p1 p2.170 171 inductive ActionLabel : Type[0] ≝172 | call_act : FunctionName → CallCostLabel → ActionLabel173 | ret_act : option(ReturnPostCostLabel ) → ActionLabel174 | cost_act : option NonFunctionalLabel → ActionLabel.257 eq_costlabel p p1 p2 ≡ eqb X p1 p2. 258 259 inductive ActionLabel (p : label_params) : Type[0] ≝ 260 | call_act : FunctionName → CallCostLabel p → ActionLabel p 261 | ret_act : option(ReturnPostCostLabel p) → ActionLabel p 262 | cost_act : option (NonFunctionalLabel p) → ActionLabel p. 175 263 176 definition is_cost_label : ActionLabel → Prop ≝ 177 λact.match act with [ cost_act nf ⇒ True | _ ⇒ False ]. 178 179 definition is_non_silent_cost_act : ActionLabel → Prop ≝ 180 λact. ∃l. act = cost_act (Some ? l). 181 182 definition is_cost_act : ActionLabel → Prop ≝ 183 λact.∃l.act = cost_act l. 184 185 definition is_call_act : ActionLabel → Prop ≝ 186 λact.∃f,l.act = call_act f l. 187 188 definition is_labelled_ret_act : ActionLabel → Prop ≝ 189 λact.∃l.act = ret_act (Some ? l). 190 191 definition is_unlabelled_ret_act : ActionLabel → Prop ≝ 192 λact.act = ret_act (None ?). 193 194 definition is_costlabelled_act : ActionLabel → Prop ≝ 195 λact.match act with [ call_act _ _ ⇒ True 264 definition eq_ActionLabel : ∀p.ActionLabel p → ActionLabel p → bool ≝ 265 λp,c1,c2. 266 match c1 with 267 [ call_act f l ⇒ match c2 with [ call_act f' l' ⇒ f == f' ∧ l == l' | _ ⇒ false] 268 | ret_act x ⇒ match c2 with [ret_act y ⇒ match x with [ None ⇒ match y with [ None ⇒ true | _ ⇒ false] 269 | Some l ⇒ match y with [ Some l' ⇒ l == l' | _ ⇒ false] 270 ] 271 | _ ⇒ false 272 ] 273 | cost_act x ⇒ match c2 with [cost_act y ⇒ match x with [ None ⇒ match y with [ None ⇒ true | _ ⇒ false] 274 | Some l ⇒ match y with [ Some l' ⇒ l == l' | _ ⇒ false] 275 ] 276 | _ ⇒ false 277 ] 278 ]. 279 280 definition is_cost_label : ∀p.ActionLabel p → Prop ≝ 281 λp,act.match act with [ cost_act nf ⇒ True | _ ⇒ False ]. 282 283 definition is_non_silent_cost_act : ∀p.ActionLabel p → Prop ≝ 284 λp,act. ∃l. act = cost_act … (Some ? l). 285 286 definition is_cost_act : ∀p.ActionLabel p → Prop ≝ 287 λp,act.∃l.act = cost_act … l. 288 289 definition is_call_act : ∀p.ActionLabel p → Prop ≝ 290 λp,act.∃f,l.act = call_act … f l. 291 292 definition is_labelled_ret_act : ∀p.(ActionLabel p) → Prop ≝ 293 λp,act.∃l.act = ret_act … (Some ? l). 294 295 definition is_unlabelled_ret_act : ∀p.ActionLabel p → Prop ≝ 296 λp,act.act = ret_act … (None ?). 297 298 definition is_costlabelled_act : ∀p.ActionLabel p → Prop ≝ 299 λp,act.match act with [ call_act _ _ ⇒ True 196 300 | ret_act x ⇒ match x with [ Some _ ⇒ True | None ⇒ False ] 197 301 | cost_act x ⇒ match x with [ Some _ ⇒ True | None ⇒ False ] 198 302 ]. 199 303 200 lemma is_costlabelled_act_elim : 201 ∀P : ActionLabel → Prop. 202 (∀l. is_costlabelled_act l → P l) → 203 (∀l. ¬is_costlabelled_act l → P l) → 204 ∀l.P l. 304 definition is_ret_call_act : ∀p.ActionLabel p → Prop ≝ 305 λp,a.match a with [ret_act _ ⇒ True | call_act _ _ ⇒ True | _ ⇒ False ]. 306 307 definition is_silent_cost_act_b : ∀p.ActionLabel p → bool ≝ 308 λp,act. match act with 309 [ cost_act x ⇒ match x with [None ⇒ true | _ ⇒ false ] | _ ⇒ false]. 310 311 lemma is_costlabelled_act_elim : ∀p. 312 ∀P : ActionLabel p → Prop. 313 (∀l. is_costlabelled_act p l → P l) → 314 (∀l. ¬is_costlabelled_act p l → P l) → 315 ∀l.P l. #p 205 316 #P #H1 #H2 * /2/ 206 317 [ * /2/ @H2 % *] * /2/ @H2 %* 207 318 qed. 319 320 lemma eq_a_call_label : ∀p. 321 ∀c1,c2.c1 == c2 → a_call p c1 == a_call p c2. 322 #p #c1 #c2 #EQ cases(eqb_true ? c1 c2) #H1 #_ lapply(H1 (eq_true_to_b … EQ)) -EQ 323 #EQ destruct >(\b (refl …)) @I 324 qed. 325 326 lemma eq_a_non_functional_label : ∀p. 327 ∀c1,c2.c1 == c2 → a_non_functional_label p c1 == a_non_functional_label p c2. 328 #p #c1 #c2 #EQ cases(eqb_true (DEQNonFunctionalLabel ?) c1 c2) #H1 #_ lapply(H1 (eq_true_to_b … EQ)) -EQ 329 #EQ destruct >(\b (refl …)) @I 330 qed. 331 332 (* TOOLS FOR FRESHING LABELS *) 333 334 definition fresh_nf_label : ∀p : label_params.p → NonFunctionalLabel p × p ≝ 335 λp,x.〈a_non_functional_label … (succ_label … p x),(succ_label … p x)〉. 336 337 definition fresh_cc_labe : ∀p : label_params.p → CallCostLabel p × p ≝ 338 λp,x.〈a_call_label … (succ_label … p x),(succ_label … p x)〉. 339 340 definition fresh_rc_label : ∀p : label_params.p → ReturnPostCostLabel p × p ≝ 341 λp,x.〈a_return_cost_label … (succ_label … p x),(succ_label … p x)〉. 208 342 209 343 … … 215 349 | cl_other : status_class. 216 350 217 record abstract_status : Type[2] ≝351 record abstract_status (p : label_params) : Type[2] ≝ 218 352 { as_status :> Type[0] 219 ; as_execute : ActionLabel→ relation as_status353 ; as_execute : (ActionLabel p) → relation as_status 220 354 ; as_syntactical_succ : relation as_status 221 355 ; as_classify : as_status → status_class … … 225 359 ; jump_emits : ∀s1,s2,l. 226 360 as_classify … s1 = cl_jump → 227 as_execute l s1 s2 → is_non_silent_cost_act l361 as_execute l s1 s2 → is_non_silent_cost_act … l 228 362 ; io_entering : ∀s1,s2,l.as_classify … s2 = cl_io → 229 as_execute l s1 s2 → is_non_silent_cost_act l363 as_execute l s1 s2 → is_non_silent_cost_act … l 230 364 ; io_exiting : ∀s1,s2,l.as_classify … s1 = cl_io → 231 as_execute l s1 s2 → is_non_silent_cost_act l365 as_execute l s1 s2 → is_non_silent_cost_act … l 232 366 }. 233 367 234 368 (* RAW TRACES *) 235 369 236 inductive raw_trace (S : abstract_status) : S → S → Type[0] ≝ 237 | t_base : ∀st : S.raw_trace S st st 238 | t_ind : ∀st1,st2,st3 : S.∀l : ActionLabel. 239 as_execute S l st1 st2 → raw_trace S st2 st3 → 240 raw_trace S st1 st3. 241 242 let rec append_on_trace (S : abstract_status) (st1 : S) (st2 : S) (st3 : S) 243 (t1 : raw_trace S st1 st2) on t1 : raw_trace S st2 st3 → raw_trace S st1 st3 ≝ 370 inductive raw_trace (p : label_params) (S : abstract_status p) : S → S → Type[0] ≝ 371 | t_base : ∀st : S.raw_trace … S st st 372 | t_ind : ∀st1,st2,st3 : S.∀l : ActionLabel p. 373 as_execute … S l st1 st2 → raw_trace … S st2 st3 → 374 raw_trace … S st1 st3. 375 376 377 let rec append_on_trace (p : label_params) (S : abstract_status p) (st1 : S) (st2 : S) (st3 : S) 378 (t1 : raw_trace … S st1 st2) on t1 : raw_trace … S st2 st3 → raw_trace … S st1 st3 ≝ 244 379 match t1 245 return λst1,st2,tr.raw_trace S st2 st3 → raw_traceS st1 st3380 return λst1,st2,tr.raw_trace … S st2 st3 → raw_trace … S st1 st3 246 381 with 247 382 [ t_base st ⇒ λt2.t2 … … 249 384 ]. 250 385 251 interpretation "trace_append" 'append t1 t2 = (append_on_trace ???? t1 t2). 252 253 lemma append_nil_is_nil : ∀S : abstract_status. 386 interpretation "trace_append" 'append t1 t2 = (append_on_trace ????? t1 t2). 387 388 let rec len (p : label_params) (s : abstract_status p) (st1 : s) (st2 : s) (t : raw_trace p s st1 st2) 389 on t : ℕ ≝ 390 match t with 391 [ t_base s ⇒ O 392 | t_ind s1 s2 s3 l prf lt ⇒ S (len … lt) 393 ]. 394 395 lemma len_append : ∀p.∀S : abstract_status p.∀st1,st2,st3 : S. 396 ∀t1 : raw_trace … st1 st2.∀t2 : raw_trace … st2 st3. 397 len ???? (t1 @ t2) = (len ???? t1) + (len ???? t2). 398 #p #S #st1 #st2 #st3 #t1 elim t1 normalize // 399 qed. 400 401 lemma append_nil_is_nil : ∀p. ∀S : abstract_status p. 254 402 ∀s1,s2 : S.∀t1 : raw_trace … s1 s2.∀t2 : raw_trace … s2 s1. 255 403 t1 @ t2 = t_base … s1 → s1 = s2 ∧ t1 ≃ t_base … s1 ∧ t2 ≃ t_base … s1. 256 # S #s1 #s2 #t1 elim t1 -t1 -s1 -s2404 #p #S #s1 #s2 #t1 elim t1 -t1 -s1 -s2 257 405 [ #st #t2 normalize #EQ destruct /3 by refl_jmeq, conj/] 258 406 #s1 #s2 #s3 #l #prf #t2 #IH #t2 normalize #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct 259 407 qed. 260 408 261 lemma append_associative : ∀ S,st1,st2,st3,st4.262 ∀t1 : raw_trace S st1 st2.∀t2 : raw_traceS st2 st3.263 ∀t3 : raw_trace S st3 st4.(t1 @ t2) @ t3 = t1 @ (t2 @ t3).264 # S #st1 #st2 #st3 #st4 #t1 elim t1 -t1409 lemma append_associative : ∀p.∀S,st1,st2,st3,st4. 410 ∀t1 : raw_trace p S st1 st2.∀t2 : raw_trace … S st2 st3. 411 ∀t3 : raw_trace … S st3 st4.(t1 @ t2) @ t3 = t1 @ (t2 @ t3). 412 #p #S #st1 #st2 #st3 #st4 #t1 elim t1 -t1 265 413 [ #st #t2 #t3 %] #st1' #st2' #st3' #l #prf #tl #IH #t2 #t3 whd in ⊢ (??%%); >IH % 266 414 qed. 267 415 268 definition trace_prefix: ∀ S : abstract_status.∀st1,st2,st3 : S.raw_trace … st1 st2 →416 definition trace_prefix: ∀p.∀S : abstract_status p.∀st1,st2,st3 : S.raw_trace … st1 st2 → 269 417 raw_trace … st1 st3 → Prop≝ 270 λ S,st1,st2,st3,t1,t2.∃t3 : raw_trace … st2 st3.t2 = t1 @ t3.271 272 definition trace_suffix : ∀ S : abstract_status.∀st1,st2,st3 : S.raw_trace … st2 st3 →418 λp,S,st1,st2,st3,t1,t2.∃t3 : raw_trace … st2 st3.t2 = t1 @ t3. 419 420 definition trace_suffix : ∀p.∀S : abstract_status p.∀st1,st2,st3 : S.raw_trace … st2 st3 → 273 421 raw_trace … st1 st3 → Prop≝ 274 λ S,st1,st2,st3,t1,t2.∃t3 : raw_trace … st1 st2.t2 = t3 @ t1.275 276 definition actionlabel_to_costlabel ≝ 277 λl : ActionLabel .match l with422 λp,S,st1,st2,st3,t1,t2.∃t3 : raw_trace … st1 st2.t2 = t3 @ t1. 423 424 definition actionlabel_to_costlabel ≝ λp. 425 λl : ActionLabel p.match l with 278 426 [ call_act f c ⇒ [ c ] 279 427 | ret_act x ⇒ match x with 280 [ Some c ⇒ [ a_return_post c ]428 [ Some c ⇒ [ a_return_post … c ] 281 429 | None ⇒ [] 282 430 ] 283 431 | cost_act x ⇒ match x with 284 [ Some c ⇒ [ a_non_functional_label c ]432 [ Some c ⇒ [ a_non_functional_label … c ] 285 433 | None ⇒ [] 286 434 ] 287 435 ]. 288 436 289 let rec get_costlabels_of_trace ( S : abstract_status) (st1 : S) (st2 : S)290 (t : raw_trace S st1 st2) on t : list CostLabel≝437 let rec get_costlabels_of_trace (p : label_params) (S : abstract_status p) (st1 : S) (st2 : S) 438 (t : raw_trace p S st1 st2) on t : list (CostLabel p) ≝ 291 439 match t with 292 440 [ t_base st ⇒ [ ] 293 441 | t_ind st1' st2' st3' l prf t' ⇒ 294 442 let tl ≝ get_costlabels_of_trace … t' in 295 actionlabel_to_costlabel l @ tl443 actionlabel_to_costlabel … l @ tl 296 444 ]. 297 445 298 lemma get_cost_label_append : ∀ S : abstract_status.∀st1,st2,st3 : S.446 lemma get_cost_label_append : ∀p.∀S : abstract_status p.∀st1,st2,st3 : S. 299 447 ∀t1 : raw_trace … st1 st2. ∀t2 : raw_trace … st2 st3. 300 448 get_costlabels_of_trace … (t1 @ t2) = 301 449 append … (get_costlabels_of_trace … t1) (get_costlabels_of_trace … t2). 302 # S #st1 #st2 #st3 #t1 elim t1 [ #st #t2 % ] #st1' #st2' #st3' *450 #p #S #st1 #st2 #st3 #t1 elim t1 [ #st #t2 % ] #st1' #st2' #st3' * 303 451 [#f * #l | * [| * #l] | * [| #l] ] #prf #tl #IH #t2 normalize try @eq_f @IH 304 452 qed. 305 453 306 lemma append_tind_commute : ∀ S : abstract_status.∀st1,st2,st3,st4 : S.∀l.307 ∀prf : as_execute … l st1 st2. ∀t1 : raw_trace S st2 st3.308 ∀t2 : raw_trace S st3 st4.t_ind … prf (t1 @ t2) = (t_ind … prf t1) @ t2.309 # S #st1 #st2 #st3 #st4 #l #prf #t1 #t2 % qed.454 lemma append_tind_commute : ∀p.∀S : abstract_status p.∀st1,st2,st3,st4 : S.∀l. 455 ∀prf : as_execute … l st1 st2. ∀t1 : raw_trace p S st2 st3. 456 ∀t2 : raw_trace p S st3 st4.t_ind … prf (t1 @ t2) = (t_ind … prf t1) @ t2. 457 #p #S #st1 #st2 #st3 #st4 #l #prf #t1 #t2 % qed. 310 458 311 459 lemma get_costlabelled_is_costlabelled : 312 ∀ S : abstract_status.∀s1,s2,s3 : S. ∀l.460 ∀p.∀S : abstract_status p.∀s1,s2,s3 : S. ∀l. 313 461 ∀prf : as_execute … l s1 s2.∀t : raw_trace … s2 s3. 314 is_costlabelled_act l →462 is_costlabelled_act p l → 315 463 |get_costlabels_of_trace … (t_ind … prf t)| = 1 + |get_costlabels_of_trace … t|. 316 # S #s1 #s2 #s3 * normalize464 #p #S #s1 #s2 #s3 * normalize 317 465 [ /2 by monotonic_pred/ 318 466 | * normalize /2 by lt_to_le, monotonic_pred/ #_ #t * … … 323 471 (*SILENT TRACES*) 324 472 325 inductive pre_silent_trace ( S : abstract_status) :326 ∀st1,st2 : S.raw_trace S st1 st2 → Prop ≝327 | silent_empty : ∀st : S.as_classify … st ≠ cl_io → pre_silent_trace … (t_base S st)328 | silent_cons : ∀st1,st2,st3 : S.∀prf : as_execute S (cost_act(None ?)) st1 st2.329 ∀tl : raw_trace S st2 st3. as_classify … st1 ≠ cl_io → pre_silent_trace … tl →473 inductive pre_silent_trace (p : label_params) (S : abstract_status p) : 474 ∀st1,st2 : S.raw_trace p S st1 st2 → Prop ≝ 475 | silent_empty : ∀st : S.as_classify … st ≠ cl_io → pre_silent_trace … (t_base p S st) 476 | silent_cons : ∀st1,st2,st3 : S.∀prf : as_execute p S (cost_act … (None ?)) st1 st2. 477 ∀tl : raw_trace p S st2 st3. as_classify … st1 ≠ cl_io → pre_silent_trace … tl → 330 478 pre_silent_trace … (t_ind … prf … tl). 331 479 332 definition is_trace_non_empty : ∀ S : abstract_status.∀st1,st2.333 raw_trace S st1 st2 → bool ≝334 λ S,st1,st2,t.match t with [ t_base _ ⇒ false | _ ⇒ true ].335 336 definition silent_trace : ∀ S:abstract_status.∀s1,s2 : S.337 raw_trace S s1 s2 → Prop ≝ λS,s1,s2,t.pre_silent_trace … t ∨480 definition is_trace_non_empty : ∀p.∀S : abstract_status p.∀st1,st2. 481 raw_trace p S st1 st2 → bool ≝ 482 λp,S,st1,st2,t.match t with [ t_base _ ⇒ false | _ ⇒ true ]. 483 484 definition silent_trace : ∀p.∀S:abstract_status p.∀s1,s2 : S. 485 raw_trace p S s1 s2 → Prop ≝ λp,S,s1,s2,t.pre_silent_trace … t ∨ 338 486 ¬ (bool_to_Prop (is_trace_non_empty … t)). 339 487 340 lemma pre_silent_io : ∀ S : abstract_status.488 lemma pre_silent_io : ∀p.∀S : abstract_status p. 341 489 ∀s1,s2 : S. ∀t : raw_trace … s1 s2. pre_silent_trace … t → 342 490 as_classify … s2 ≠ cl_io. 343 # S #s1 #s2 #t elim t [ #st #H inversion H // #st1 #st2 #st3 #prf #tl #H1 #sil_tl #_ #EQ1 #EQ2 #EQ3 destruct]491 #p #S #s1 #s2 #t elim t [ #st #H inversion H // #st1 #st2 #st3 #prf #tl #H1 #sil_tl #_ #EQ1 #EQ2 #EQ3 destruct] 344 492 #st1 #st2 #st3 #l #prf #tl #IH #H inversion H // 345 493 #st1' #st2' #st3' #prf' #tl' #Hclass #silent_tl' #_ #EQ1 #EQ2 destruct … … 347 495 qed. 348 496 349 lemma append_silent : ∀ S : abstract_status.497 lemma append_silent : ∀p.∀S : abstract_status p. 350 498 ∀s1,s2,s3 : S. ∀t1 : raw_trace … s1 s2. 351 499 ∀t2 : raw_trace … s2 s3.silent_trace … t1 → 352 500 silent_trace … t2 → 353 501 silent_trace … (t1 @ t2). 354 # S #s1 #s2 #s3 #t1 elim t1 -t1 -s1 -s2502 #p #S #s1 #s2 #s3 #t1 elim t1 -t1 -s1 -s2 355 503 [ #st #t2 #_ * /2 by or_introl, or_intror/ ] 356 504 #st1 #st2 #st3 #l #prf #tl #IH #t2 * … … 364 512 qed. 365 513 366 lemma silent_append_l2 : ∀ S : abstract_status.514 lemma silent_append_l2 : ∀p.∀S : abstract_status p. 367 515 ∀s1,s2,s3 :S.∀t1 : raw_trace … s1 s2.∀t2 : raw_trace … s2 s3. 368 516 silent_trace … (t1 @ t2) → silent_trace … t2. 369 # S #s1 #s2 #s3 #t1 elim t1 // -s1 -s2517 #p #S #s1 #s2 #s3 #t1 elim t1 // -s1 -s2 370 518 #st1 #st2 #st3 #l #prf #tl #IH #t2 * [2: * #H cases(H I)] 371 519 #H @IH % inversion H in ⊢ ?; … … 375 523 qed. 376 524 377 lemma silent_append_l1 : ∀ S : abstract_status.525 lemma silent_append_l1 : ∀p.∀S : abstract_status p. 378 526 ∀s1,s2,s3 :S.∀t1 : raw_trace … s1 s2.∀t2 : raw_trace … s2 s3. 379 527 silent_trace … (t1 @ t2) → silent_trace … t1. 380 # S #s1 #s2 #s3 #t1 elim t1 -s1 -s2528 #p #S #s1 #s2 #s3 #t1 elim t1 -s1 -s2 381 529 [ #st #t2 #_ %2 % *] 382 530 #st1 #st2 #st3 #l #prf #tl #IH #t2 * [2: * #H cases(H I)] #H … … 391 539 qed. 392 540 393 lemma get_cost_label_silent_is_empty : ∀ S : abstract_status.394 ∀st1,st2.∀t : raw_trace S st1 st2.silent_trace … t → get_costlabels_of_trace … t = [ ].395 # S #st1 #st2 #t elim t [//] #s1 #s2 #s3 #l #prf #tl #IH * [2: * #ABS @⊥ @ABS % ]541 lemma get_cost_label_silent_is_empty : ∀p.∀S : abstract_status p. 542 ∀st1,st2.∀t : raw_trace p S st1 st2.silent_trace … t → get_costlabels_of_trace … t = [ ]. 543 #p #S #st1 #st2 #t elim t [//] #s1 #s2 #s3 #l #prf #tl #IH * [2: * #ABS @⊥ @ABS % ] 396 544 #H inversion H 397 545 [#s #cl #EQ1 #EQ2 #EQ3 destruct] #s1' #s2' #s3' #prf' #tl' #cl' #Htl #_ #EQ1 #EQ2 #EQ3 … … 399 547 qed. 400 548 401 lemma get_cost_label_invariant_for_append_silent_trace_l :∀ S : abstract_status.402 ∀st1,st2,st3 : S.∀t1 : raw_trace S st1 st2.∀t2 : raw_trace Sst2 st3.549 lemma get_cost_label_invariant_for_append_silent_trace_l :∀p.∀S : abstract_status p. 550 ∀st1,st2,st3 : S.∀t1 : raw_trace … st1 st2.∀t2 : raw_trace … st2 st3. 403 551 silent_trace … t1 → 404 552 get_costlabels_of_trace … (t1 @ t2) = get_costlabels_of_trace … t2. 405 # S #st1 #st2 #st3 #t1 elim t1553 #p #S #st1 #st2 #st3 #t1 elim t1 406 554 [#st #t2 #_ %] #st1' #st2' #st3' #l' #prf #tl #IH #t2 * 407 555 [2: whd in match is_trace_non_empty; normalize nodelta * #ABS @⊥ @ABS %] … … 409 557 #st1'' #st2'' #st3'' #prf'' #tl'' #Hclass #Htl'' #_ #EQ1 #EQ2 #EQ3 destruct 410 558 #_ whd in ⊢ (??%?); >IH [%] %1 assumption 411 qed. 412 413 lemma silent_suffix_cancellable : ∀ S : abstract_status.559 qed. 560 561 lemma silent_suffix_cancellable : ∀p.∀S : abstract_status p. 414 562 ∀s1,s2,s2',s3,s3',s4 : S.∀l,l'. 415 563 ∀t1 : raw_trace … s1 s2. ∀prf : as_execute … l s2 s3. … … 417 565 ∀t1' :raw_trace … s1 s2'.∀prf' : as_execute … l' s2' s3'. 418 566 ∀t2' : raw_trace … s3' s4. 419 is_costlabelled_act l → is_costlabelled_actl' →567 is_costlabelled_act … l → is_costlabelled_act … l' → 420 568 silent_trace … t2 → silent_trace … t2' → 421 569 t1 @ (t_ind … prf t2) = t1' @ (t_ind … prf' t2') → 422 570 s2 = s2' ∧ l = l' ∧ t1 ≃ t1'. 423 # S #s1 #s2 #s2' #s3 #s3' #s4 #l #l' #t1 #prf #t2 #t1' #prf' #t2' #Hl #Hl' #sil_t2 #sil_t2'571 #p #S #s1 #s2 #s2' #s3 #s3' #s4 #l #l' #t1 #prf #t2 #t1' #prf' #t2' #Hl #Hl' #sil_t2 #sil_t2' 424 572 lapply prf -prf lapply prf' -prf' lapply t1' -t1' elim t1 425 573 [ normalize #st * normalize … … 445 593 (*PRE-MEASURABLE TRACES*) 446 594 447 inductive pre_measurable_trace ( S : abstract_status) :448 ∀st1,st2 : S.raw_trace ?st1 st2 → Prop ≝449 | pm_empty : ∀st : S. as_classify … st ≠ cl_io → pre_measurable_trace … (t_base ?st)450 | pm_cons_cost_act : ∀st1,st2,st3 : S.∀l : ActionLabel .451 ∀prf : as_execute S l st1 st2.∀tl : raw_trace Sst2 st3.452 as_classify … st1 ≠ cl_io → is_cost_act l → pre_measurable_trace … tl →595 inductive pre_measurable_trace (p : label_params) (S : abstract_status p) : 596 ∀st1,st2 : S.raw_trace … st1 st2 → Prop ≝ 597 | pm_empty : ∀st : S. as_classify … st ≠ cl_io → pre_measurable_trace … (t_base … st) 598 | pm_cons_cost_act : ∀st1,st2,st3 : S.∀l : ActionLabel p. 599 ∀prf : as_execute … l st1 st2.∀tl : raw_trace … st2 st3. 600 as_classify … st1 ≠ cl_io → is_cost_act … l → pre_measurable_trace … tl → 453 601 pre_measurable_trace … (t_ind … prf … tl) 454 | pm_cons_lab_ret : ∀st1,st2,st3 : S.∀l : ActionLabel .602 | pm_cons_lab_ret : ∀st1,st2,st3 : S.∀l : ActionLabel p. 455 603 as_classify … st1 ≠ cl_io → 456 ∀prf : as_execute S l st1 st2.∀tl : raw_trace Sst2 st3.457 is_labelled_ret_act l → pre_measurable_trace … tl →604 ∀prf : as_execute … l st1 st2.∀tl : raw_trace … st2 st3. 605 is_labelled_ret_act … l → pre_measurable_trace … tl → 458 606 pre_measurable_trace … (t_ind … prf … tl) 459 | pm_cons_post_call : ∀st1,st2,st3 : S.∀l : ActionLabel .460 ∀prf : as_execute S l st1 st2.∀tl : raw_trace Sst2 st3.461 as_classify … st1 ≠ cl_io → is_call_act l → is_call_post_label … st1 →607 | pm_cons_post_call : ∀st1,st2,st3 : S.∀l : ActionLabel p. 608 ∀prf : as_execute … l st1 st2.∀tl : raw_trace … st2 st3. 609 as_classify … st1 ≠ cl_io → is_call_act … l → is_call_post_label … st1 → 462 610 pre_measurable_trace … tl → 463 611 pre_measurable_trace … (t_ind … prf … tl) 464 612 | pm_balanced_call : ∀st1,st2,st3,st4,st5.∀l1,l2. 465 ∀prf : as_execute S l1 st1 st2.∀t1 : raw_trace Sst2 st3.466 ∀t2 : raw_trace S st4 st5.∀prf' : as_execute Sl2 st3 st4.613 ∀prf : as_execute … l1 st1 st2.∀t1 : raw_trace … st2 st3. 614 ∀t2 : raw_trace … st4 st5.∀prf' : as_execute … l2 st3 st4. 467 615 as_classify … st1 ≠ cl_io → as_classify … st3 ≠ cl_io 468 → is_call_act l1 → ¬ is_call_post_label … st1 →616 → is_call_act … l1 → ¬ is_call_post_label … st1 → 469 617 pre_measurable_trace … t1 → pre_measurable_trace … t2 → 470 as_syntactical_succ Sst1 st4 →471 is_unlabelled_ret_act l2 →618 as_syntactical_succ … st1 st4 → 619 is_unlabelled_ret_act … l2 → 472 620 pre_measurable_trace … (t_ind … prf … (t1 @ (t_ind … prf' … t2))). 473 474 lemma pre_measurable_trace_inv : ∀S : abstract_status. 475 ∀st1,st2 : S.∀t : raw_trace … st1 st2. pre_measurable_trace … t → 476 (st1 = st2 ∧ as_classify … st1 ≠ cl_io ∧ t ≃ t_base … st1) ∨ 477 (∃st1' : S.∃l.∃prf : as_execute S l st1 st1'.∃tl. 478 t = t_ind … prf … tl ∧ as_classify … st1 ≠ cl_io ∧ is_cost_act l ∧ 479 pre_measurable_trace … tl) ∨ 480 (∃st1' : S.∃l.∃prf : as_execute S l st1 st1'.∃tl. 481 t = t_ind … prf … tl ∧ 482 as_classify … st1 ≠ cl_io ∧ is_labelled_ret_act l ∧ pre_measurable_trace … tl) ∨ 483 (∃st1' : S .∃l.∃prf : as_execute S l st1 st1'.∃tl. 484 t = t_ind … prf … tl ∧ as_classify … st1 ≠ cl_io ∧ is_call_act l ∧ 485 (bool_to_Prop (is_call_post_label … st1)) ∧ pre_measurable_trace … tl) ∨ 486 (∃st1',st1'',st1''' : S.∃l1,l2.∃prf : as_execute S l1 st1 st1'. 487 ∃t1 : raw_trace S st1' st1''.∃t2 : raw_trace S st1''' st2. 488 ∃prf' : as_execute S l2 st1'' st1'''. 489 t = t_ind … prf … (t1 @ (t_ind … prf' … t2)) ∧ as_classify … st1 ≠cl_io ∧ 490 as_classify … st1'' ≠ cl_io ∧ is_call_act l1 ∧ 491 bool_to_Prop (¬ is_call_post_label … st1) ∧ 492 pre_measurable_trace … t1 ∧ pre_measurable_trace … t2 ∧ 493 as_syntactical_succ S st1 st1''' ∧ is_unlabelled_ret_act l2). 494 #S #st1 #st2 #t #H inversion H 495 [ #st #Hclass #EQ1 #EQ2 destruct #EQ destruct #_ %%%%% // % // 496 | #st1' #st2' #st3' #l #prf #tl #Hst1' #Hl #Htl #_ #EQ1 #EQ2 #EQ3 destruct #_ 497 %%%%2 %{st2'} %{l} %{prf} %{tl} % // % // % // 498 | #st1' #st2' #st3' #l #Hst1 #prf #tl #Hl #Htl #_ #EQ1 #EQ2 #EQ3 destruct #_ 499 %%%2 %{st2'} %{l} %{prf} %{tl} % // % // % // 500 | #st1' #st2' #st3' #l #prf #tl #Hst1 #Hl #H1st1 #Htl #_ #EQ1 #EQ2 #EQ3 destruct #_ 501 % %2 %{st2'} %{l} %{prf} %{tl} % // % // % // % // 502 | #st1' #st2' #st3' #st4' #st5' #l1 #l2 #prf #t1 #t2 #prf' #Hst1' #Hst3' #Hl1 503 #H1st1' #Ht1 #Ht2 #succ #Hl2 #_ #_ #EQ1 #EQ2 #EQ3 destruct #_ %2 504 %{st2'} %{st3'} %{st4'} %{l1} %{(ret_act (None ?))} %{prf} %{t1} %{t2} 505 %{prf'} /12 by conj/ 506 ] 507 qed. 508 509 lemma append_premeasurable_silent : ∀S : abstract_status. 510 ∀st1',st1,st2 : S.∀t : raw_trace S st1 st2.∀t' : raw_trace S st1' st1. 621 622 lemma append_premeasurable : ∀p.∀S : abstract_status p. 623 ∀st1,st2,st3 : S.∀t1 : raw_trace … st1 st2.∀t2: raw_trace … st2 st3. 624 pre_measurable_trace … t1 → pre_measurable_trace … t2 → 625 pre_measurable_trace … (t1 @ t2). 626 #p #S #st1 #st2 #st3 #t1 #t2 #Hpre lapply t2 -t2 elim Hpre -Hpre // 627 [ #s1 #s2 #s3 #l #prf #tl #Hclass #Hcost #pre_tl #IH #t2 #pr3_t2 628 %2 // @IH // 629 | #s1 #s2 #s3 #l #Hclass #prf #tl #ret_l #pre_tl #IH #t2 #pre_t2 %3 // @IH // 630 | #s1 #s2 #s3 #l #prf #tl #Hclass #call #callp #pre_tl #IH #t2 #pre_t2 %4 // @IH // 631 | #s1 #s2 #s3 #s4 #s5 #l1 #l2 #prf1 #tl1 #tl2 #prf2 #Hclass1 #Hclass3 #call_l1 632 #nopost #pre_tl1 #pre_tl2 #succ #unlab #IH1 #IH2 #t2 #pre_t2 633 normalize >append_associative in ⊢ (?????(????????%)); %5 // @IH2 // 634 ] 635 qed. 636 637 lemma append_premeasurable_silent : ∀p.∀S : abstract_status p. 638 ∀st1',st1,st2 : S.∀t : raw_trace … st1 st2.∀t' : raw_trace … st1' st1. 511 639 pre_measurable_trace … t → silent_trace … t' → 512 640 pre_measurable_trace … (t' @ t). 513 # S #st1' #st1 #st2 #t #t' lapply t -t elim t'514 [ #st #t #Hpre #_ whd in ⊢ (???? %); assumption]641 #p #S #st1' #st1 #st2 #t #t' lapply t -t elim t' 642 [ #st #t #Hpre #_ whd in ⊢ (?????%); assumption] 515 643 #s1 #s2 #s3 #l #prf #tl #IH #t #Hpre * [2: * #H @⊥ @H %] #H inversion H in ⊢ ?; 516 644 [ #s #H1 #EQ1 #EQ2 destruct #EQ destruct] 517 645 #s1' #s2' #s3' #prf' #tl' #Hclass #silent_tl' #_ #EQ1 #EQ2 #EQ3 518 destruct #_ whd in ⊢ (???? %); %2646 destruct #_ whd in ⊢ (?????%); %2 519 647 [ assumption 520 648 | %{(None ?)} % … … 524 652 qed. 525 653 526 lemma pre_silent_is_premeasurable : ∀ S : abstract_status.527 ∀st1,st2 : S. ∀t : raw_trace Sst1 st2.pre_silent_trace … t654 lemma pre_silent_is_premeasurable : ∀p.∀S : abstract_status p. 655 ∀st1,st2 : S. ∀t : raw_trace … st1 st2.pre_silent_trace … t 528 656 → pre_measurable_trace … t. 529 # S #st1 #st2 #t elim t657 #p #S #st1 #st2 #t elim t 530 658 [ #st #H inversion H in ⊢ ?; [ #st' #Hst #EQ1 #EQ2 #EQ3 destruct #_ %1 @Hst ] 531 659 #st' #st'' #st''' #prf #tl #Hst'' #pre_tl #_ #EQ1 #EQ2 #EQ3 destruct ] … … 536 664 qed. 537 665 538 lemma append_silent_premeasurable : ∀ S : abstract_status.539 ∀st1',st1,st2 : S.∀t : raw_trace S st1 st2.∀t' : raw_trace Sst1' st1.666 lemma append_silent_premeasurable : ∀p.∀S : abstract_status p. 667 ∀st1',st1,st2 : S.∀t : raw_trace … st1 st2.∀t' : raw_trace … st1' st1. 540 668 pre_measurable_trace … t' → silent_trace … t → 541 669 pre_measurable_trace … (t' @ t). 542 # S #st1' #st1 #st2 #t #t' #Ht' lapply t -t elim Ht'670 #p #S #st1' #st1 #st2 #t #t' #Ht' lapply t -t elim Ht' 543 671 [ #st #Hst #r * [ #H1 @pre_silent_is_premeasurable assumption ] 544 672 inversion r [2: #H1 #H2 #H3 #H4 #H5 #H6 #H7 #H8 #H9 #H10 #H11 cases H11 … … 552 680 #r #pre_r normalize 553 681 >append_tind_commute >append_tind_commute >append_associative 554 <append_tind_commute in ⊢ (???? (??????%)); %5682 <append_tind_commute in ⊢ (?????(???????%)); %5 555 683 try assumption [1,3: whd [ %{f} %{l1'} ] % ] @IH2 assumption 556 684 qed. 557 685 558 lemma head_of_premeasurable_is_not_io : ∀ S : abstract_status.686 lemma head_of_premeasurable_is_not_io : ∀p.∀S : abstract_status p. 559 687 ∀st1,st2 : S. ∀t : raw_trace … st1 st2.pre_measurable_trace … t → 560 as_classify … st1 ≠ cl_io. 688 as_classify … st1 ≠ cl_io. #p 561 689 #S #st1 #st2 #t #H inversion H // 562 690 qed. 563 691 564 lemma get_costlabels_of_trace_empty : ∀ S : abstract_status.∀s1,s2 : S.∀t : raw_trace … s1 s2.692 lemma get_costlabels_of_trace_empty : ∀p.∀S : abstract_status p.∀s1,s2 : S.∀t : raw_trace … s1 s2. 565 693 get_costlabels_of_trace … t = nil ? → pre_measurable_trace … t → silent_trace … t. 566 # S #s1 #s2 #t elim t [ #st #_ #_ %2 % * ] #st1 #st2 #st3 #l #prf #t' #IH #EQ #H -s1 -s2 inversion H694 #p #S #s1 #s2 #t elim t [ #st #_ #_ %2 % * ] #st1 #st2 #st3 #l #prf #t' #IH #EQ #H -s1 -s2 inversion H 567 695 [ #H1 #H2 #H3 #H4 #H5 #H6 destruct 568 696 | #s1 #s2 #s3 #l1 #prf1 #tl #Hclass * #lbl #EQ destruct #pre_meas_tl #_ #EQ1 #EQ2 #EQ3 #EQ4 destruct … … 581 709 (*NO-IO TRACES*) 582 710 583 inductive no_io_trace ( S : abstract_status) : ∀st1,st2 : S.raw_trace … st1 st2 → Prop ≝584 t_base_io :∀st : S.as_classify … st ≠ cl_io → no_io_trace S … (t_base … st)711 inductive no_io_trace (p : label_params) (S : abstract_status p) : ∀st1,st2 : S.raw_trace … st1 st2 → Prop ≝ 712 t_base_io :∀st : S.as_classify … st ≠ cl_io → no_io_trace … S … (t_base … st) 585 713 | t_ind_io : ∀st1,st2,st3 : S.∀l,prf.∀tl : raw_trace … st2 st3.as_classify … st1 ≠ cl_io → 586 714 no_io_trace … tl → no_io_trace … (t_ind … st1 … l prf tl). 587 715 588 lemma no_io_append : ∀ S : abstract_status.716 lemma no_io_append : ∀p.∀S : abstract_status p. 589 717 ∀st1,st2,st3 : S.∀t1 : raw_trace … st1 st2. 590 718 ∀t2 : raw_trace … st2 st3. 591 719 no_io_trace … t1 → no_io_trace … t2 → 592 720 no_io_trace … (t1 @ t2). 593 # S #st1 #st2 #st3 #t1 lapply st3 elim t1721 #p #S #st1 #st2 #st3 #t1 lapply st3 elim t1 594 722 [ #st #st3 #t2 #_ //] 595 723 #s #s' #s'' #l #prf #tl #IH #st3 #t2 #H inversion H in ⊢ ?; … … 599 727 qed. 600 728 601 lemma pre_measurable_no_io : ∀ S : abstract_status.729 lemma pre_measurable_no_io : ∀p.∀S : abstract_status p. 602 730 ∀st1,st2 : S. ∀t : raw_trace … st1 st2. 603 731 pre_measurable_trace … t → 604 732 no_io_trace … t. 605 # S #st1 #st2 #t #H elim H /2/ -st1-st2733 #p #S #st1 #st2 #t #H elim H /2/ -st1-st2 606 734 #st1 #st2 #st3 #st4 #st5 #l1 #l2 #prf1 #t1 #t2 #prf' #H1 #H2 #H3 #H4 #pre1 #pre2 #H5 607 735 #H6 #IH1 #IH2 %2 // @no_io_append // %2 // 608 736 qed. 609 737 610 lemma head_of_no_io_is_no_io : ∀ S : abstract_status.738 lemma head_of_no_io_is_no_io : ∀p.∀S : abstract_status p. 611 739 ∀st1,st2 : S. ∀t : raw_trace … st1 st2. 612 740 no_io_trace … t → as_classify … st1 ≠ cl_io. 613 # S #st1 #st2 #t #H elim H //614 qed. 615 616 617 lemma no_io_append_l1 : ∀ S : abstract_status.741 #p #S #st1 #st2 #t #H elim H // 742 qed. 743 744 745 lemma no_io_append_l1 : ∀p.∀S : abstract_status p. 618 746 ∀st1,st2,st3 : S.∀t1 : raw_trace … st1 st2. 619 747 ∀t2 : raw_trace … st2 st3.no_io_trace … (t1 @ t2) → 620 748 no_io_trace … t1. 621 # S #st1 #st2 #st3 #t1 lapply st3 elim t1 [ #st #st3 #t2 normalize #H % /2/]749 #p #S #st1 #st2 #st3 #t1 lapply st3 elim t1 [ #st #st3 #t2 normalize #H % /2/] 622 750 #s #s' #s'' #l #prf #tl #IH #st3 #t2 #H inversion H in ⊢ ?; 623 751 [ #H9 #H10 #H11 #H12 #H13 #H14 destruct normalize in H13; destruct ] … … 626 754 qed. 627 755 628 lemma no_io_append_l2 : ∀ S : abstract_status.756 lemma no_io_append_l2 : ∀p.∀S : abstract_status p. 629 757 ∀st1,st2,st3 : S.∀t1 : raw_trace … st1 st2. 630 758 ∀t2 : raw_trace … st2 st3.no_io_trace … (t1 @ t2) → 631 759 no_io_trace … t2. 632 # S #st1 #st2 #st3 #t1 lapply st3 elim t1 [ #st #st3 #t2 normalize #H assumption ]760 #p #S #st1 #st2 #st3 #t1 lapply st3 elim t1 [ #st #st3 #t2 normalize #H assumption ] 633 761 #s #s' #s'' #l #prf #tl #IH #st3 #t2 #H @IH inversion H in ⊢ ?; 634 762 [ #H9 #H10 #H11 #H12 #H13 #H14 destruct normalize in H13; destruct ] … … 638 766 (*MEASURABLE TRACES*) 639 767 640 definition measurable : 641 ∀S: abstract_status . ∀si,s1,s3,sn : S. raw_trace … si sn → Prop ≝642 λ S,si,s1,s3,sn,t.768 definition measurable : ∀p. 769 ∀S: abstract_status p. ∀si,s1,s3,sn : S. raw_trace … si sn → Prop ≝ 770 λp,S,si,s1,s3,sn,t. 643 771 ∃s0,s2:S. ∃ti0 : raw_trace … si s0.∃t12 : raw_trace … s1 s2.∃t3n : raw_trace … s3 sn. 644 772 ∃l1,l2,prf1,prf2. 645 pre_measurable_trace … t12 ∧646 t = ti0 @ t_ind ? s0 s1 sn l1 prf1 … (t12 @ (t_ind ?s2 s3 sn l2 prf2 … t3n))647 ∧ is_costlabelled_act l1 ∧ is_costlabelled_act l2 ∧ (is_call_actl2 → bool_to_Prop (is_call_post_label … s2))648 ∧ silent_trace … ti0 ∧ silent_trace … t3n ∧ (is_call_act l1 → bool_to_Prop (is_call_post_label … s0)).773 pre_measurable_trace p … t12 ∧ 774 t = ti0 @ t_ind … s0 s1 sn l1 prf1 … (t12 @ (t_ind … s2 s3 sn l2 prf2 … t3n)) 775 ∧ is_costlabelled_act … l1 ∧ is_costlabelled_act … l2 ∧ (is_call_act … l2 → bool_to_Prop (is_call_post_label … s2)) 776 ∧ silent_trace … ti0 ∧ silent_trace … t3n ∧ (is_call_act … l1 → bool_to_Prop (is_call_post_label … s0)). 649 777 650 778 651 definition measurable_prefix ≝ 652 λS : abstract_status .779 definition measurable_prefix ≝ λp. 780 λS : abstract_status p. 653 781 λs1,s4 :S. 654 782 λt : raw_trace … s1 s4. … … 656 784 ∃l.∃prf : as_execute … l s2 s3. 657 785 ∃t34 : raw_trace … s3 s4. 658 silent_trace … t12 ∧ t = t12 @ (t_ind … prf t34) ∧ is_costlabelled_act l.786 silent_trace … t12 ∧ t = t12 @ (t_ind … prf t34) ∧ is_costlabelled_act … l. 659 787 660 788 lemma measurable_prefix_of_premeasurable : 661 ∀S : abstract_status. 789 ∀p. 790 ∀S : abstract_status p. 662 791 ∀s1,s4 : S. 663 792 ∀t : raw_trace … s1 s4. … … 665 794 get_costlabels_of_trace … t ≠ nil ? → 666 795 measurable_prefix … t. 667 # S #s1 #s4 #t elim t796 #p #S #s1 #s4 #t elim t 668 797 [ #st #_ * #H @⊥ @H %] 669 798 #st1 #st2 #st3 #l @(is_costlabelled_act_elim … l) -l … … 709 838 710 839 definition measurable_suffix ≝ 711 λS : abstract_status.λs1,s1' : S.λt : raw_trace … s1 s1'.840 λp. λS : abstract_status p.λs1,s1' : S.λt : raw_trace … s1 s1'. 712 841 ∃s1_em : S. 713 842 ∃t_pre_em : raw_trace … s1 s1_em. 714 843 ∃s1_postem : S. 715 844 ∃t_post_em : raw_trace … s1_postem s1'. 716 silent_trace S ?? t_post_em ∧717 ∃l_em : ActionLabel .845 silent_trace … S ?? t_post_em ∧ 846 ∃l_em : ActionLabel p. 718 847 ∃ex_em : as_execute … l_em s1_em s1_postem. 719 848 t = t_pre_em @ (t_ind … ex_em t_post_em) ∧ 720 (is_call_act l_em → bool_to_Prop(is_call_post_label … s1_em)) ∧721 is_costlabelled_act l_em.849 (is_call_act … l_em → bool_to_Prop(is_call_post_label … s1_em)) ∧ 850 is_costlabelled_act … l_em. 722 851 723 852 724 lemma measurable_suffix_tind : ∀ S : abstract_status.725 ∀s1,s2,s3 : S.∀l : ActionLabel .∀prf : as_execute … l s1 s2.∀t : raw_trace … s2 s3.726 measurable_suffix … (t_ind … prf t) → (is_costlabelled_act l \wedgesilent_trace … t) ∨ measurable_suffix … t.727 # S #s1 #s2 #s3 #l #prf #t lapply prf -prf lapply s1 -s1 cases t -t -s2 -s3853 lemma measurable_suffix_tind : ∀p.∀S : abstract_status p. 854 ∀s1,s2,s3 : S.∀l : ActionLabel p.∀prf : as_execute … l s1 s2.∀t : raw_trace … s2 s3. 855 measurable_suffix … (t_ind … prf t) → (is_costlabelled_act … l ∧ silent_trace … t) ∨ measurable_suffix … t. 856 #p #S #s1 #s2 #s3 #l #prf #t lapply prf -prf lapply s1 -s1 cases t -t -s2 -s3 728 857 [ #st #st1 #prf * #s_em * #t_pre * #s_post * #t_post * #sil_post * #l_em * #ex_em ** 729 858 inversion t_pre in ⊢ ?; … … 747 876 qed. 748 877 749 lemma measurable_suffix_append : ∀ S : abstract_status.878 lemma measurable_suffix_append : ∀p.∀S : abstract_status p. 750 879 ∀s1,s2,s3 : S.∀t1 : raw_trace … s1 s2.∀t2 : raw_trace … s2 s3. 751 880 measurable_suffix … t2 → measurable_suffix … (t1 @ t2). 752 # S #s1 #s2 #s3 #t1 #t2 * #s_em * #pre_t * #s_post * #t_post881 #p #S #s1 #s2 #s3 #t1 #t2 * #s_em * #pre_t * #s_post * #t_post 753 882 * #sil_tpost * #l_em * #prf ** #EQ destruct #Hcall #Hcost 754 883 %{s_em} %{(t1@pre_t)} %{s_post} %{t_post} /7 by ex_intro, conj/ 755 884 qed. 756 885 757 lemma measurable_suffix_append_l1 : ∀ S : abstract_status.886 lemma measurable_suffix_append_l1 : ∀p.∀S : abstract_status p. 758 887 ∀s1,s2,s3 : S.∀t1 : raw_trace … s1 s2.∀t2 : raw_trace … s2 s3. 759 888 measurable_suffix … (t1 @ t2) → silent_trace … t2 → measurable_suffix … t1. 760 # S #s1 #s2 #s3 #t1 elim t1 -t1 -s1 -s2889 #p #S #s1 #s2 #s3 #t1 elim t1 -t1 -s1 -s2 761 890 [#st #t2 * #s_em * #t_pre * #s_post * #t_post * #sil_tpost * #l_em * #exe ** #EQ normalize in EQ; destruct 762 891 #Hcall #Hcost #H lapply(silent_append_l2 … H) -H * [2: * #H cases(H I)] #H inversion H … … 771 900 | #s #s' #s'' #l' #prf' #tl' #_ #EQ1 #EQ2 #EQ3 destruct * #s_post * #t_post * #sil_tpost * #l_em * #ex_em 772 901 ** #EQ normalize in EQ; lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct #Hcall #Hcost #H 773 change with ((t_ind ?????? (t_base …)) @ tl) in ⊢ (????%); @measurable_suffix_append @(IH t2) //902 change with ((t_ind ??????? (t_base …)) @ tl) in ⊢ (?????%); @measurable_suffix_append @(IH t2) // 774 903 %{s''} %{tl'} %{s_post} %{t_post} %{sil_tpost} %{l_em} %{ex_em} % // % assumption 775 904 ] 776 905 qed. 777 906 778 lemma measurable_suffix_append_case : ∀ S : abstract_status.907 lemma measurable_suffix_append_case : ∀p.∀S : abstract_status p. 779 908 ∀s1,s2,s3 : S. ∀t1 : raw_trace … s1 s2.∀t2 : raw_trace … s2 s3. 780 909 measurable_suffix … (t1 @ t2) → (measurable_suffix … t1 ∧ silent_trace … t2) ∨ measurable_suffix … t2. 781 # S #s1 #s2 #s3 #t1 elim t1 -s1 -s2910 #p #S #s1 #s2 #s3 #t1 elim t1 -s1 -s2 782 911 [ #st #t2 #H %2 assumption] 783 912 #st1 #st2 #st3 #l #prf #tl #IH #t2 whd in match (append_on_trace ??????); #Hsuff 784 913 cases(measurable_suffix_tind … Hsuff) 785 [ * #_ #H %% [2: /2 by silent_append_l2/] change with ((t_ind ?????? tl) @ t2) in Hsuff : (????%);914 [ * #_ #H %% [2: /2 by silent_append_l2/] change with ((t_ind ??????? tl) @ t2) in Hsuff : (?????%); 786 915 @(measurable_suffix_append_l1 … Hsuff) /2/ 787 916 | #H cases(IH … H) 788 [ * #H1 #H2 %% // change with ((t_ind ?????? (t_base …)) @ tl) in ⊢ (????%); @measurable_suffix_append //917 [ * #H1 #H2 %% // change with ((t_ind ??????? (t_base …)) @ tl) in ⊢ (?????%); @measurable_suffix_append // 789 918 | /2/ 790 919 ] … … 792 921 qed. 793 922 794 lemma measurable_is_measurable_suffix : ∀ S : abstract_status.923 lemma measurable_is_measurable_suffix : ∀p.∀S : abstract_status p. 795 924 ∀si,s1,s3,sn : S. ∀t : raw_trace … si sn.measurable …s1 s3 … t → measurable_suffix … t. 796 # S #si #s1 #s3 #sn #t * #s0 * #s2 * #ti0 * #t12 * #t3n * #l1 * #l2 * #prf1 * #prf2 ******* #_925 #p #S #si #s1 #s3 #sn #t * #s0 * #s2 * #ti0 * #t12 * #t3n * #l1 * #l2 * #prf1 * #prf2 ******* #_ 797 926 #EQ destruct /11 width=20 by conj,ex_intro/ 798 927 qed. 799 928 800 lemma prefix_of_measurable_suffix_is_premeasurable : ∀ S : abstract_status.929 lemma prefix_of_measurable_suffix_is_premeasurable : ∀p.∀S : abstract_status p. 801 930 ∀s1,s2,s3,s4 :S.∀l.∀prf : as_execute … l s2 s3. 802 931 ∀t1 : raw_trace … s1 s2.∀t2 : raw_trace … s3 s4.∀t : raw_trace … s1 s4. 803 pre_measurable_trace … t → t = (t1 @ (t_ind … prf t2)) → is_costlabelled_act l →804 silent_trace … t2 → (is_call_act l → bool_to_Prop(is_call_post_label … s2)) →932 pre_measurable_trace … t → t = (t1 @ (t_ind … prf t2)) → is_costlabelled_act … l → 933 silent_trace … t2 → (is_call_act … l → bool_to_Prop(is_call_post_label … s2)) → 805 934 pre_measurable_trace … t1. 806 # S #s1 #s2 #s3 #s4 #l #prf #t1 #t2 #t #H lapply t1 -t1 lapply t2 -t2 lapply prf -prf lapply s2 -s2935 #p #S #s1 #s2 #s3 #s4 #l #prf #t1 #t2 #t #H lapply t1 -t1 lapply t2 -t2 lapply prf -prf lapply s2 -s2 807 936 lapply s3 -s3 elim H -t -s1 -s4 808 937 [ #st #_ #s1 #s2 #prf #r #p inversion p in ⊢ ?; … … 856 985 [ **] -Hmeas -IH1 -EQ * #s_em * #t_pre * #s_post * #t_post * #sil_tpost 857 986 * #l_em * #ex_em ** #EQ destruct #H1 #H2 858 change with (t_ind … (t_base …) @ ?) in match (t_ind ??????? ) in e0;987 change with (t_ind … (t_base …) @ ?) in match (t_ind ????????) in e0; 859 988 <append_associative in e0; <append_associative #e0 860 989 cases(silent_suffix_cancellable … e0) // -e0 * #EQ1 #EQ2 #EQ3 destruct -
LTS/Vm.ma
r3535 r3549 17 17 ; io_instr : Type[0] }. 18 18 19 inductive AssemblerInstr (p : assembler_params) : Type[0] ≝20 | Seq : seq_instr p → option (NonFunctionalLabel ) → AssemblerInstrp21 | Ijmp: ℕ → AssemblerInstr p 22 | CJump : jump_condition p → ℕ → NonFunctionalLabel → NonFunctionalLabel → AssemblerInstrp23 | Iio : NonFunctionalLabel → io_instr p → NonFunctionalLabel → AssemblerInstrp24 | Icall: FunctionName → AssemblerInstr p 25 | Iret: AssemblerInstr p .26 27 definition labels_pc_of_instr : ∀p .AssemblerInstr p → ℕ → list (CostLabel× ℕ) ≝28 λp, i,program_counter.19 inductive AssemblerInstr (p : assembler_params) (l_p : label_params) : Type[0] ≝ 20 | Seq : seq_instr p → option (NonFunctionalLabel l_p) → AssemblerInstr p l_p 21 | Ijmp: ℕ → AssemblerInstr p l_p 22 | CJump : jump_condition p → ℕ → NonFunctionalLabel l_p → NonFunctionalLabel l_p → AssemblerInstr p l_p 23 | Iio : NonFunctionalLabel l_p → io_instr p → NonFunctionalLabel l_p → AssemblerInstr p l_p 24 | Icall: FunctionName → AssemblerInstr p l_p 25 | Iret: AssemblerInstr p l_p. 26 27 definition labels_pc_of_instr : ∀p,l_p.AssemblerInstr p l_p → ℕ → list (CostLabel l_p × ℕ) ≝ 28 λp,l_p,i,program_counter. 29 29 match i with 30 30 [ Seq _ opt_l ⇒ match opt_l with 31 [ Some lbl ⇒ [〈(a_non_functional_label lbl),S program_counter〉]31 [ Some lbl ⇒ [〈(a_non_functional_label … lbl),S program_counter〉] 32 32 | None ⇒ [ ] 33 33 ] 34 34 | Ijmp _ ⇒ [ ] 35 | CJump _ newpc ltrue lfalse ⇒ [〈(a_non_functional_label ltrue),newpc〉;36 〈(a_non_functional_label lfalse),S program_counter〉]37 | Iio lin _ lout ⇒ [〈(a_non_functional_label lin),program_counter〉;38 〈(a_non_functional_label lout),S program_counter〉]35 | CJump _ newpc ltrue lfalse ⇒ [〈(a_non_functional_label … ltrue),newpc〉; 36 〈(a_non_functional_label … lfalse),S program_counter〉] 37 | Iio lin _ lout ⇒ [〈(a_non_functional_label … lin),program_counter〉; 38 〈(a_non_functional_label … lout),S program_counter〉] 39 39 | Icall f ⇒ [ ] 40 40 | Iret ⇒ [ ] 41 41 ]. 42 42 43 let rec labels_pc (p : assembler_params) 44 (prog : list (AssemblerInstr p )) (call_label_fun : list (ℕ × CallCostLabel))45 (return_label_fun : list (ℕ × ReturnPostCostLabel)) (i_act : NonFunctionalLabel)46 (program_counter : ℕ) on prog : list ( CostLabel× ℕ) ≝43 let rec labels_pc (p : assembler_params) (l_p : label_params) 44 (prog : list (AssemblerInstr p l_p)) (call_label_fun : list (ℕ × (CallCostLabel l_p))) 45 (return_label_fun : list (ℕ × (ReturnPostCostLabel l_p))) (i_act : NonFunctionalLabel l_p) 46 (program_counter : ℕ) on prog : list ((CostLabel l_p) × ℕ) ≝ 47 47 match prog with 48 [ nil ⇒ [〈a_non_functional_label (i_act),O〉] @49 map … (λx.let〈y,z〉 ≝ x in 〈(a_call z),y〉) (call_label_fun) @50 map … (λx.let〈y,z〉 ≝ x in 〈(a_return_post z),y〉) (return_label_fun)51 | cons i is ⇒ (labels_pc_of_instr … i program_counter)@labels_pc p is call_label_fun return_label_fun i_act (S program_counter)48 [ nil ⇒ [〈a_non_functional_label … (i_act),O〉] @ 49 map … (λx.let〈y,z〉 ≝ x in 〈(a_call … z),y〉) (call_label_fun) @ 50 map … (λx.let〈y,z〉 ≝ x in 〈(a_return_post … z),y〉) (return_label_fun) 51 | cons i is ⇒ (labels_pc_of_instr … i program_counter)@labels_pc p l_p is call_label_fun return_label_fun i_act (S program_counter) 52 52 ]. 53 53 54 54 include "basics/lists/listb.ma". 55 55 56 (*doppione da mettere a posto*) 57 let rec no_duplicates (A : DeqSet) (l : list A) on l : Prop ≝ 58 match l with 59 [ nil ⇒ True 60 | cons x xs ⇒ ¬ (bool_to_Prop (x ∈ xs)) ∧ no_duplicates … xs 61 ]. 62 63 64 record AssemblerProgram (p : assembler_params) : Type[0] ≝ 65 { instructions : list (AssemblerInstr p) 56 57 record AssemblerProgram (p : assembler_params) (l_p : label_params) : Type[0] ≝ 58 { instructions : list (AssemblerInstr p l_p) 66 59 ; endmain : ℕ 67 60 ; endmain_ok : endmain < |instructions| 68 61 ; entry_of_function : FunctionName → ℕ 69 ; call_label_fun : list (ℕ × CallCostLabel)70 ; return_label_fun : list (ℕ × ReturnPostCostLabel)71 ; in_act : NonFunctionalLabel 62 ; call_label_fun : list (ℕ × (CallCostLabel l_p)) 63 ; return_label_fun : list (ℕ × (ReturnPostCostLabel l_p)) 64 ; in_act : NonFunctionalLabel l_p 72 65 ; asm_no_duplicates : no_duplicates … (map ?? \fst … (labels_pc … instructions call_label_fun return_label_fun in_act O)) 73 66 }. 74 67 75 68 76 definition fetch: ∀p .AssemblerProgram p → ℕ → option (AssemblerInstrp) ≝77 λp,l ,n. nth_opt ? n (instructions … l).69 definition fetch: ∀p,l_p.AssemblerProgram p l_p → ℕ → option (AssemblerInstr p l_p) ≝ 70 λp,l_p,l,n. nth_opt ? n (instructions … l). 78 71 79 72 definition stackT: Type[0] ≝ list (nat). … … 102 95 103 96 104 inductive vmstep (p : assembler_params) (p' : sem_params p) 105 (prog : AssemblerProgram p ) :106 ActionLabel → relation (vm_state p p') ≝97 inductive vmstep (p : assembler_params) (p' : sem_params p) (l_p : label_params) 98 (prog : AssemblerProgram p l_p) : 99 ActionLabel l_p → relation (vm_state p p') ≝ 107 100 | vm_seq : ∀st1,st2 : vm_state p p'.∀i,l. 108 fetch … prog (pc … st1) = return (Seq p i l) →101 fetch … prog (pc … st1) = return (Seq p l_p i l) → 109 102 asm_is_io … st1 = false → 110 103 eval_asm_seq p p' i (asm_store … st1) = return asm_store … st2 → … … 112 105 asm_is_io … st1 = asm_is_io … st2 → 113 106 S (pc … st1) = pc … st2 → 114 vmstep … (cost_act l) st1 st2107 vmstep … (cost_act … l) st1 st2 115 108 | vm_ijump : ∀st1,st2 : vm_state p p'.∀new_pc : ℕ. 116 fetch … prog (pc … st1) = return (Ijmp p new_pc) →109 fetch … prog (pc … st1) = return (Ijmp p l_p new_pc) → 117 110 asm_is_io … st1 = false → 118 111 asm_store … st1 = asm_store … st2 → … … 120 113 asm_is_io … st1 = asm_is_io … st2 → 121 114 new_pc = pc … st2 → 122 vmstep … (cost_act (None ?)) st1 st2115 vmstep … (cost_act … (None ?)) st1 st2 123 116 | vm_cjump_true : 124 117 ∀st1,st2 : vm_state p p'.∀cond,new_pc,ltrue,lfalse. 125 118 eval_asm_cond p p' cond (asm_store … st1) = return true→ 126 fetch … prog (pc … st1) = return (CJump p cond new_pc ltrue lfalse) →119 fetch … prog (pc … st1) = return (CJump p l_p cond new_pc ltrue lfalse) → 127 120 asm_is_io … st1 = false → 128 121 asm_store … st1 = asm_store … st2 → … … 130 123 asm_is_io … st1 = asm_is_io … st2 → 131 124 pc … st2 = new_pc → 132 vmstep … (cost_act (Some ? ltrue)) st1 st2125 vmstep … (cost_act … (Some ? ltrue)) st1 st2 133 126 | vm_cjump_false : 134 127 ∀st1,st2 : vm_state p p'.∀cond,new_pc,ltrue,lfalse. 135 128 eval_asm_cond p p' cond (asm_store … st1) = return false→ 136 fetch … prog (pc … st1) = return (CJump p cond new_pc ltrue lfalse) →129 fetch … prog (pc … st1) = return (CJump p l_p cond new_pc ltrue lfalse) → 137 130 asm_is_io … st1 = false → 138 131 asm_store … st1 = asm_store … st2 → … … 140 133 asm_is_io … st1 = asm_is_io … st2 → 141 134 S (pc … st1) = pc … st2 → 142 vmstep … (cost_act (Some ? lfalse)) st1 st2135 vmstep … (cost_act … (Some ? lfalse)) st1 st2 143 136 | vm_io_in : 144 137 ∀st1,st2 : vm_state p p'.∀lin,io,lout. 145 fetch … prog (pc … st1) = return (Iio p l in io lout) →138 fetch … prog (pc … st1) = return (Iio p l_p lin io lout) → 146 139 asm_is_io … st1 = false → 147 140 asm_store … st1 = asm_store … st2 → … … 149 142 true = asm_is_io … st2 → 150 143 pc … st1 = pc … st2 → 151 vmstep … (cost_act (Some ? lin)) st1 st2144 vmstep … (cost_act … (Some ? lin)) st1 st2 152 145 | vm_io_out : 153 146 ∀st1,st2 : vm_state p p'.∀lin,io,lout. 154 fetch … prog (pc … st1) = return (Iio p l in io lout) →147 fetch … prog (pc … st1) = return (Iio p l_p lin io lout) → 155 148 asm_is_io … st1 = true → 156 149 eval_asm_io … io (asm_store … st1) = return asm_store … st2 → … … 158 151 false = asm_is_io … st2 → 159 152 S (pc … st1) = pc … st2 → 160 vmstep … (cost_act (Some ? lout)) st1 st2153 vmstep … (cost_act … (Some ? lout)) st1 st2 161 154 | vm_call : 162 155 ∀st1,st2 : vm_state p p'.∀f,lbl. 163 fetch … prog (pc … st1) = return (Icall p f) →156 fetch … prog (pc … st1) = return (Icall p l_p f) → 164 157 asm_is_io … st1 = false → 165 158 asm_store … st1 = asm_store … st2 → … … 167 160 asm_is_io … st1 = asm_is_io … st2 → 168 161 entry_of_function … prog f = pc … st2 → 169 label_of_pc ?(call_label_fun … prog) (entry_of_function … prog f) = return lbl →170 vmstep … (call_act f lbl) st1 st2162 label_of_pc … (call_label_fun … prog) (entry_of_function … prog f) = return lbl → 163 vmstep … (call_act … f lbl) st1 st2 171 164 | vm_ret : 172 165 ∀st1,st2 : vm_state p p'.∀newpc,lbl. 173 fetch … prog (pc … st1) = return (Iret p ) →166 fetch … prog (pc … st1) = return (Iret p l_p) → 174 167 asm_is_io … st1 = false → 175 168 asm_store … st1 = asm_store … st2 → … … 177 170 asm_is_io … st1 = asm_is_io … st2 → 178 171 newpc = pc … st2 → 179 label_of_pc ?(return_label_fun … prog) newpc = return lbl →180 vmstep … (ret_act (Some ? lbl)) st1 st2.181 182 definition eval_vmstate : ∀p : assembler_params.∀p' : sem_params p. 183 AssemblerProgram p → vm_state p p' → option (ActionLabel× (vm_state p p')) ≝184 λp,p', prog,st.172 label_of_pc … (return_label_fun … prog) newpc = return lbl → 173 vmstep … (ret_act … (Some ? lbl)) st1 st2. 174 175 definition eval_vmstate : ∀p : assembler_params.∀p' : sem_params p.∀l_p : label_params. 176 AssemblerProgram p l_p → vm_state p p' → option ((ActionLabel l_p) × (vm_state p p')) ≝ 177 λp,p',l_p,prog,st. 185 178 ! i ← fetch … prog (pc … st); 186 179 match i with … … 190 183 else 191 184 ! new_store ← eval_asm_seq p p' x (asm_store … st); 192 return 〈cost_act opt_l,185 return 〈cost_act … opt_l, 193 186 mk_vm_state ?? (S (pc … st)) (asm_stack … st) new_store false〉 194 187 | Ijmp newpc ⇒ … … 196 189 None ? 197 190 else 198 return 〈cost_act (None ?),191 return 〈cost_act … (None ?), 199 192 mk_vm_state ?? newpc (asm_stack … st) (asm_store … st) false〉 200 193 | CJump cond newpc ltrue lfalse ⇒ … … 204 197 ! b ← eval_asm_cond p p' cond (asm_store … st); 205 198 if b then 206 return 〈cost_act (Some ? ltrue),199 return 〈cost_act …(Some ? ltrue), 207 200 mk_vm_state ?? newpc (asm_stack … st) (asm_store … st) false〉 208 201 else 209 return 〈cost_act (Some ? lfalse),202 return 〈cost_act … (Some ? lfalse), 210 203 mk_vm_state ?? (S (pc … st)) (asm_stack … st) (asm_store … st) false〉 211 204 | Iio lin io lout ⇒ 212 205 if asm_is_io … st then 213 206 ! new_store ← eval_asm_io … io (asm_store … st); 214 return 〈cost_act (Some ? lout),207 return 〈cost_act … (Some ? lout), 215 208 mk_vm_state ?? (S (pc … st)) (asm_stack … st) 216 209 new_store false〉 217 210 else 218 return 〈cost_act (Some ? lin),211 return 〈cost_act … (Some ? lin), 219 212 mk_vm_state ?? (pc … st) (asm_stack … st) (asm_store … st) 220 213 true〉 … … 224 217 else 225 218 ! lbl ← label_of_pc ? (call_label_fun … prog) (entry_of_function … prog f); 226 return 〈call_act f lbl,219 return 〈call_act … f lbl, 227 220 mk_vm_state ?? (entry_of_function … prog f) 228 221 ((S (pc … st)) :: (asm_stack … st)) … … 233 226 ! 〈newpc,tl〉 ← option_pop … (asm_stack … st); 234 227 ! lbl ← label_of_pc ? (return_label_fun … prog) newpc; 235 return 〈ret_act (Some ? lbl),228 return 〈ret_act … (Some ? lbl), 236 229 mk_vm_state ?? newpc tl (asm_store … st) false〉 237 230 ]. 238 231 239 lemma eval_vmstate_to_Prop : ∀p,p', prog,st1,st2,l.240 eval_vmstate p p' prog st1 = return 〈l,st2〉 → vmstep … prog l st1 st2.241 #p #p' # prog #st1 #st2 #l whd in match eval_vmstate; normalize nodelta232 lemma eval_vmstate_to_Prop : ∀p,p',l_p,prog,st1,st2,l. 233 eval_vmstate p p' l_p prog st1 = return 〈l,st2〉 → vmstep … prog l st1 st2. 234 #p #p' #l_p #prog #st1 #st2 #l whd in match eval_vmstate; normalize nodelta 242 235 #H cases(bind_inversion ????? H) -H * normalize nodelta 243 236 [ #seq #opt_l * #EQfetch inversion(asm_is_io ???) normalize nodelta … … 269 262 270 263 271 lemma vm_step_to_eval : ∀p,p', prog,st1,st2,l.vmstep … prog l st1 st2 →272 eval_vmstate p p' prog st1 = return 〈l,st2〉.273 #p #p' # prog * #pc1 #stack1 #store1 #io1264 lemma vm_step_to_eval : ∀p,p',l_p,prog,st1,st2,l.vmstep … prog l st1 st2 → 265 eval_vmstate p p' l_p prog st1 = return 〈l,st2〉. 266 #p #p' #l_p #prog * #pc1 #stack1 #store1 #io1 274 267 * #pc2 #stack2 #store2 #io2 #l #H inversion H 275 268 [ #s1 #s2 #i #opt_l #EQfetch #EQio #EQstore #EQstack #EQio1 #EQpc … … 316 309 317 310 definition ass_vmstep ≝ 318 λp,p', prog.311 λp,p',l_p,prog. 319 312 λl.λs1,s2 : vm_ass_state p p'. 320 313 match s1 with … … 322 315 match s2 with 323 316 [ STATE st2 ⇒ 324 (eqb (pc ?? st1) (endmain … prog)) = false ∧ vmstep p p' prog l st1 st2317 (eqb (pc ?? st1) (endmain … prog)) = false ∧ vmstep p p' l_p prog l st1 st2 325 318 | INITIAL ⇒ False 326 319 | FINAL ⇒ eqb (pc … st1) (endmain … prog) = true ∧ 327 l = cost_act (Some … (in_act … prog))320 l = cost_act … (Some … (in_act … prog)) 328 321 ] 329 322 | INITIAL ⇒ match s2 with 330 323 [ STATE st2 ⇒ eqb (pc … st2) O = true ∧ 331 l = cost_act (Some … (in_act … prog))324 l = cost_act … (Some … (in_act … prog)) 332 325 | _ ⇒ False 333 326 ] … … 335 328 ]. 336 329 337 definition asm_operational_semantics : ∀p .sem_params p → AssemblerProgram p → abstract_status≝338 λp, p',prog.let init_act ≝ cost_act(Some ? (in_act … prog)) in339 let end_act ≝ cost_act (Some ? (in_act … prog)) in340 mk_abstract_status 330 definition asm_operational_semantics : ∀p,l_p.sem_params p → AssemblerProgram p l_p → abstract_status l_p ≝ 331 λp,l_p,p',prog.let init_act ≝ cost_act … (Some ? (in_act … prog)) in 332 let end_act ≝ cost_act … (Some ? (in_act … prog)) in 333 mk_abstract_status … 341 334 (vm_ass_state p p') 342 335 (ass_vmstep … prog) … … 427 420 428 421 definition asm_concrete_trans_sys ≝ 429 λp,p', prog.mk_concrete_transition_sys …430 (asm_operational_semantics p p' prog).422 λp,p',l_p,prog.mk_concrete_transition_sys … 423 (asm_operational_semantics p p' l_p prog). 431 424 432 425 433 426 definition emits_labels ≝ 434 λp .λinstr : AssemblerInstrp.match instr with427 λp,l_p.λinstr : AssemblerInstr p l_p.match instr with 435 428 [ Seq _ opt_l ⇒ match opt_l with 436 429 [ None ⇒ Some ? (λpc.S pc) … … 441 434 ]. 442 435 443 definition fetch_state : ∀p,p' .AssemblerProgram p → vm_state p p' → option (AssemblerInstrp) ≝444 λp,p', prog,st.fetch … prog (pc … st).445 446 record asm_galois_connection (p: assembler_params) (p': sem_params p) ( prog: AssemblerProgramp) : Type[2] ≝436 definition fetch_state : ∀p,p',l_p.AssemblerProgram p l_p → vm_state p p' → option (AssemblerInstr p l_p) ≝ 437 λp,p',l_p,prog,st.fetch … prog (pc … st). 438 439 record asm_galois_connection (p: assembler_params) (p': sem_params p) (l_p : label_params) (prog: AssemblerProgram p l_p) : Type[2] ≝ 447 440 { aabs_d : abstract_transition_sys 448 ; agalois_rel:> galois_rel (asm_concrete_trans_sysp p' prog) aabs_d441 ; agalois_rel:> galois_rel … (asm_concrete_trans_sys p l_p p' prog) aabs_d 449 442 }. 450 443 451 444 definition galois_connection_of_asm_galois_connection: 452 ∀p,p', prog. asm_galois_connection p p' prog → galois_connection≝453 λp,p', prog,agc.454 mk_galois_connection 455 (asm_concrete_trans_sys p p' prog)445 ∀p,p',l_p,prog. asm_galois_connection p p' l_p prog → galois_connection l_p ≝ 446 λp,p',l_p,prog,agc. 447 mk_galois_connection … 448 (asm_concrete_trans_sys p l_p p' prog) 456 449 (aabs_d … agc) 457 450 (agalois_rel … agc). … … 460 453 461 454 definition ass_fetch ≝ 462 λp,p', prog.463 λs.match s with [ STATE st ⇒ if eqb (pc … st) (endmain …prog) then455 λp,p',l_p,prog. 456 λs.match s with [ STATE st ⇒ if eqb (pc … st) (endmain p l_p prog) then 464 457 Some ? (None ?) 465 else ! x ← fetch_state p p' prog st; Some ? (Some ? x)458 else ! x ← fetch_state p p' l_p prog st; Some ? (Some ? x) 466 459 | INITIAL ⇒ Some ? (None ?) 467 460 | FINAL ⇒ None ? ]. 468 461 469 462 definition ass_instr_map ≝ 470 λp,p', prog.λasm_galois_conn: asm_galois_connection p p'prog.471 λinstr_map: AssemblerInstr p → (*option*) (abs_instr … (abs_dasm_galois_conn)).463 λp,p',l_p,prog.λasm_galois_conn: asm_galois_connection p p' l_p prog. 464 λinstr_map: AssemblerInstr p l_p → (*option*) (abs_instr … (abs_d … asm_galois_conn)). 472 465 (λi.match i with [None ⇒ (*Some …*) (e …) |Some x ⇒ instr_map … x]). 473 466 474 record asm_abstract_interpretation (p: assembler_params) (p': sem_params p) ( prog: AssemblerProgramp) : Type[2] ≝475 { asm_galois_conn: asm_galois_connection p p' prog476 ; instr_map : AssemblerInstr p → (*option*) (abs_instr … (abs_dasm_galois_conn))467 record asm_abstract_interpretation (p: assembler_params) (p': sem_params p) (l_p : label_params) (prog: AssemblerProgram p l_p) : Type[2] ≝ 468 { asm_galois_conn: asm_galois_connection p p' l_p prog 469 ; instr_map : AssemblerInstr p l_p → (*option*) (abs_instr … (abs_d … asm_galois_conn)) 477 470 ; instr_map_ok : 478 471 ∀s,s': concr … asm_galois_conn. ∀a: abs_d … asm_galois_conn.∀l,i. … … 484 477 485 478 definition abstract_interpretation_of_asm_abstract_interpretation: 486 ∀p,p', prog. asm_abstract_interpretation p p' prog → abstract_interpretation479 ∀p,p',l_p,prog. asm_abstract_interpretation p p' l_p prog → abstract_interpretation l_p 487 480 ≝ 488 λp,p', prog,aai.489 mk_abstract_interpretation 490 (asm_galois_conn … aai) (option (AssemblerInstr p )) (ass_fetch p p'prog)481 λp,p',l_p,prog,aai. 482 mk_abstract_interpretation … 483 (asm_galois_conn … aai) (option (AssemblerInstr p l_p)) (ass_fetch p p' l_p prog) 491 484 (ass_instr_map … prog … (instr_map … aai)) (instr_map_ok … aai). 492 485 … … 496 489 λA,l.match l with [ nil ⇒ false | _ ⇒ true ]. 497 490 498 let rec block_cost (p : assembler_params) 499 (prog: AssemblerProgram p ) (abs_t : monoid)500 (instr_m : AssemblerInstr p → abs_t)491 let rec block_cost (p : assembler_params) (l_p : label_params) 492 (prog: AssemblerProgram p l_p) (abs_t : monoid) 493 (instr_m : AssemblerInstr p l_p → abs_t) 501 494 (prog_c: option ℕ) 502 495 (program_size: ℕ) … … 532 525 533 526 534 lemma labels_pc_ok : ∀p, prog,l1,l2,i_act,i,lbl,pc,m.527 lemma labels_pc_ok : ∀p,l_p,prog,l1,l2,i_act,i,lbl,pc,m. 535 528 nth_opt ? pc prog = return i → 536 529 mem ? lbl (labels_pc_of_instr … i (m+pc)) → 537 mem ? lbl (labels_pc p prog l1 l2 i_act m).538 #p # instrs #l1 #l2 #iact #i #lbl #pc530 mem ? lbl (labels_pc p l_p prog l1 l2 i_act m). 531 #p #l_p #instrs #l1 #l2 #iact #i #lbl #pc 539 532 whd in match fetch; normalize nodelta lapply pc -pc 540 533 elim instrs … … 547 540 qed. 548 541 549 lemma labels_pf_in_act: ∀p, prog,pc.550 mem CostLabel (in_actp prog)551 (map ( CostLabel×ℕ) CostLabel\fst552 (labels_pc p (instructions p prog) (call_label_fun pprog)553 (return_label_fun p prog) (in_act pprog) pc)).554 #p # prog elim (instructions pprog) normalize /2/555 qed. 556 557 lemma labels_pc_return: ∀p, prog,l1,l2,iact,x1,x2.558 label_of_pc ReturnPostCostLabell2 x1=return x2 →542 lemma labels_pf_in_act: ∀p,l_p,prog,pc. 543 mem (CostLabel l_p) (in_act p l_p prog) 544 (map ((CostLabel l_p)×ℕ) (CostLabel l_p) \fst 545 (labels_pc p … (instructions p … prog) (call_label_fun p … prog) 546 (return_label_fun p … prog) (in_act p … prog) pc)). 547 #p #l_p #prog elim (instructions … prog) normalize /2/ 548 qed. 549 550 lemma labels_pc_return: ∀p,l_p,prog,l1,l2,iact,x1,x2. 551 label_of_pc (ReturnPostCostLabel l_p) l2 x1=return x2 → 559 552 ∀m. 560 mem … 〈(a_return_post x2),x1〉 (labels_pcp prog l1 l2 iact m).561 #p #l #l1 #l2 #iact whd in match (labels_pc ???); #x1 #x2 #H elim l553 mem … 〈(a_return_post … x2),x1〉 (labels_pc p l_p prog l1 l2 iact m). 554 #p #l_p #l #l1 #l2 #iact whd in match (labels_pc ???); #x1 #x2 #H elim l 562 555 [ #m @mem_append_l2 @mem_append_l2 whd in H:(??%?); 563 556 elim l2 in H; [ whd in ⊢ (??%% → ?); #EQ destruct] … … 569 562 qed. 570 563 571 lemma labels_pc_call: ∀p, prog,l1,l2,iact,x1,x2.572 label_of_pc CallCostLabell1 x1=return x2 →564 lemma labels_pc_call: ∀p,l_p,prog,l1,l2,iact,x1,x2. 565 label_of_pc (CallCostLabel l_p) l1 x1=return x2 → 573 566 ∀m. 574 mem … 〈(a_call x2),x1〉 (labels_pcp prog l1 l2 iact m).575 #p #l #l1 #l2 #iact whd in match (labels_pc ???); #x1 #x2 #H elim l567 mem … 〈(a_call l_p x2),x1〉 (labels_pc p l_p prog l1 l2 iact m). 568 #p #l_p #l #l1 #l2 #iact whd in match (labels_pc ???); #x1 #x2 #H elim l 576 569 [ #m @mem_append_l2 @mem_append_l1 whd in H:(??%?); 577 570 elim l1 in H; [ whd in ⊢ (??%% → ?); #EQ destruct] … … 609 602 ]. 610 603 611 definition static_analisys : ∀p : assembler_params.∀ abs_t : monoid.(AssemblerInstr p → abs_t) →612 ∀mT : cost_map_T DEQCostLabel abs_t.AssemblerProgramp → option mT ≝613 λp, abs_t,instr_m,mT,prog.604 definition static_analisys : ∀p : assembler_params.∀l_p : label_params.∀abs_t : monoid.(AssemblerInstr p l_p → abs_t …) → 605 ∀mT : cost_map_T (DEQCostLabel l_p) (abs_t …).AssemblerProgram p l_p → option mT ≝ 606 λp,l_p,abs_t,instr_m,mT,prog. 614 607 let prog_size ≝ S (|instructions … prog|) in 615 m_foldr Option ?? (λx,m.let 〈z,w〉≝ x in ! k ← block_cost p prog abs_t instr_m (Some ? w) prog_size;608 m_foldr Option ?? (λx,m.let 〈z,w〉≝ x in ! k ← block_cost p l_p prog abs_t instr_m (Some ? w) prog_size; 616 609 return set_map … m z k) (labels_pc … (instructions … prog) 617 610 (call_label_fun … prog) (return_label_fun … prog) (in_act … prog) O) 618 611 (empty_map ?? mT). 619 612 620 621 (*falso: necessita di no_duplicates*)622 613 623 614 definition eq_deq_prod : ∀D1,D2 : DeqSet.D1 × D2 → D1 × D2 → bool ≝ … … 658 649 qed. 659 650 660 lemma static_analisys_ok : ∀p, abs_t,instr_m,mT,prog,lbl,pc,map.661 static_analisys p abs_t instr_m mT prog = return map →651 lemma static_analisys_ok : ∀p,l_p,abs_t,instr_m,mT,prog,lbl,pc,map. 652 static_analisys p l_p abs_t instr_m mT prog = return map → 662 653 mem … 〈lbl,pc〉 (labels_pc … (instructions … prog) (call_label_fun … prog) 663 654 (return_label_fun … prog) (in_act … prog) O) → … … 665 656 block_cost … prog abs_t instr_m (Some ? pc) (S (|(instructions … prog)|)) ∧ 666 657 block_cost … prog abs_t instr_m (Some ? pc) (S (|(instructions … prog)|)) ≠ None ?. 667 #p # abs_t #instr_m #mT * #prog whd in match static_analisys; normalize nodelta658 #p #l_p #abs_t #instr_m #mT * #prog whd in match static_analisys; normalize nodelta 668 659 #endmain #Hendmain #entry_fun #call_label_fun #return_label_fun #inact 669 660 #nodup generalize in match nodup in ⊢ (∀_.∀_.∀_. (??(????%??)?) → %); #Hnodup lapply nodup -nodup 670 lapply (labels_pc ?????? ) #l elim l [ #x #y #z #w #h * ]661 lapply (labels_pc ???????) #l elim l [ #x #y #z #w #h * ] 671 662 * #hd1 #hd2 #tl #IH * #H1 #H2 #lbl #pc #map #H 672 663 cases(bind_inversion ????? H) -H #map1 * #EQmap1 normalize nodelta #H … … 683 674 qed. 684 675 685 definition terminated_trace : ∀ S : abstract_status.∀s1,s3 : S.raw_trace … s1 s3 → Prop ≝686 λ S,s1,s3,t.∃s2: S.∃t1 : raw_trace … s1 s2.∃ l,prf.t = t1 @ (t_ind? s2 s3 s3 l prf … (t_base … s3))687 ∧ is_costlabelled_act l ∧ pre_measurable_trace … t1.676 definition terminated_trace : ∀p : label_params.∀S : abstract_status p.∀s1,s3 : S.raw_trace … s1 s3 → Prop ≝ 677 λp,S,s1,s3,t.∃s2: S.∃t1 : raw_trace … s1 s2.∃ l,prf.t = t1 @ (t_ind p ? s2 s3 s3 l prf … (t_base … s3)) 678 ∧ is_costlabelled_act … l ∧ pre_measurable_trace … t1. 688 679 689 680 definition big_op: ∀M: monoid. list M → M ≝ … … 716 707 qed. 717 708 718 lemma monotonicity_of_block_cost : ∀p, prog,abs_t,instr_m,pc,size,k.719 block_cost p prog abs_t instr_m (Some ? pc) size = return k →709 lemma monotonicity_of_block_cost : ∀p,l_p,prog,abs_t,instr_m,pc,size,k. 710 block_cost p l_p prog abs_t instr_m (Some ? pc) size = return k → 720 711 ∀size'.size ≤ size' → 721 block_cost p prog abs_t instr_m (Some ? pc) size' = return k.722 #p # prog #abs_t #instr_m #pc #size lapply pc elim size712 block_cost p l_p prog abs_t instr_m (Some ? pc) size' = return k. 713 #p #l_p #prog #abs_t #instr_m #pc #size lapply pc elim size 723 714 [ #pc #k whd in ⊢ (??%% → ?); #EQ destruct] 724 715 #n #IH #pc #k whd in ⊢ (??%% → ?); @eqb_elim … … 741 732 qed. 742 733 743 lemma step_emit : ∀p,p', prog,st1,st2,l,i.744 fetch p prog (pc … st1) = return i →734 lemma step_emit : ∀p,p',l_p,prog,st1,st2,l,i. 735 fetch p l_p prog (pc … st1) = return i → 745 736 eval_vmstate p p' … prog st1 = return 〈l,st2〉 → 746 737 emits_labels … i = None ? → ∃x. 747 match l in ActionLabel return λ_: ActionLabel.(list CostLabel) with748 [call_act f c ⇒ [a_call c]738 match l in ActionLabel return λ_: (ActionLabel l_p).(list (CostLabel l_p)) with 739 [call_act f c ⇒ [a_call … c] 749 740 |ret_act x ⇒ 750 match x with [None⇒[]|Some c⇒[a_return_post c]]741 match x with [None⇒[]|Some c⇒[a_return_post … c]] 751 742 |cost_act x ⇒ 752 match x with [None⇒[]|Some c⇒[a_non_functional_label c]]743 match x with [None⇒[]|Some c⇒[a_non_functional_label … c]] 753 744 ] = [x] ∧ 754 (mem … 〈x,pc … st2〉 (labels_pc p (instructions … prog) (call_label_fun … prog) (return_label_fun … prog) (in_act … prog) O)).755 #p #p' # prog #st1 #st2 #l #i #EQi whd in match eval_vmstate; normalize nodelta745 (mem … 〈x,pc … st2〉 (labels_pc p l_p (instructions … prog) (call_label_fun … prog) (return_label_fun … prog) (in_act … prog) O)). 746 #p #p' #l_p #prog #st1 #st2 #l #i #EQi whd in match eval_vmstate; normalize nodelta 756 747 >EQi >m_return_bind normalize nodelta cases i in EQi; -i normalize nodelta 757 748 [ #seq * [|#lab] … … 775 766 qed. 776 767 777 lemma step_non_emit : ∀p,p', prog,st1,st2,l,i,f.778 fetch p prog (pc … st1) = return i →768 lemma step_non_emit : ∀p,p',l_p,prog,st1,st2,l,i,f. 769 fetch p l_p prog (pc … st1) = return i → 779 770 eval_vmstate p p' … prog st1 = return 〈l,st2〉 → 780 771 emits_labels … i = Some ? f → 781 match l in ActionLabel return λ_: ActionLabel.(list CostLabel) with782 [call_act f c ⇒ [a_call c]772 match l in ActionLabel return λ_: (ActionLabel l_p).(list (CostLabel l_p)) with 773 [call_act f c ⇒ [a_call … c] 783 774 |ret_act x ⇒ 784 match x with [None⇒[]|Some c⇒[a_return_post c]]775 match x with [None⇒[]|Some c⇒[a_return_post … c]] 785 776 |cost_act x ⇒ 786 match x with [None⇒[]|Some c⇒[a_non_functional_label c]]777 match x with [None⇒[]|Some c⇒[a_non_functional_label … c]] 787 778 ] = [ ] ∧ pc … st2 = f (pc … st1). 788 #p #p' # prog #st1 #st2 #l #i #f #EQi whd in match eval_vmstate; normalize nodelta779 #p #p' #l_p #prog #st1 #st2 #l #i #f #EQi whd in match eval_vmstate; normalize nodelta 789 780 >EQi >m_return_bind normalize nodelta cases i in EQi; -i normalize nodelta 790 781 [ #seq * [|#lab] … … 807 798 808 799 lemma labels_of_trace_are_in_code : 809 ∀p,p', prog.∀st1,st2 : vm_ass_state p p'.∀t : raw_trace (asm_operational_semanticsp p' prog) … st1 st2.800 ∀p,p',l_p,prog.∀st1,st2 : vm_ass_state p p'.∀t : raw_trace … (asm_operational_semantics p l_p p' prog) … st1 st2. 810 801 ∀lbl. 811 802 mem … lbl (get_costlabels_of_trace … t) → 812 mem … lbl (map … \fst … (labels_pc … (instructions p prog) (call_label_fun … prog) (return_label_fun … prog) (in_act … prog) O)).813 #p #p' # prog #st1 #st2 #t elim t803 mem … lbl (map … \fst … (labels_pc … (instructions p … prog) (call_label_fun … prog) (return_label_fun … prog) (in_act … prog) O)). 804 #p #p' #l_p #prog #st1 #st2 #t elim t 814 805 [ #st #lbl * 815 806 | #st1 #st2 #st3 #l whd in ⊢ (% → ?); … … 862 853 863 854 864 lemma tbase_tind_append : ∀ S : abstract_status.∀st1,st2,st3.∀t : raw_trace … st1 st2.855 lemma tbase_tind_append : ∀p : label_params.∀S : abstract_status p.∀st1,st2,st3 : S.∀t : raw_trace … st1 st2. 865 856 ∀l,prf.∀t'. 866 t_base … st1 = t @ t_ind S st2 st3 st1 l prf t' → False.867 # S #st1 #st2 #st3 *857 t_base … st1 = t @ t_ind … S st2 st3 st1 l prf t' → False. 858 #p #S #st1 #st2 #st3 * 868 859 [ #st #l #prf #t' normalize #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct ] 869 860 #s1 #s2 #s3 #l1 #prf1 #t2 #l2 #prf2 #t3 normalize #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct … … 893 884 *) 894 885 895 lemma get_cost_label_of_trace_tind : ∀ S : abstract_status.886 lemma get_cost_label_of_trace_tind : ∀p. ∀S : abstract_status p. 896 887 ∀st1,st2,st3 : S.∀l,prf,t. 897 888 get_costlabels_of_trace … (t_ind … st1 st2 st3 l prf t) = 898 actionlabel_to_costlabel l @ get_costlabels_of_trace … t.899 # S #st1 #st2 #st3 * // qed.889 actionlabel_to_costlabel … l @ get_costlabels_of_trace … t. 890 #p #S #st1 #st2 #st3 * // qed. 900 891 901 892 lemma actionlabel_ok : 902 ∀ l : ActionLabel.903 is_costlabelled_act l → ∃c.actionlabel_to_costlabell = [c].904 * /2/ * /2/ *905 qed. 906 907 lemma i_act_in_map : ∀p, prog,iact,l1,l2.908 mem ? 〈a_non_functional_label iact,O〉 (labels_pcp prog l1 l2 iact O).909 #p # instr #iact #l1 #l2 generalize in match O in ⊢ (???%); elim instr893 ∀p.∀l : ActionLabel p. 894 is_costlabelled_act … l → ∃c.actionlabel_to_costlabel … l = [c]. 895 #p * /2/ * /2/ * 896 qed. 897 898 lemma i_act_in_map : ∀p,l_p,prog,iact,l1,l2. 899 mem ? 〈a_non_functional_label l_p iact,O〉 (labels_pc p l_p prog l1 l2 iact O). 900 #p #l_p #instr #iact #l1 #l2 generalize in match O in ⊢ (???%); elim instr 910 901 [ normalize /2/] #i #xs #IH #m whd in match (labels_pc ???); 911 902 @mem_append_l2 @IH … … 916 907 lemma static_dynamic_inv : 917 908 (* Given an assembly program *) 918 ∀p,p', prog.909 ∀p,p',l_p,prog. 919 910 920 911 (* Given an abstraction interpretation framework for the program *) 921 ∀R: asm_abstract_interpretation p p' prog.912 ∀R: asm_abstract_interpretation p p' l_p prog. 922 913 923 914 (* If the static analysis does not fail *) … … 925 916 926 917 (* For every pre_measurable, terminated trace *) 927 ∀st1,st2. ∀t: raw_trace (asm_operational_semantics … prog) … st1 st2.918 ∀st1,st2. ∀t: raw_trace l_p (asm_operational_semantics … prog) … st1 st2. 928 919 terminated_trace … t → 929 920 … … 933 924 (* Let k be the statically computed abstract action of the prefix of the trace 934 925 up to the first label *) 935 ∀k.block_cost p prog … (instr_map … R) (get_pc … st1 (endmain … prog)) (S (|(instructions … prog)|)) = return k →926 ∀k.block_cost p … prog … (instr_map … R) (get_pc … st1 (endmain … prog)) (S (|(instructions … prog)|)) = return k → 936 927 937 928 (* Let abs_actions be the list of statically computed abstract actions … … 954 945 @(proj2 … (static_analisys_ok … EQmap … Hmem)) 955 946 ] 956 #p #p' # prog #R #mT #map1 #EQmap #st1 #st2 #t * #st3 * #t1 *947 #p #p' #l_p #prog #R #mT #map1 #EQmap #st1 #st2 #t * #st3 * #t1 * 957 948 #l1 * #prf1 ** #EQ destruct #Hlabelled 958 949 >(rewrite_in_dependent_map ??? (get_costlabels_of_trace … t1)) … … 1002 993 whd in match (dependent_map ????); #costs #EQ destruct #a1 #good_st_a1 1003 994 >neutral_r >act_neutral 1004 lapply(instr_map_ok R … (refl …) good_st_a1)995 lapply(instr_map_ok … l_p … R … (refl …) good_st_a1) 1005 996 [1,7,13,19,25,31,37: whd in ⊢ (??%%); @eqb_elim normalize nodelta 1006 997 [1,3,5,7,9,11,13: #EQ cases(absurd ? EQ Hpc) ] #_ whd in match fetch_state; … … 1009 1000 | >EQfetch in ⊢ (??%?); | >EQfetch in ⊢ (??%?); | >EQfetch in ⊢ (??%?); 1010 1001 | >EQfetch in ⊢ (??%?); ] % 1011 |3,9,15,21,27,33,39: skip |*: try assumption // ]]]1002 |3,9,15,21,27,33,39: skip |*: try assumption #H @H] ]] 1012 1003 | -st1 * [3: #st1] #st3 #st4 #l [3: *] cases st3 -st3 1013 1004 [1,2,4,5: * #H1 #H2 #tl #_ #l1 #exe @⊥ lapply tl -tl lapply(refl ? (FINAL p p')) … … 1040 1031 [2: cases(static_analisys_ok … c … (pc … st3) … EQmap) // #EQ #_ <EQ whd in match (big_op ??); 1041 1032 >neutral_l assumption 1042 |3,6: [ >neutral_r] lapply(instr_map_ok R … (refl …) good_a1)1033 |3,6: [ >neutral_r] lapply(instr_map_ok … l_p … R … (refl …) good_a1) 1043 1034 [1,7: whd in ⊢ (??%?); @eqb_elim 1044 1035 [1,3: #ABS cases(absurd ? ABS Hpc) ] #_ normalize nodelta whd in match fetch_state; … … 1078 1069 [2: lapply EQpc_st3 @eqb_elim [2: #_ #EQ destruct] #EQ #_ >EQ @i_act_in_map ] 1079 1070 #EQ #_ <EQ whd in match (big_op ??); >neutral_l assumption 1080 | lapply(instr_map_ok R … (refl …) good_a1)1071 | lapply(instr_map_ok … l_p … R … (refl …) good_a1) 1081 1072 [1: % | assumption |*:] normalize in ⊢ (% → ?); #H @H 1082 1073 ] … … 1094 1085 qed. *) 1095 1086 1096 lemma execute_mem_label_pc :∀p,p', prog.∀st0,st1 :vm_state p p'.∀l1,c1.1097 actionlabel_to_costlabel l 1 = [c1] →1098 vmstep p p' prog l1 st0 st1 →1087 lemma execute_mem_label_pc :∀p,p',l_p,prog.∀st0,st1 :vm_state p p'.∀l1,c1. 1088 actionlabel_to_costlabel l_p l1 = [c1] → 1089 vmstep p p' l_p prog l1 st0 st1 → 1099 1090 mem … 〈c1,pc p p' st1〉 1100 1091 (labels_pc … (instructions … prog) (call_label_fun … prog) 1101 1092 (return_label_fun … prog) (in_act … prog) O). 1102 #p #p' # prog #st0 #st1 #l1 #c1 #EQc1 #H lapply(vm_step_to_eval … H) -H1093 #p #p' #l_p #prog #st0 #st1 #l1 #c1 #EQc1 #H lapply(vm_step_to_eval … H) -H 1103 1094 #H cases(bind_inversion ????? H); #i * #EQi #Hi cases(step_emit … EQi H) 1104 1095 [ #c2 whd in match actionlabel_to_costlabel in EQc1; normalize nodelta in EQc1; … … 1124 1115 include "Simulation.ma". 1125 1116 1126 lemma sub_trace_premeasurable_l1 : ∀p,p', prog.1127 ∀s1,s2,s3. ∀t1: raw_trace (asm_operational_semanticsp p' prog) … s1 s2.1128 ∀t2 : raw_trace (asm_operational_semanticsp p' prog) … s2 s3.1117 lemma sub_trace_premeasurable_l1 : ∀p,p',l_p,prog. 1118 ∀s1,s2,s3. ∀t1: raw_trace l_p (asm_operational_semantics p l_p p' prog) … s1 s2. 1119 ∀t2 : raw_trace l_p (asm_operational_semantics p l_p p' prog) … s2 s3. 1129 1120 pre_measurable_trace … (t1 @ t2) → 1130 1121 pre_measurable_trace … t1. 1131 #p #p' # prog #s1 #s2 #s3 #t1 lapply s3 -s3 elim t11132 [ #st #st3 #t2 whd in ⊢ (???? % → ?); #H %1 @(head_of_premeasurable_is_not_io … H) ]1122 #p #p' #l_p #prog #s1 #s2 #s3 #t1 lapply s3 -s3 elim t1 1123 [ #st #st3 #t2 whd in ⊢ (?????% → ?); #H %1 @(head_of_premeasurable_is_not_io … H) ] 1133 1124 #st1 #st2 #st3 #l #exe #tl #IH #st4 #t2 #H inversion H in ⊢ ?; 1134 1125 [ #st #class #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in EQ3 : (??%?%); destruct … … 1144 1135 qed. 1145 1136 1146 lemma sub_trace_premeasurable_l2 : ∀p,p', prog.1147 ∀s1,s2,s3. ∀t1: raw_trace (asm_operational_semanticsp p' prog) … s1 s2.1148 ∀t2 : raw_trace (asm_operational_semanticsp p' prog) … s2 s3.1137 lemma sub_trace_premeasurable_l2 : ∀p,p',l_p,prog. 1138 ∀s1,s2,s3. ∀t1: raw_trace l_p (asm_operational_semantics p l_p p' prog) … s1 s2. 1139 ∀t2 : raw_trace l_p (asm_operational_semantics p l_p p' prog) … s2 s3. 1149 1140 pre_measurable_trace … (t1 @ t2) → 1150 1141 pre_measurable_trace … t2. 1151 #p #p' # prog #s1 #s2 #s3 #t1 lapply s3 -s3 elim t11152 [ #st #st3 #t2 whd in ⊢ (???? % → ?); #H @H ]1142 #p #p' #l_p #prog #s1 #s2 #s3 #t1 lapply s3 -s3 elim t1 1143 [ #st #st3 #t2 whd in ⊢ (?????% → ?); #H @H ] 1153 1144 #st1 #st2 #st3 #l #exe #tl #IH #st4 #t2 #H inversion H in ⊢ ?; 1154 1145 [ #st #class #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in EQ3 : (??%?%); destruct … … 1168 1159 theorem static_dynamic : 1169 1160 (* Given an assembly program *) 1170 ∀p,p', prog.1161 ∀p,p',l_p,prog. 1171 1162 1172 1163 (* Given an abstraction interpretation framework for the program *) 1173 ∀R: asm_abstract_interpretation p p' prog.1164 ∀R: asm_abstract_interpretation p p' l_p prog. 1174 1165 1175 1166 (* If the static analysis does not fail *) … … 1178 1169 (* For every measurable trace whose second state is st1 or, equivalently, 1179 1170 whose first state after the initial labelled transition is st1 *) 1180 ∀si,s1,s2,sn. ∀t: raw_trace (asm_operational_semantics … prog) … si sn.1171 ∀si,s1,s2,sn. ∀t: raw_trace l_p (asm_operational_semantics … prog) … si sn. 1181 1172 measurable … s1 s2 … t → 1182 1173 … … 1203 1194 @(proj2 … (static_analisys_ok … EQmap … Hmem)) 1204 1195 ] 1205 #p #p' # prog #R #mT #map1 #EQmap #si #s1 #s2 #sn #t #measurable1196 #p #p' #l_p #prog #R #mT #map1 #EQmap #si #s1 #s2 #sn #t #measurable 1206 1197 cases measurable #s0 * #s3 * #ti0 * #t13 * #t2n* #l1 * #l2 * #prf1 * #prf2 1207 1198 ******* #pre_t13 #EQ destruct #Hl1 #Hl2 #Hcall_l2 #sil_ti0 #sil_t2n #Hcall_l1 -
LTS/costs.ma
r3535 r3549 13 13 (**************************************************************************) 14 14 15 include "basics/types.ma". 16 include "basics/deqsets.ma". 15 17 include "Traces.ma". 16 18 17 record monoid: Type[1] ≝ 18 { carrier :> DeqSet 19 ; op: carrier → carrier → carrier 20 ; e: carrier 21 ; neutral_r : ∀x. op … x e = x 22 ; neutral_l : ∀x. op … e x = x 23 ; is_associative: ∀x,y,z. op … (op … x y) z = op … x (op … y z) 24 }. 25 26 record monoid_action (I : monoid) (M : Type[0]) : Type[0] ≝ 27 { act :2> I → M → M 28 ; act_neutral : ∀x.act (e …) x = x 29 ; act_op : ∀i,j,x.act (op … i j) x = act j (act i x) 30 }. 31 32 record concrete_transition_sys : Type[2] ≝ 33 { trans_sys :> abstract_status 19 record concrete_transition_sys (p : label_params) : Type[2] ≝ 20 { trans_sys :> abstract_status p 34 21 }. 35 22 … … 43 30 interpretation "act_abs" 'sem f x = (act ?? (act_abs ?) f x). 44 31 45 record galois_rel (C : concrete_transition_sys) 32 record galois_rel (p : label_params) 33 (C : concrete_transition_sys p) 46 34 (A : abstract_transition_sys) : Type[0] ≝ 47 35 { rel_galois :2> C → A → Prop 48 36 }. 49 37 50 record galois_connection : Type[2] ≝51 { concr : concrete_transition_sys 38 record galois_connection (p : label_params) : Type[2] ≝ 39 { concr : concrete_transition_sys p 52 40 ; abs_d : abstract_transition_sys 53 ; galois_rel:> galois_rel concr abs_d41 ; galois_rel:> galois_rel p concr abs_d 54 42 }. 55 43 56 record abstract_interpretation : Type[2] ≝57 { galois_conn:> galois_connection 44 record abstract_interpretation (p : label_params) : Type[2] ≝ 45 { galois_conn:> galois_connection p 58 46 ; instr_t : Type[0] 59 ; fetch : concr galois_conn → option instr_t60 ; instr_map : instr_t → abs_instr … (abs_d galois_conn)47 ; fetch : concr … galois_conn → option instr_t 48 ; instr_map : instr_t → abs_instr … (abs_d … galois_conn) 61 49 ; instr_map_ok : 62 50 ∀s,s': concr … galois_conn. ∀a: abs_d … galois_conn.∀l,i.
Note: See TracChangeset
for help on using the changeset viewer.