include "costs.ma". include "basics/lists/list.ma". include "../src/utilities/option.ma". include "basics/jmeq.ma". lemma bind_inversion : ∀A,B : Type[0].∀m : option A. ∀f : A → option B.∀y : B. ! x ← m; f x = return y → ∃ x.(m = return x) ∧ (f x = return y). #A #B * [| #a] #f #y normalize #EQ [destruct] %{a} %{(refl …)} // qed. record assembler_params : Type[1] ≝ { seq_instr : Type[0] ; jump_condition : Type[0] ; io_instr : Type[0] }. inductive AssemblerInstr (p : assembler_params) : Type[0] ≝ | Seq : seq_instr p → option (NonFunctionalLabel) → AssemblerInstr p | Ijmp: ℕ → AssemblerInstr p | CJump : jump_condition p → ℕ → NonFunctionalLabel → NonFunctionalLabel → AssemblerInstr p | Iio : NonFunctionalLabel → io_instr p → NonFunctionalLabel → AssemblerInstr p | Icall: FunctionName → AssemblerInstr p | Iret: AssemblerInstr p. definition labels_pc_of_instr : ∀p.AssemblerInstr p → ℕ → list (CostLabel × ℕ) ≝ λp,i,program_counter. match i with [ Seq _ opt_l ⇒ match opt_l with [ Some lbl ⇒ [〈(a_non_functional_label lbl),S program_counter〉] | None ⇒ [ ] ] | Ijmp _ ⇒ [ ] | CJump _ newpc ltrue lfalse ⇒ [〈(a_non_functional_label ltrue),newpc〉; 〈(a_non_functional_label lfalse),S program_counter〉] | Iio lin _ lout ⇒ [〈(a_non_functional_label lin),program_counter〉; 〈(a_non_functional_label lout),S program_counter〉] | Icall f ⇒ [ ] | Iret ⇒ [ ] ]. let rec labels_pc (p : assembler_params) (prog : list (AssemblerInstr p)) (call_label_fun : list (ℕ × CallCostLabel)) (return_label_fun : list (ℕ × ReturnPostCostLabel)) (i_act : NonFunctionalLabel) (program_counter : ℕ) on prog : list (CostLabel × ℕ) ≝ match prog with [ nil ⇒ [〈a_non_functional_label (i_act),O〉] @ map … (λx.let〈y,z〉 ≝ x in 〈(a_call z),y〉) (call_label_fun) @ map … (λx.let〈y,z〉 ≝ x in 〈(a_return_post z),y〉) (return_label_fun) | 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) ]. include "basics/lists/listb.ma". (*doppione da mettere a posto*) let rec no_duplicates (A : DeqSet) (l : list A) on l : Prop ≝ match l with [ nil ⇒ True | cons x xs ⇒ ¬ (bool_to_Prop (x ∈ xs)) ∧ no_duplicates … xs ]. record AssemblerProgram (p : assembler_params) : Type[0] ≝ { instructions : list (AssemblerInstr p) ; endmain : ℕ ; endmain_ok : endmain < |instructions| ; entry_of_function : FunctionName → ℕ ; call_label_fun : list (ℕ × CallCostLabel) ; return_label_fun : list (ℕ × ReturnPostCostLabel) ; in_act : NonFunctionalLabel ; asm_no_duplicates : no_duplicates … (map ?? \fst … (labels_pc … instructions call_label_fun return_label_fun in_act O)) }. definition fetch: ∀p.AssemblerProgram p → ℕ → option (AssemblerInstr p) ≝ λp,l,n. nth_opt ? n (instructions … l). definition stackT: Type[0] ≝ list (nat). record sem_params (p : assembler_params) : Type[1] ≝ { m : monoid ; asm_store_type : Type[0] ; eval_asm_seq : seq_instr p → asm_store_type → option asm_store_type ; eval_asm_cond : jump_condition p → asm_store_type → option bool ; eval_asm_io : io_instr p → asm_store_type → option asm_store_type ; cost_of_io : io_instr p → asm_store_type → m ; cost_of : AssemblerInstr p → asm_store_type → m }. record vm_state (p : assembler_params) (p' : sem_params p) : Type[0] ≝ { pc : ℕ ; asm_stack : stackT ; asm_store : asm_store_type … p' ; asm_is_io : bool ; cost : m … p' }. definition label_of_pc ≝ λL.λl.λpc.find … (λp.let 〈x,y〉 ≝ p in if eqb x pc then Some L y else None ? ) l. definition option_pop ≝ λA.λl:list A. match l with [ nil ⇒ None ? | cons a tl ⇒ Some ? (〈a,tl〉) ]. inductive vmstep (p : assembler_params) (p' : sem_params p) (prog : AssemblerProgram p) : ActionLabel → relation (vm_state p p') ≝ | vm_seq : ∀st1,st2 : vm_state p p'.∀i,l. fetch … prog (pc … st1) = return (Seq p i l) → asm_is_io … st1 = false → eval_asm_seq p p' i (asm_store … st1) = return asm_store … st2 → asm_stack … st1 = asm_stack … st2 → asm_is_io … st1 = asm_is_io … st2 → S (pc … st1) = pc … st2 → op … (cost … st1) (cost_of p p' (Seq p i l) (asm_store … st1)) = cost … st2 → vmstep … (cost_act l) st1 st2 | vm_ijump : ∀st1,st2 : vm_state p p'.∀new_pc : ℕ. fetch … prog (pc … st1) = return (Ijmp p new_pc) → asm_is_io … st1 = false → asm_store … st1 = asm_store … st2 → asm_stack … st1 = asm_stack … st2 → asm_is_io … st1 = asm_is_io … st2 → new_pc = pc … st2 → op … (cost … st1) (cost_of p p' (Ijmp … new_pc) (asm_store … st1)) = cost … st2 → vmstep … (cost_act (None ?)) st1 st2 | vm_cjump_true : ∀st1,st2 : vm_state p p'.∀cond,new_pc,ltrue,lfalse. eval_asm_cond p p' cond (asm_store … st1) = return true→ fetch … prog (pc … st1) = return (CJump p cond new_pc ltrue lfalse) → asm_is_io … st1 = false → asm_store … st1 = asm_store … st2 → asm_stack … st1 = asm_stack … st2 → asm_is_io … st1 = asm_is_io … st2 → pc … st2 = new_pc → op … (cost … st1) (cost_of p p' (CJump p cond new_pc ltrue lfalse) (asm_store … st1)) = cost … st2 → vmstep … (cost_act (Some ? ltrue)) st1 st2 | vm_cjump_false : ∀st1,st2 : vm_state p p'.∀cond,new_pc,ltrue,lfalse. eval_asm_cond p p' cond (asm_store … st1) = return false→ fetch … prog (pc … st1) = return (CJump p cond new_pc ltrue lfalse) → asm_is_io … st1 = false → asm_store … st1 = asm_store … st2 → asm_stack … st1 = asm_stack … st2 → asm_is_io … st1 = asm_is_io … st2 → S (pc … st1) = pc … st2 → op … (cost … st1) (cost_of p p' (CJump … cond new_pc ltrue lfalse) (asm_store … st1)) = cost … st2 → vmstep … (cost_act (Some ? lfalse)) st1 st2 | vm_io_in : ∀st1,st2 : vm_state p p'.∀lin,io,lout. fetch … prog (pc … st1) = return (Iio p lin io lout) → asm_is_io … st1 = false → asm_store … st1 = asm_store … st2 → asm_stack … st1 = asm_stack … st2 → true = asm_is_io … st2 → pc … st1 = pc … st2 → cost … st1 = cost … st2 → vmstep … (cost_act (Some ? lin)) st1 st2 | vm_io_out : ∀st1,st2 : vm_state p p'.∀lin,io,lout. fetch … prog (pc … st1) = return (Iio p lin io lout) → asm_is_io … st1 = true → eval_asm_io … io (asm_store … st1) = return asm_store … st2 → asm_stack … st1 = asm_stack … st2 → false = asm_is_io … st2 → S (pc … st1) = pc … st2 → op … (cost … st1) (cost_of_io p p' io (asm_store … st1)) = cost … st2 → vmstep … (cost_act (Some ? lout)) st1 st2 | vm_call : ∀st1,st2 : vm_state p p'.∀f,lbl. fetch … prog (pc … st1) = return (Icall p f) → asm_is_io … st1 = false → asm_store … st1 = asm_store … st2 → S (pc … st1) :: asm_stack … st1 = asm_stack … st2 → asm_is_io … st1 = asm_is_io … st2 → entry_of_function … prog f = pc … st2 → op … (cost … st1) (cost_of p p' (Icall p f) (asm_store … st1)) = cost … st2 → label_of_pc ? (call_label_fun … prog) (entry_of_function … prog f) = return lbl → vmstep … (call_act f lbl) st1 st2 | vm_ret : ∀st1,st2 : vm_state p p'.∀newpc,lbl. fetch … prog (pc … st1) = return (Iret p) → asm_is_io … st1 = false → asm_store … st1 = asm_store … st2 → asm_stack … st1 = newpc :: asm_stack … st2 → asm_is_io … st1 = asm_is_io … st2 → newpc = pc … st2 → label_of_pc ? (return_label_fun … prog) newpc = return lbl → op … (cost … st1) (cost_of p p' (Iret p) (asm_store … st1)) = cost … st2 → vmstep … (ret_act (Some ? lbl)) st1 st2. definition eval_vmstate : ∀p : assembler_params.∀p' : sem_params p. AssemblerProgram p → vm_state p p' → option (ActionLabel × (vm_state p p')) ≝ λp,p',prog,st. ! i ← fetch … prog (pc … st); match i with [ Seq x opt_l ⇒ if asm_is_io … st then None ? else ! new_store ← eval_asm_seq p p' x (asm_store … st); return 〈cost_act opt_l, mk_vm_state ?? (S (pc … st)) (asm_stack … st) new_store false (op … (cost … st) (cost_of p p' (Seq p x opt_l) (asm_store … st)))〉 | Ijmp newpc ⇒ if asm_is_io … st then None ? else return 〈cost_act (None ?), mk_vm_state ?? newpc (asm_stack … st) (asm_store … st) false (op … (cost … st) (cost_of p p' (Ijmp … newpc) (asm_store … st)))〉 | CJump cond newpc ltrue lfalse ⇒ if asm_is_io … st then None ? else ! b ← eval_asm_cond p p' cond (asm_store … st); if b then return 〈cost_act (Some ? ltrue), mk_vm_state ?? newpc (asm_stack … st) (asm_store … st) false (op … (cost … st) (cost_of p p' (CJump … cond newpc ltrue lfalse) (asm_store … st)))〉 else return 〈cost_act (Some ? lfalse), mk_vm_state ?? (S (pc … st)) (asm_stack … st) (asm_store … st) false (op … (cost … st) (cost_of p p' (CJump … cond newpc ltrue lfalse) (asm_store … st)))〉 | Iio lin io lout ⇒ if asm_is_io … st then ! new_store ← eval_asm_io … io (asm_store … st); return 〈cost_act (Some ? lout), mk_vm_state ?? (S (pc … st)) (asm_stack … st) new_store false (op … (cost … st) (cost_of_io p p' io (asm_store … st)))〉 else return 〈cost_act (Some ? lin), mk_vm_state ?? (pc … st) (asm_stack … st) (asm_store … st) true (cost … st)〉 | Icall f ⇒ if asm_is_io … st then None ? else ! lbl ← label_of_pc ? (call_label_fun … prog) (entry_of_function … prog f); return 〈call_act f lbl, mk_vm_state ?? (entry_of_function … prog f) ((S (pc … st)) :: (asm_stack … st)) (asm_store … st) false (op … (cost … st) (cost_of p p' (Icall p f) (asm_store … st)))〉 | Iret ⇒ if asm_is_io … st then None ? else ! 〈newpc,tl〉 ← option_pop … (asm_stack … st); ! lbl ← label_of_pc ? (return_label_fun … prog) newpc; return 〈ret_act (Some ? lbl), mk_vm_state ?? newpc tl (asm_store … st) false (op … (cost … st) (cost_of p p' (Iret p) (asm_store … st)))〉 ]. lemma eval_vmstate_to_Prop : ∀p,p',prog,st1,st2,l. eval_vmstate p p' prog st1 = return 〈l,st2〉 → vmstep … prog l st1 st2. #p #p' #prog #st1 #st2 #l whd in match eval_vmstate; normalize nodelta #H cases(bind_inversion ????? H) -H * normalize nodelta [ #seq #opt_l * #EQfetch inversion(asm_is_io ???) normalize nodelta [ #_ whd in ⊢ (??%% → ?); #EQ destruct] #EQio #H cases(bind_inversion ????? H) #newstore * #EQnewstore whd in ⊢ (??%% → ?); #EQ destruct @vm_seq // | #newpc * #EQfetch inversion(asm_is_io ???) normalize nodelta #EQio [ whd in ⊢ (??%% → ?); #EQ destruct] whd in ⊢ (??%% → ?); #EQ destruct @vm_ijump // | #cond #new_pc #ltrue #lfase * #EQfetch inversion(asm_is_io ???) normalize nodelta [ #_ whd in ⊢ (??%% → ?); #EQ destruct] #EQio #H cases(bind_inversion ????? H) -H * normalize nodelta * #EQcond whd in ⊢ (??%% → ?); #EQ destruct [ @(vm_cjump_true … EQfetch) // | @(vm_cjump_false … EQfetch) //] | #lin #io #lout * #EQfetch inversion(asm_is_io ???) normalize nodelta #EQio [ #H cases(bind_inversion ????? H) -H #newstore * #EQnewstore ] whd in ⊢ (??%% → ?); #EQ destruct [ @(vm_io_out … EQfetch) // | @(vm_io_in … EQfetch) // ] | #f * #EQfetch inversion(asm_is_io ???) #EQio normalize nodelta [ whd in ⊢ (??%% → ?); #EQ destruct ] #H cases (bind_inversion ????? H) -H #lbl * #EQlb whd in ⊢ (??%% → ?); #EQ destruct @(vm_call … EQfetch) // | * #EQfetch inversion(asm_is_io ???) normalize nodelta #EQio [ whd in ⊢ (??%% → ?); #EQ destruct] #H cases(bind_inversion ????? H) -H * #newpc #tl * whd in match option_pop; normalize nodelta inversion(asm_stack ???) normalize nodelta [#_ whd in ⊢ (??%% → ?); #EQ destruct] #newpc1 #tl1 #_ #EQstack whd in ⊢ (??%% → ?); #EQ destruct #H cases(bind_inversion ????? H) #lbl * #EQlbl whd in ⊢ (??%% → ?); #EQ destruct @vm_ret // ] qed. lemma vm_step_to_eval : ∀p,p',prog,st1,st2,l.vmstep … prog l st1 st2 → eval_vmstate p p' prog st1 = return 〈l,st2〉. #p #p' #prog * #pc1 #stack1 #store1 #io1 #cost1 * #pc2 #stack2 #store2 #io2 #cost2 #l #H inversion H [ #s1 #s2 #i #opt_l #EQfetch #EQio #EQstore #EQstack #EQio1 #EQpc #EQcost #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in match eval_vmstate; normalize nodelta >EQfetch >m_return_bind normalize nodelta >EQio normalize nodelta >EQstore >m_return_bind EQio EQstack >EQcost % | #s1 #s2 #newpc #EQfetch #EQio1 #EQstore #EQstack #EQio2 #EQnewpc #EQcost #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in match eval_vmstate; normalize nodelta >EQfetch >m_return_bind normalize nodelta >EQio1 normalize nodelta EQio1 >EQstore >EQstack EQstore % |3,4: #s1 #s2 #cond #newoc #ltrue #lfalse #EQev_cond #EQfetch #EQio1 #EQstore #EQstack #EQio2 #EQnewoc #EQcost #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in match eval_vmstate; normalize nodelta >EQfetch >m_return_bind normalize nodelta >EQio1 normalize nodelta >EQev_cond >m_return_bind normalize nodelta EQio2 >EQstore >EQstack EQstore [%] >EQnewoc % |5,6: #s1 #s2 #lin #io #lout #EQfetch #EQio1 #EQstore #EQstack #EQio2 #EQpc #EQcost #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in match eval_vmstate; normalize nodelta >EQfetch >m_return_bind normalize nodelta >EQio1 normalize nodelta >EQstack EQcost [ >EQstore EQstore >m_return_bind EQfetch >m_return_bind normalize nodelta >EQio1 normalize nodelta >EQlab_pc >m_return_bind >EQentry >EQcost EQio1 EQstore % | #s1 #s2 #newpc #lbl #EQfetch #EQio1 #EQstore #EQstack #EQio2 #EQnewpc #EQlab_pc #EQcosts #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in match eval_vmstate; normalize nodelta >EQfetch >m_return_bind normalize nodelta >EQio1 normalize nodelta >EQstack whd in match option_pop; normalize nodelta >m_return_bind >EQlab_pc >m_return_bind >EQcosts >EQstore EQio1 % ] qed. coercion vm_step_to_eval. include "../src/utilities/hide.ma". discriminator option. inductive vm_ass_state (p : assembler_params) (p' : sem_params p) : Type[0] ≝ | INITIAL : vm_ass_state p p' | FINAL : vm_ass_state p p' | STATE : vm_state p p' → vm_ass_state p p'. definition ass_vmstep ≝ λp,p',prog. λl.λs1,s2 : vm_ass_state p p'. match s1 with [ STATE st1 ⇒ match s2 with [ STATE st2 ⇒ (eqb (pc ?? st1) (endmain … prog)) = false ∧ vmstep p p' prog l st1 st2 | INITIAL ⇒ False | FINAL ⇒ eqb (pc … st1) (endmain … prog) = true ∧ l = cost_act (Some … (in_act … prog)) ] | INITIAL ⇒ match s2 with [ STATE st2 ⇒ eqb (pc … st2) O = true ∧ l = cost_act (Some … (in_act … prog)) | _ ⇒ False ] | FINAL ⇒ False ]. definition asm_operational_semantics : ∀p.sem_params p → AssemblerProgram p → abstract_status ≝ λp,p',prog.let init_act ≝ cost_act (Some ? (in_act … prog)) in let end_act ≝ cost_act (Some ? (in_act … prog)) in mk_abstract_status (vm_ass_state p p') (ass_vmstep … prog) (λ_.λ_.True) (λst.match st with [ INITIAL ⇒ cl_other | FINAL ⇒ cl_other | STATE s ⇒ match fetch … prog (pc … s) with [ Some i ⇒ match i with [ Seq _ _ ⇒ cl_other | Ijmp _ ⇒ cl_other | CJump _ _ _ _ ⇒ cl_jump | Iio _ _ _ ⇒ if asm_is_io … s then cl_io else cl_other | Icall _ ⇒ cl_other | Iret ⇒ cl_other ] | None ⇒ cl_other ] ] ) (λ_.true) (λs.match s with [ INITIAL ⇒ true | _ ⇒ false]) (λs.match s with [ FINAL ⇒ true | _ ⇒ false]) ???. @hide_prf [ #s1 #s2 #l cases s1 normalize nodelta [1,2: #abs destruct] -s1 #s1 cases s2 normalize nodelta [1: #_ * |2: #_ * #_ #EQ >EQ normalize /2/ ] #s2 inversion(fetch ???) normalize nodelta [ #_ #EQ destruct] * normalize nodelta [ #seq #lbl #_ | #n #_ | #cond #newpc #ltrue #lfalse #EQfetch | #lin #io #lout #_ cases (asm_is_io ??) normalize nodelta | #f #_ | #_ ] #EQ destruct * #_ #H lapply(vm_step_to_eval … H) whd in match eval_vmstate; normalize nodelta >EQfetch >m_return_bind normalize nodelta cases(asm_is_io ??) normalize nodelta [ whd in ⊢ (??%% → ?); #EQ destruct] #H cases(bind_inversion ????? H) -H * * #_ normalize nodelta whd in ⊢ (??%% → ?); #EQ destruct % // | #s1 #s2 #l cases s1 normalize nodelta [2: #_ * |3: #s1 ] cases s2 normalize nodelta [1,2,4,5: #abs destruct ] #s2 [2: #_ * #_ /2/ ] inversion(fetch ???) normalize nodelta [ #_ #EQ destruct] * normalize nodelta [ #seq #lbl #_ | #n #_ | #cond #newpc #ltrue #lfalse #_ | #lin #io #lout #EQfetch inversion (asm_is_io ??) #EQio normalize nodelta | #f #_ | #_ ] #EQ destruct * #_ #H lapply(vm_step_to_eval … H) whd in match eval_vmstate; normalize nodelta #H cases(bind_inversion ????? H) -H * [ #seq1 #lbl1 | #n1 | #cond1 #newpc1 #ltrue1 #lfalse1 | #lin1 #io1 #lout1 | #f | ] normalize nodelta * #_ cases(asm_is_io ??) normalize nodelta [1,3,5,9,11: whd in ⊢ (??%% → ?); #EQ destruct |2,6,7,10,12: #H cases(bind_inversion ????? H) -H #x * #_ [2: cases x normalize nodelta |5: #H cases(bind_inversion ????? H) -H #y * #_ ] ] whd in ⊢ (??%% → ?); #EQ destruct destruct % // | #s1 #s2 #l cases s1 normalize nodelta [1,2: #abs destruct ] #s1 cases s2 normalize nodelta [ #_ * | #_ * /2/ ] #s2 inversion(fetch ???) normalize nodelta [ #_ #EQ destruct] * normalize nodelta [ #seq #lbl #_ | #n #_ | #cond #newpc #ltrue #lfalse #_ | #lin #io #lout #EQfetch inversion(asm_is_io ??) normalize nodelta #EQio | #f #_ | #_ ] #EQ destruct * #_ #H lapply(vm_step_to_eval … H) whd in match eval_vmstate; normalize nodelta >EQfetch >m_return_bind normalize nodelta >EQio normalize nodelta #H cases(bind_inversion ????? H) -H #x * #_ whd in ⊢ (??%% → ?); #EQ destruct % // qed. definition asm_concrete_trans_sys ≝ λp,p',prog.mk_concrete_transition_sys … (asm_operational_semantics p p' prog) (m … p') (λs.match s with [STATE st ⇒ cost … st | _ ⇒ e …] ). definition emits_labels ≝ λp.λinstr : AssemblerInstr p.match instr with [ Seq _ opt_l ⇒ match opt_l with [ None ⇒ Some ? (λpc.S pc) | Some _ ⇒ None ? ] | Ijmp newpc ⇒ Some ? (λ_.newpc) | _ ⇒ None ? ]. definition fetch_state : ∀p,p'.AssemblerProgram p → vm_state p p' → option (AssemblerInstr p) ≝ λp,p',prog,st.fetch … prog (pc … st). record asm_galois_connection (p: assembler_params) (p': sem_params p) (prog: AssemblerProgram p) : Type[2] ≝ { aabs_d : abstract_transition_sys (m … p') ; agalois_rel:> galois_rel (asm_concrete_trans_sys p p' prog) aabs_d }. definition galois_connection_of_asm_galois_connection: ∀p,p',prog. asm_galois_connection p p' prog → galois_connection ≝ λp,p',prog,agc. mk_galois_connection (asm_concrete_trans_sys p p' prog) (aabs_d … agc) (agalois_rel … agc). coercion galois_connection_of_asm_galois_connection. definition ass_fetch ≝ λp,p',prog. λs.match s with [ STATE st ⇒ if eqb (pc … st) (endmain … prog) then Some ? (None ?) else ! x ← fetch_state p p' prog st; Some ? (Some ? x) | INITIAL ⇒ Some ? (None ?) | FINAL ⇒ None ? ]. definition ass_instr_map ≝ λp,p',prog.λasm_galois_conn: asm_galois_connection p p' prog. λinstr_map: AssemblerInstr p → (*option*) (abs_instr … (abs_d asm_galois_conn)). (λi.match i with [None ⇒ (*Some …*) (e …) |Some x ⇒ instr_map … x]). record asm_abstract_interpretation (p: assembler_params) (p': sem_params p) (prog: AssemblerProgram p) : Type[2] ≝ { asm_galois_conn: asm_galois_connection p p' prog ; instr_map : AssemblerInstr p → (*option*) (abs_instr … (abs_d asm_galois_conn)) ; instr_map_ok : ∀s,s': concr … asm_galois_conn. ∀a: abs_d … asm_galois_conn.∀l,i. as_execute … l s s' → ass_fetch … prog s = Some ? i → ∀I. ass_instr_map … instr_map i = (*Some ?*) I → asm_galois_conn s a → asm_galois_conn s' (〚I〛 a) }. definition abstract_interpretation_of_asm_abstract_interpretation: ∀p,p',prog. asm_abstract_interpretation p p' prog → abstract_interpretation ≝ λp,p',prog,aai. mk_abstract_interpretation (asm_galois_conn … aai) (option (AssemblerInstr p)) (ass_fetch p p' prog) (ass_instr_map … prog … (instr_map … aai)) (instr_map_ok … aai). coercion abstract_interpretation_of_asm_abstract_interpretation. definition non_empty_list : ∀A.list A → bool ≝ λA,l.match l with [ nil ⇒ false | _ ⇒ true ]. let rec block_cost (p : assembler_params) (prog: AssemblerProgram p) (abs_t : monoid) (instr_m : AssemblerInstr p → abs_t) (prog_c: option ℕ) (program_size: ℕ) on program_size: option abs_t ≝ match prog_c with [ None ⇒ return e … abs_t | Some program_counter ⇒ match program_size with [ O ⇒ None ? | S program_size' ⇒ if eqb program_counter (endmain … prog) then return e … abs_t else ! instr ← fetch … prog program_counter; ! n ← (match emits_labels … instr with [ Some f ⇒ block_cost … prog abs_t instr_m (Some ? (f program_counter)) program_size' | None ⇒ return e … ]); return (op … abs_t (instr_m … instr) n) ] ]. record cost_map_T (dom : DeqSet) (codom : Type[0]) : Type[1] ≝ { map_type :> Type[0] ; empty_map : map_type ; get_map : map_type → dom → option codom ; set_map : map_type → dom → codom → map_type ; get_set_hit : ∀k,v,m.get_map (set_map m k v) k = return v ; get_set_miss : ∀k1,k2,v,m.(k1 == k2) = false → get_map (set_map m k1 v) k2 = get_map m k2 }. lemma labels_pc_ok : ∀p,prog,l1,l2,i_act,i,lbl,pc,m. nth_opt ? pc prog = return i → mem ? lbl (labels_pc_of_instr … i (m+pc)) → mem ? lbl (labels_pc p prog l1 l2 i_act m). #p #instrs #l1 #l2 #iact #i #lbl #pc whd in match fetch; normalize nodelta lapply pc -pc elim instrs [ #pc #m whd in ⊢ (??%% → ?); #EQ destruct] #x #xs #IH * [|#pc'] #m whd in ⊢ (??%% → ?); [ #EQ destruct #lbl_addr whd in match (labels_pc ???); /2 by mem_append_l1/ | #EQ #H2 whd in match (labels_pc ???); @mem_append_l2 @(IH … EQ) // ] qed. lemma labels_pf_in_act: ∀p,prog,pc. mem CostLabel (in_act p prog) (map (CostLabel×ℕ) CostLabel \fst (labels_pc p (instructions p prog) (call_label_fun p prog) (return_label_fun p prog) (in_act p prog) pc)). #p #prog elim (instructions p prog) normalize /2/ qed. lemma labels_pc_return: ∀p,prog,l1,l2,iact,x1,x2. label_of_pc ReturnPostCostLabel l2 x1=return x2 → ∀m. mem … 〈(a_return_post x2),x1〉 (labels_pc p prog l1 l2 iact m). #p #l #l1 #l2 #iact whd in match (labels_pc ???); #x1 #x2 #H elim l [ #m @mem_append_l2 @mem_append_l2 whd in H:(??%?); elim l2 in H; [ whd in ⊢ (??%% → ?); #EQ destruct] * #x #y #tl #IH whd in ⊢ (??%? → %); normalize nodelta @eqb_elim normalize nodelta [ #EQ whd in ⊢ (??%% → ?); #EQ2 destruct /2/ | #NEQ #H2 %2 @IH // ] | #hd #tl #IH #m @mem_append_l2 @IH ] qed. lemma labels_pc_call: ∀p,prog,l1,l2,iact,x1,x2. label_of_pc CallCostLabel l1 x1=return x2 → ∀m. mem … 〈(a_call x2),x1〉 (labels_pc p prog l1 l2 iact m). #p #l #l1 #l2 #iact whd in match (labels_pc ???); #x1 #x2 #H elim l [ #m @mem_append_l2 @mem_append_l1 whd in H:(??%?); elim l1 in H; [ whd in ⊢ (??%% → ?); #EQ destruct] * #x #y #tl #IH whd in ⊢ (??%? → %); normalize nodelta @eqb_elim normalize nodelta [ #EQ whd in ⊢ (??%% → ?); #EQ2 destruct /2/ | #NEQ #H2 %2 @IH // ] | #hd #tl #IH #m @mem_append_l2 @IH ] qed. (* lemma labels_pc_bounded : ∀p.∀prog : AssemblerProgram p.∀lbl,pc.∀m. mem ? 〈lbl,pc〉 (labels_pc p (instructions … prog) m) → (m + pc) < (|(instructions … prog)|). #p * #instr #endmain #_ #H1 #H2 elim instr [ #H3 @⊥ /2/ ] #x #xs #IH #_ #lbl #pc #m whd in match (labels_pc ???); #H cases(mem_append ???? H) -H [ whd in match labels_pc_of_instr; normalize nodelta cases x normalize nodelta [ #seq * [|#lab] | #newpc | #cond #newpc #ltrue #lfalse | #lin #io #lout | #f | ] normalize [1,3,6,7: *] * [2,4,6: * [2,4:*] ] #EQ destruct *) let rec m_foldr (M : Monad) (X,Y : Type[0]) (f : X→Y→M Y) (l : list X) (y : Y) on l : M Y ≝ match l with [ nil ⇒ return y | cons x xs ⇒ ! z ← m_foldr M X Y f xs y; f x z ]. definition static_analisys : ∀p : assembler_params.∀abs_t : monoid.(AssemblerInstr p → abs_t) → ∀mT : cost_map_T DEQCostLabel abs_t.AssemblerProgram p → option mT ≝ λp,abs_t,instr_m,mT,prog. let prog_size ≝ S (|instructions … prog|) in m_foldr Option ?? (λx,m.let 〈z,w〉≝ x in ! k ← block_cost p prog abs_t instr_m (Some ? w) prog_size; return set_map … m z k) (labels_pc … (instructions … prog) (call_label_fun … prog) (return_label_fun … prog) (in_act … prog) O) (empty_map ?? mT). (*falso: necessita di no_duplicates*) definition eq_deq_prod : ∀D1,D2 : DeqSet.D1 × D2 → D1 × D2 → bool ≝ λD1,D2,x,y.\fst x == \fst y ∧ \snd x == \snd y. definition DeqProd ≝ λD1 : DeqSet.λD2 : DeqSet. mk_DeqSet (D1 × D2) (eq_deq_prod D1 D2) ?. @hide_prf * #x1 #x2 * #y1 #y2 whd in match eq_deq_prod; normalize nodelta % [2: #EQ destruct @andb_Prop >(\b (refl …)) %] inversion (? ==?) #EQ1 whd in match (andb ??); #EQ2 destruct >(\P EQ1) >(\P EQ2) % qed. (* unification hint 0 ≔ D1,D2 ; X ≟ DeqProd D1 D2 (* ---------------------------------------- *) ⊢ D1 × D2 ≡ carr X. unification hint 0 ≔ D1,D2,p1,p2; X ≟ DeqProd D1 D2 (* ---------------------------------------- *) ⊢ eq_deq_prod D1 D2 p1 p2 ≡ eqb X p1 p2. definition deq_prod_to_prod : ∀D1,D2 : DeqSet.DeqProd D1 D2 → D1 × D2 ≝ λD1,D2,x.x. coercion deq_prod_to_prod. *) lemma map_mem : ∀A,B,f,l,a.mem A a l → ∃b : B.mem B b (map A B f l) ∧ b = f a. #A #B #f #l elim l [ #a *] #x #xs #IH #a * [ #EQ destruct %{(f x)} % // % // | #H cases(IH … H) #b * #H1 #EQ destruct %{(f a)} % // %2 // ] qed. lemma static_analisys_ok : ∀p,abs_t,instr_m,mT,prog,lbl,pc,map. static_analisys p abs_t instr_m mT prog = return map → mem … 〈lbl,pc〉 (labels_pc … (instructions … prog) (call_label_fun … prog) (return_label_fun … prog) (in_act … prog) O) → get_map … map lbl = block_cost … prog abs_t instr_m (Some ? pc) (S (|(instructions … prog)|)) ∧ block_cost … prog abs_t instr_m (Some ? pc) (S (|(instructions … prog)|)) ≠ None ?. #p #abs_t #instr_m #mT * #prog whd in match static_analisys; normalize nodelta #endmain #Hendmain #entry_fun #call_label_fun #return_label_fun #inact #nodup generalize in match nodup in ⊢ (∀_.∀_.∀_. (??(????%??)?) → %); #Hnodup lapply nodup -nodup lapply (labels_pc ??????) #l elim l [ #x #y #z #w #h * ] * #hd1 #hd2 #tl #IH * #H1 #H2 #lbl #pc #map #H cases(bind_inversion ????? H) -H #map1 * #EQmap1 normalize nodelta #H cases(bind_inversion ????? H) -H #elem * #EQelem whd in ⊢ (??%% → ?); #EQ destruct * [ #EQ destruct % [ >get_set_hit >EQelem % | >EQelem % whd in ⊢ (??%% → ?); #EQ destruct] | #H % [ >get_set_miss [ @(proj1 … (IH …)) //] inversion (? == ?) [2: #_ %] #ABS cases H1 -H1 #H1 @⊥ @H1 >(\P ABS) >mem_to_memb // cases(map_mem … \fst … H) #z1 * #Hz1 #EQ destruct @Hz1 | @(proj2 … (IH …)) // ] ] qed. include "Simulation.ma". definition terminated_trace : ∀S : abstract_status.∀s1,s3 : S.raw_trace … s1 s3 → Prop ≝ λ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)) ∧ is_costlabelled_act l. definition big_op: ∀M: monoid. list M → M ≝ λM. foldl … (op … M) (e … M). lemma big_op_associative: ∀M:monoid. ∀l1,l2. big_op M (l1@l2) = op M (big_op … l1) (big_op … l2). #M #l1 whd in match big_op; normalize nodelta generalize in match (e M) in ⊢ (? → (??%(??%?))); elim l1 [ #c #l2 whd in match (append ???); normalize lapply(neutral_r … c) generalize in match c in ⊢ (??%? → ???%); lapply c -c lapply(e M) elim l2 normalize [ #c #c1 #c2 #EQ @sym_eq // | #x #xs #IH #c1 #c2 #c3 #EQ act_neutral // | #hd #tl #IH #l2 #x change with ([?]@(tl@l2)) in match ([?]@(tl@l2)); >big_op_associative >act_op >IH change with ([hd]@tl) in match ([hd]@tl); >big_op_associative >act_op in ⊢ (???%); % ] qed. lemma monotonicity_of_block_cost : ∀p,prog,abs_t,instr_m,pc,size,k. block_cost p prog abs_t instr_m (Some ? pc) size = return k → ∀size'.size ≤ size' → block_cost p prog abs_t instr_m (Some ? pc) size' = return k. #p #prog #abs_t #instr_m #pc #size lapply pc elim size [ #pc #k whd in ⊢ (??%% → ?); #EQ destruct] #n #IH #pc #k whd in ⊢ (??%% → ?); @eqb_elim [ #EQ destruct normalize nodelta whd in ⊢ (??%% → ?); #EQ destruct #size' * [2: #m #_] whd in ⊢ (??%%); @eqb_elim try( #_ %) * #H @⊥ @H % | #Hpc normalize nodelta #H cases(bind_inversion ????? H) -H #i * #EQi #H cases(bind_inversion ????? H) -H #elem * #EQelem whd in ⊢ (??%% → ?); #EQ destruct #size' * [ whd in ⊢ (??%?); @eqb_elim [ #EQ @⊥ @(absurd ?? Hpc) assumption ] #_ normalize nodelta >EQi >m_return_bind >EQelem % | #m #Hm whd in ⊢ (??%?); @eqb_elim [ #EQ @⊥ @(absurd ?? Hpc) assumption ] #_ normalize nodelta >EQi >m_return_bind cases (emits_labels ??) in EQelem; normalize nodelta [ whd in ⊢ (??%%→ ??%%); #EQ destruct %] #f #EQelem >(IH … EQelem) [2: /2/ ] % ] ] qed. lemma step_emit : ∀p,p',prog,st1,st2,l,i. fetch p prog (pc … st1) = return i → eval_vmstate p p' … prog st1 = return 〈l,st2〉 → emits_labels … i = None ? → ∃x. match l in ActionLabel return λ_:ActionLabel.(list CostLabel) with  [call_act f c ⇒ [a_call c] |ret_act x ⇒ match x with [None⇒[]|Some c⇒[a_return_post c]] |cost_act x ⇒ match x with [None⇒[]|Some c⇒[a_non_functional_label c]] ] = [x] ∧ (mem … 〈x,pc … st2〉 (labels_pc p (instructions … prog) (call_label_fun … prog) (return_label_fun … prog) (in_act … prog) O)). #p #p' #prog #st1 #st2 #l #i #EQi whd in match eval_vmstate; normalize nodelta >EQi >m_return_bind normalize nodelta cases i in EQi; -i normalize nodelta [ #seq * [|#lab] | #newpc | #cond #newpc #ltrue #lfalse | #lin #io #lout | #f | ] #EQi cases(asm_is_io ???) normalize nodelta [1,3,5,7,11,13: whd in ⊢ (??%% → ?); #EQ destruct |2,4,8,9,12,14: #H cases(bind_inversion ????? H) -H #x1 * #EQx1 [3: cases x1 in EQx1; -x1 #EQx1 normalize nodelta |6: #H cases(bind_inversion ????? H) -H #x2 * #EQx2 ] ] whd in ⊢ (??%% → ?); #EQ destruct whd in match emits_labels; normalize nodelta #EQ destruct % [2,4,6,8,10,12,14: % try % |*:] [1,2,4,5,7: @(labels_pc_ok … EQi) normalize /3 by or_introl,or_intror/ ] /2 by labels_pc_return, labels_pc_call/ qed. lemma step_non_emit : ∀p,p',prog,st1,st2,l,i,f. fetch p prog (pc … st1) = return i → eval_vmstate p p' … prog st1 = return 〈l,st2〉 → emits_labels … i = Some ? f → match l in ActionLabel return λ_:ActionLabel.(list CostLabel) with  [call_act f c ⇒ [a_call c] |ret_act x ⇒ match x with [None⇒[]|Some c⇒[a_return_post c]] |cost_act x ⇒ match x with [None⇒[]|Some c⇒[a_non_functional_label c]] ] = [ ] ∧ pc … st2 = f (pc … st1). #p #p' #prog #st1 #st2 #l #i #f #EQi whd in match eval_vmstate; normalize nodelta >EQi >m_return_bind normalize nodelta cases i in EQi; -i normalize nodelta [ #seq * [|#lab] | #newpc | #cond #newpc #ltrue #lfalse | #lin #io #lout | #f | ] #EQi cases(asm_is_io ???) normalize nodelta [1,3,5,7,11,13: whd in ⊢ (??%% → ?); #EQ destruct |2,4,8,9,12,14: #H cases(bind_inversion ????? H) -H #x1 * #EQx1 [3: cases x1 in EQx1; -x1 #EQx1 normalize nodelta |6: #H cases(bind_inversion ????? H) -H #x2 * #EQx2 ] ] whd in ⊢ (??%% → ?); #EQ destruct whd in match emits_labels; normalize nodelta #EQ destruct /2 by refl, conj/ qed. lemma labels_of_trace_are_in_code : ∀p,p',prog.∀st1,st2 : vm_ass_state p p'.∀t : raw_trace (asm_operational_semantics p p' prog) … st1 st2. ∀lbl. mem … lbl (get_costlabels_of_trace … t) → mem … lbl (map … \fst … (labels_pc … (instructions p prog) (call_label_fun … prog) (return_label_fun … prog) (in_act … prog) O)). #p #p' #prog #st1 #st2 #t elim t [ #st #lbl * | #st1 #st2 #st3 #l whd in ⊢ (% → ?); cases st1 -st1 normalize nodelta [2: * |3: #st1] cases st2 -st2 normalize nodelta [1,4,5: * |2: * #HN1 #HN2 >HN2 -HN2 |6: #st2 * #HN1 #HN2 >HN2 -HN2 |3: #st2 ] [3: * #H1 #H2 #tl #IH #lbl whd in match (get_costlabels_of_trace ????); #H cases(mem_append ???? H) -H [2: #H @IH //] lapply(vm_step_to_eval … H2) whd in match eval_vmstate; normalize nodelta #H cases(bind_inversion ????? H) -H #i * #EQi #_ inversion(emits_labels … i) [ #EQemit cases(step_emit … (vm_step_to_eval … H2)) // #x * #EQ1 #EQ2 >EQ1 * [2: *] #EQ destruct cases(map_mem … \fst … EQ2) #y * #H3 #EQ destruct // | #f #EQemit >(proj1 … (step_non_emit … EQi (vm_step_to_eval … H2) … EQemit)) * ] |*: #tl #IH #lbl whd in match (get_costlabels_of_trace ????); * // @IH ] qed. let rec dependent_map (A,B : Type[0]) (l : list A) (f : ∀a : A.mem … a l → B) on l : list B ≝ (match l return λx.l=x → ? with [ nil ⇒ λ_.nil ? | cons x xs ⇒ λprf.(f x ?) :: dependent_map A B xs (λx,prf1.f x ?) ])(refl …). [ >prf %% | >prf %2 assumption] qed. lemma dependent_map_append : ∀A,B,l1,l2,f. dependent_map A B (l1 @ l2) (λa,prf.f a prf) = (dependent_map A B l1 (λa,prf.f a ?)) @ (dependent_map A B l2 (λa,prf.f a ?)). [2: @hide_prf /2/ | 3: @hide_prf /2/] #A #B #l1 elim l1 normalize /2/ qed. lemma rewrite_in_dependent_map : ∀A,B,l1,l2,f. ∀EQ:l1 = l2. dependent_map A B l1 (λa,prf.f a prf) = dependent_map A B l2 (λa,prf.f a ?). [2: >EQ // | #A #B #l1 #l2 #f #EQ >EQ in f; #f % ] qed. definition get_pc : ∀p,p'.vm_ass_state p p' → ℕ → option ℕ ≝ λp,p',st,endmain.match st with [ STATE s ⇒ Some ? (pc … s) | INITIAL ⇒ None ? | FINAL ⇒ Some ? endmain ]. lemma tbase_tind_append : ∀S : abstract_status.∀st1,st2,st3.∀t : raw_trace … st1 st2. ∀l,prf.∀t'. t_base … st1 = t @ t_ind S st2 st3 st1 l prf t' → False. #S #st1 #st2 #st3 * [ #st #l #prf #t' normalize #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct ] #s1 #s2 #s3 #l1 #prf1 #t2 #l2 #prf2 #t3 normalize #EQ lapply(eq_to_jmeq ??? EQ) -EQ #EQ destruct qed. let rec chop (A : Type[0]) (l : list A) on l : list A ≝ match l with [ nil ⇒ nil ? | cons x xs ⇒ match xs with [ nil ⇒ nil ? | cons _ _ ⇒ x :: chop … xs] ]. lemma chop_append_singleton : ∀A : Type[0].∀x : A.∀l : list A.chop ? (l@ [x]) = l. #A #x #l elim l normalize // #y * normalize // qed. lemma chop_mem : ∀A : Type[0].∀x : A.∀l : list A. mem … x (chop ? l) → mem … x l. #A #x #l elim l [*] #y * [ normalize /2/] #z #zs #IH * [/2/] /3/ qed. definition actionlabel_to_costlabel : ActionLabel → list CostLabel ≝ λa.match a with [ call_act f l ⇒ [a_call l] | ret_act opt_l ⇒ match opt_l with [None ⇒ [ ] | Some l ⇒ [a_return_post l]] | cost_act opt_l ⇒ match opt_l with [None ⇒ [ ] | Some l ⇒ [a_non_functional_label l]] ]. lemma get_cost_label_of_trace_tind : ∀S : abstract_status. ∀st1,st2,st3 : S.∀l,prf,t. get_costlabels_of_trace … (t_ind … st1 st2 st3 l prf t) = actionlabel_to_costlabel l @ get_costlabels_of_trace … t. #S #st1 #st2 #st3 * // qed. lemma actionlabel_ok : ∀l : ActionLabel. is_costlabelled_act l → ∃c.actionlabel_to_costlabel l = [c]. * /2/ * /2/ * qed. lemma i_act_in_map : ∀p,prog,iact,l1,l2. mem ? 〈a_non_functional_label iact,O〉 (labels_pc p prog l1 l2 iact O). #p #instr #iact #l1 #l2 generalize in match O in ⊢ (???%); elim instr [ normalize /2/] #i #xs #IH #m whd in match (labels_pc ???); @mem_append_l2 @IH qed. coercion big_op : ∀M:monoid. ∀l: list M. M ≝ big_op on _l: list ? to ?. lemma static_dynamic_inv : (* Given an assembly program *) ∀p,p',prog. (* Given an abstraction interpretation framework for the program *) ∀R: asm_abstract_interpretation p p' prog. (* If the static analysis does not fail *) ∀mT,map1. ∀EQmap : static_analisys … (instr_map … R) mT prog = return map1. (* For every pre_measurable, terminated trace *) ∀st1,st2. ∀t: raw_trace (asm_operational_semantics … prog) … st1 st2. terminated_trace … t → pre_measurable_trace … t → (* Let labels be the costlabels observed in the trace (last one excluded) *) let labels ≝ chop … (get_costlabels_of_trace … t) in (* Let k be the statically computed abstract action of the prefix of the trace up to the first label *) ∀k.block_cost p prog … (instr_map … R) (get_pc … st1 (endmain … prog)) (S (|(instructions … prog)|)) = return k → (* Let abs_actions be the list of statically computed abstract actions associated to each label in labels. *) ∀abs_actions. abs_actions = dependent_map … labels (λlbl,prf.(opt_safe … (get_map … map1 lbl) …)) → (* Given an abstract state in relation with the first state of the trace *) ∀a1.R st1 a1 → (* The final state of the trace is in relation with the one obtained by first applying k to a1, and then applying every semantics in abs_trace. *) R st2 (〚abs_actions〛 (〚k〛 a1)). [2: @hide_prf cases(mem_map ????? (labels_of_trace_are_in_code … (chop_mem … prf))) * #lbl' #pc * #Hmem #EQ destruct >(proj1 … (static_analisys_ok … EQmap … Hmem)) @(proj2 … (static_analisys_ok … EQmap … Hmem)) ] #p #p' #prog #R #mT #map1 #EQmap #st1 #st2 #t * #st3 * #t1 * #l1 * #prf1 * #EQ destruct #Hlabelled >(rewrite_in_dependent_map ??? (get_costlabels_of_trace … t1)) [2: >get_cost_label_append >get_cost_label_of_trace_tind >append_nil cases(actionlabel_ok … Hlabelled) #c #EQc >EQc // ] lapply Hlabelled lapply prf1 -prf1 lapply l1 -l1 elim t1 -st3 [ * [3: #st] #l #prf #H1 #_ #k whd in ⊢ (??%? → ?); [3: cases prf |2: whd in ⊢ (??%% → ?); #EQ destruct #labels whd in ⊢ (??%% → ?); #EQlabels #a1 #rel_fin lapply(instr_map_ok … R … prf … (refl …) rel_fin) [ %|] cases st2 in prf; -st2 [3: #st2] * #EQpc #EQ destruct #H >act_neutral >act_neutral normalize in H; <(act_neutral … (act_abs …) a1) @H | @eqb_elim normalize nodelta [ #EQpc whd in ⊢(??%% → ?); #EQ destruct #labels whd in ⊢ (??%% → ?); #EQ destruct #a1 #good_st_a1>act_neutral >act_neutral whd in prf; cases st2 in prf; -st2 [3: #st2] normalize nodelta * >EQpc @eqb_elim [2,4: * #ABS @⊥ @ABS %] #_ #EQ destruct #EQ destruct whd in EQc : (??%%); destruct lapply(instr_map_ok … R … (refl …) good_st_a1) [5: @(FINAL …) |2: whd % [2: % | // ] | whd whd in ⊢ (??%%); @eqb_elim [2: * #ABS @⊥ @ABS assumption | #_ % ] |3,4: skip] whd in ⊢ (% → ?); >act_neutral #H @H | #Hpc lapply prf whd in ⊢ (% → ?); cases st2 in prf; -st2 [3: #st2] #prf normalize nodelta [2:* |3: * #ABS @⊥ lapply ABS -ABS @eqb_elim [#EQ #_ @(absurd ? EQ Hpc) | #_ #EQ destruct ] ] * #INUTILE #H4 #H cases(bind_inversion ????? H) -H * [ #seq * [|#lbl1] | #newpc | #cond #newpc #ltrue #lfalse | #lin #io #lout | #f | ] * #EQfetch lapply(vm_step_to_eval … H4) whd in match eval_vmstate in ⊢ (% → ?); normalize nodelta >EQfetch >m_return_bind normalize nodelta cases(asm_is_io ??) normalize nodelta [1,3,5,7,11,13: whd in ⊢ (??%% → ?); #EQ destruct |2,4,8,9,12,14: #H cases(bind_inversion ????? H) -H #x * #_ [3: cases x normalize nodelta |6: #H cases(bind_inversion ????? H) -H #y * #_ ] ] whd in ⊢ (??%% → ?); #EQ destruct [4,8: cases H1 ] >m_return_bind whd in ⊢ (??%% → ?); #EQ destruct whd in match (dependent_map ????); #costs #EQ destruct #a1 #good_st_a1 >neutral_r >act_neutral lapply(instr_map_ok R … (refl …) good_st_a1) [1,7,13,19,25,31,37: whd in ⊢ (??%%); @eqb_elim normalize nodelta [1,3,5,7,9,11,13: #EQ cases(absurd ? EQ Hpc) ] #_ whd in match fetch_state; normalize nodelta [ >EQfetch in ⊢ (??%?); | >EQfetch in ⊢ (??%?); | >EQfetch in ⊢ (??%?); | >EQfetch in ⊢ (??%?); | >EQfetch in ⊢ (??%?); | >EQfetch in ⊢ (??%?); | >EQfetch in ⊢ (??%?); ] % |3,9,15,21,27,33,39: skip |*: try assumption // ]]] | -st1 * [3: #st1] #st3 #st4 #l [3: *] cases st3 -st3 [1,2,4,5: * #H1 #H2 #tl #_ #l1 #exe @⊥ lapply tl -tl lapply(refl ? (FINAL p p')) generalize in match (FINAL ??) in ⊢ (??%? → %); #st5 #EQst5 #tl lapply EQst5 lapply exe lapply st2 -st2 -EQst5 elim tl [ #st #st5 #ABS #EQ destruct cases ABS | #s1 #s2 #s3 #l2 #H3 #tl1 #IH #s4 #_ #EQ destruct cases H3 ] ] #st3 #exe_st1_st3 #tl #IH #l1 #exe_st4_st2 #l1_lab #pre_meas #k whd in ⊢ (??%? → ?); >rewrite_in_dependent_map [2,5: @get_cost_label_of_trace_tind |3,6: ] >dependent_map_append [ @eqb_elim [ #ABS @⊥ cases exe_st1_st3 >ABS @eqb_elim [ #_ #EQ destruct | * #ABS1 @⊥ @ABS1 %] ] #Hpc normalize nodelta #H cases(bind_inversion ????? H) -H #i * #EQi inversion(emits_labels ??) [ #EQemits whd in ⊢ (??%% → ?); #EQ destruct #labels cases(step_emit … EQi … EQemits) [4: cases exe_st1_st3 #EQ #H @(vm_step_to_eval … H) |2,3:] #c * #EQc #Hc whd in match actionlabel_to_costlabel; normalize nodelta >rewrite_in_dependent_map [2: @EQc |3:] whd in match (dependent_map ????); @opt_safe_elim #k_c #EQk_c whd in match (dependent_map ????); letin ih_labels ≝ (dependent_map ????) #EQ destruct #a1 #good_a1 >big_op_associative >act_op @IH | #f #EQemits normalize nodelta #H cases(bind_inversion ????? H) -H #k' * #EQk' whd in ⊢ (??%% → ?); #EQ destruct(EQ) #labels cases(step_non_emit … EQi… EQemits) [4: cases exe_st1_st3 #EQ #H @(vm_step_to_eval … H) |2,3:] #EQl #EQpc >(rewrite_in_dependent_map ??? []) [2: assumption] whd in match (dependent_map ????); #EQlabels #a1 #good_a1 >act_op @IH ] try // [2: cases(static_analisys_ok … c … (pc … st3) … EQmap) // #EQ #_ neutral_l assumption |3,6: [ >neutral_r] lapply(instr_map_ok R … (refl …) good_a1) [1,7: whd in ⊢ (??%?); @eqb_elim [1,3: #ABS cases(absurd ? ABS Hpc) ] #_ normalize nodelta whd in match fetch_state; normalize nodelta [ >EQi in ⊢ (??%?); | >EQi in ⊢ (??%?); ] % |2,8: assumption |*: ] normalize in ⊢ (% → ?); #H @H |5: whd in match get_pc; normalize nodelta >EQpc >(monotonicity_of_block_cost … EQk') // |*: inversion pre_meas in ⊢ ?; [1,6: #st #c #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in EQ3 : (??%??); destruct |2,7: #s1 #s2 #s3 #lbl #exe #tl1 #s1_noio * #opt_l #EQ destruct #pre_tl1 #_ #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in EQ3 : (??%??); destruct |3,8: #s1 #s2 #s3 #lbl #s1_noio #exe #tl1 * #lbl1 #EQ destruct #pre_tl1 #_ #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in EQ3 : (??%??); destruct |4,9: #s1 #s2 #s3 #lbl #exe #tl1 #s1_noio * #f * #lbl1 #EQ destruct * #pre_tl1 #_ #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in EQ3 : (??%??); destruct |5,10: #s1 #s2 #s3 #s4 #s5 #l1 #l2 #exe1 #t1 #t2 #exe2 #noio1 #noio2 #H * ] // ] | whd in ⊢ (??%% → ?); #EQ destruct cases exe_st1_st3 #EQpc_st3 #EQ destruct #labels whd in match actionlabel_to_costlabel; normalize nodelta whd in match (dependent_map ????); @opt_safe_elim #k_c #EQk_c letin ih_labels ≝ (dependent_map ????) change with ([?]@?) in match ([?]@?); #EQ #a1 #good_a1 destruct >big_op_associative >act_op @IH try // [ inversion pre_meas in ⊢ ?; [1,6: #st #c #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in EQ3 : (??%??); destruct |2,7: #s1 #s2 #s3 #lbl #exe #tl1 #s1_noio * #opt_l #EQ destruct #pre_tl1 #_ #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in EQ3 : (??%??); destruct |3,8: #s1 #s2 #s3 #lbl #s1_noio #exe #tl1 * #lbl1 #EQ destruct #pre_tl1 #_ #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in EQ3 : (??%??); destruct |4,9: #s1 #s2 #s3 #lbl #exe #tl1 #s1_noio * #f * #lbl1 #EQ destruct * #pre_tl1 #_ #EQ1 #EQ2 #EQ3 #EQ4 destruct whd in EQ3 : (??%??); destruct |5,10: #s1 #s2 #s3 #s4 #s5 #l1 #l2 #exe1 #t1 #t2 #exe2 #noio1 #noio2 #H * ] // | cases(static_analisys_ok … (in_act … prog) … (pc … st3) … EQmap) [2: lapply EQpc_st3 @eqb_elim [2: #_ #EQ destruct] #EQ #_ >EQ @i_act_in_map ] #EQ #_ neutral_l assumption | lapply(instr_map_ok R … (refl …) good_a1) [1: % | assumption |*:] normalize in ⊢ (% → ?); #H @H ] ] qed. definition measurable : ∀S: abstract_status. ∀si,s1,s3,sn : S. raw_trace … si sn → Prop ≝ λS,si,s1,s3,sn,t. pre_measurable_trace … t ∧ ∃s0,s2:S. ∃ti0 : raw_trace … si s0.∃t12 : raw_trace … s1 s2.∃t3n : raw_trace … s3 sn. ∃l1,l2,prf1,prf2. t = ti0 @ t_ind ? s0 s1 sn l1 prf1 … (t12 @ (t_ind ? s2 s3 sn l2 prf2 … t3n)) ∧ is_costlabelled_act l1 ∧ is_costlabelled_act l2 ∧ silent_trace … ti0 ∧ silent_trace … t3n. lemma chop_cons : ∀A : Type[0].∀x : A.∀xs: list A. xs ≠ [ ] → chop … (x :: xs) = x :: (chop … xs). #A #x * [ #H cases(absurd ?? H) % ] // qed. (* lemma measurable_terminated: ∀S,s1,s2,s4,t. measurable S s1 s2 s4 t → terminated_trace … t. #S #s1 #s2 #s4 #t * #_ * #s3 * #t' * #l1 * #l2 * #prf1 * #prf2 ** #EQ destruct /6 width=6 by ex_intro, conj/ qed. *) lemma execute_mem_label_pc :∀p,p',prog.∀st0,st1 :vm_state p p'.∀l1,c1. actionlabel_to_costlabel l1 = [c1] → vmstep p p' prog l1 st0 st1 → mem … 〈c1,pc p p' st1〉 (labels_pc … (instructions … prog) (call_label_fun … prog) (return_label_fun … prog) (in_act … prog) O). #p #p' #prog #st0 #st1 #l1 #c1 #EQc1 #H lapply(vm_step_to_eval … H) -H #H cases(bind_inversion ????? H); #i * #EQi #Hi cases(step_emit … EQi H) [ #c2 whd in match actionlabel_to_costlabel in EQc1; normalize nodelta in EQc1; >EQc1 * #EQ destruct // ] cases i in Hi; [ #seq * [|#lbl1] | #newpc | #cond #newpc #ltrue #lfalse | #lin #io #lout | #f | ] normalize nodelta cases(asm_is_io ??) normalize nodelta [1,3,5,7,11,13: whd in ⊢ (??%% → ?); #EQ destruct |2,4,8,9,12,14: #H cases(bind_inversion ????? H) -H #x * #_ [3: cases x normalize nodelta |6: #H cases(bind_inversion ????? H) -H #y * #_ ] ] whd in ⊢ (??%% → ?); #EQ destruct try % normalize in EQc1; destruct qed. theorem static_dynamic : (* Given an assembly program *) ∀p,p',prog. (* Given an abstraction interpretation framework for the program *) ∀R: asm_abstract_interpretation p p' prog. (* If the static analysis does not fail *) ∀mT,map1. ∀EQmap : static_analisys … (instr_map … R) mT prog = return map1. (* For every measurable trace whose second state is st1 or, equivalently, whose first state after the initial labelled transition is st1 *) ∀si,s1,s2,sn. ∀t: raw_trace (asm_operational_semantics … prog) … si sn. measurable … s1 s2 … t → (* Let labels be the costlabels observed in the trace (last one excluded) *) let labels ≝ chop … (get_costlabels_of_trace … t) in (* Let abs_actions be the list of statically computed abstract actions associated to each label in labels. *) ∀abs_actions. abs_actions = dependent_map … labels (λlbl,prf.(opt_safe … (get_map … map1 lbl) …)) → (* Given an abstract state in relation with the second state of the trace *) ∀a1.R s1 a1 → (* The final state of the trace is in relation with the one obtained by applying every semantics in abs_trace. *) R s2 (〚abs_actions〛 a1). [2: @hide_prf cases(mem_map ????? (labels_of_trace_are_in_code … (chop_mem … prf))) * #lbl' #pc * #Hmem #EQ destruct >(proj1 … (static_analisys_ok … EQmap … Hmem)) @(proj2 … (static_analisys_ok … EQmap … Hmem)) ] #p #p' #prog #R #mT #map1 #EQmap #si #s1 #s2 #sn #t #measurable cases measurable #premeas * #s0 * #s3 * #ti0 * #t12 * #t3n * #l1 * #l2 * #prf1 * #prf2 **** #EQ destruct #Hl1 #Hl2 #silent_ti0 #silent_t3n #acts cases(actionlabel_ok … Hl1) #c1 #EQc1 >rewrite_in_dependent_map [2: >get_cost_label_append in ⊢ (??%?); >(get_cost_label_silent_is_empty … silent_ti0) in ⊢ (??%?); >get_cost_label_of_trace_tind in ⊢ (??%?); >get_cost_label_append in ⊢ (??%?); >get_cost_label_of_trace_tind in ⊢ (??%?); >(get_cost_label_silent_is_empty … silent_t3n) in ⊢ (??%?); >append_nil in ⊢ (??%?); >EQc1 in ⊢ (??%?); whd in ⊢ (??(??%)?); whd in ⊢ (??(??(???%))?); >chop_cons in ⊢ (??%?); [2: cases daemon] % |3: ] whd in ⊢ (???% → ?); @opt_safe_elim #act #EQact #EQacts2 >EQacts2 >(big_op_associative ? [?]) #a1 >act_op #HR letin actsl ≝ (dependent_map ????); @(static_dynamic_inv … EQmap … (t12 @ t_ind … prf2 (t_base …)) … actsl … HR) [ /6 width=6 by conj, ex_intro/ | cases daemon (* true *) | cases s1 in prf1 t12; [3: #st1] cases s0 [3,6,9: #st0] * [2: #H #EQ lapply(refl ? (FINAL p p')) generalize in match (FINAL p p') in ⊢ (??%? → %); #st' #EQst' #tr lapply prf2 lapply EQst' -EQst' cases tr [ #st'' #EQ2 >EQ2 * | #st'' #st''' #st'''' #l3 #ABS #_ #EQ2 >EQ2 in ABS; * ] | #H1 #H2 #_ whd in match get_pc; normalize nodelta <(proj1 ?? (static_analisys_ok … EQmap …)) [ normalize in ⊢ (??%%); >neutral_l @EQact | | @execute_mem_label_pc // ] | #H #EQ destruct #_ whd in match get_pc; normalize nodelta <(proj1 ?? (static_analisys_ok … EQmap …)) [ normalize in ⊢ (??%%); >neutral_l @EQact | | whd in EQc1 : (??%%); destruct lapply H @eqb_elim [2: #_ #EQ destruct] #EQ >EQ #_ @i_act_in_map ] ] | >rewrite_in_dependent_map [2: >get_cost_label_append in ⊢ (??%?); >get_cost_label_of_trace_tind in ⊢ (??%?); >append_nil in ⊢ (??%?); % |3:] % ] qed.