include "ASM/ASMCosts.ma". include "ASM/WellLabeled.ma". include "ASM/Status.ma". include "common/StructuredTraces.ma". (* definition next_instruction_is_labelled ≝ λcost_labels: BitVectorTrie Byte 16. λs: Status. let pc ≝ program_counter … (execute_1 s) in match lookup_opt … pc cost_labels with [ None ⇒ False | _ ⇒ True ]. definition current_instruction_cost ≝ λs: Status. \snd (fetch (code_memory … s) (program_counter … s)). definition is_call_p ≝ λs. match s with [ ACALL _ ⇒ True | LCALL _ ⇒ True | JMP ⇒ True (* XXX: is function pointer call *) | _ ⇒ False ]. definition next_instruction ≝ λs: Status. let 〈instruction, ignore, ignore〉 ≝ fetch (code_memory … s) (program_counter … s) in instruction. inductive trace_ends_with_ret: Type[0] ≝ | ends_with_ret: trace_ends_with_ret | doesnt_end_with_ret: trace_ends_with_ret. definition next_instruction_is_labelled ≝ λcost_labels: BitVectorTrie Byte 16. λs: Status. let pc ≝ program_counter … (execute 1 s) in match lookup_opt … pc cost_labels with [ None ⇒ False | _ ⇒ True ]. definition next_instruction_properly_relates_program_counters ≝ λbefore: Status. λafter : Status. let size ≝ current_instruction_cost before in let pc_before ≝ program_counter … before in let pc_after ≝ program_counter … after in let sum ≝ \snd (half_add … pc_before (bitvector_of_nat … size)) in sum = pc_after. definition current_instruction ≝ λs: Status. let pc ≝ program_counter … s in \fst (\fst (fetch … (code_memory … s) pc)). definition current_instruction_is_labelled ≝ λcost_labels: BitVectorTrie Byte 16. λs: Status. let pc ≝ program_counter … s in match lookup_opt … pc cost_labels with [ None ⇒ False | _ ⇒ True ]. inductive trace_label_return (cost_labels: BitVectorTrie Byte 16): Status → Status → Type[0] ≝ | tlr_base: ∀status_before: Status. ∀status_after: Status. trace_label_label cost_labels ends_with_ret status_before status_after → trace_label_return cost_labels status_before status_after | tlr_step: ∀status_initial: Status. ∀status_labelled: Status. ∀status_final: Status. trace_label_label cost_labels doesnt_end_with_ret status_initial status_labelled → trace_label_return cost_labels status_labelled status_final → trace_label_return cost_labels status_initial status_final with trace_label_label: trace_ends_with_ret → Status → Status → Type[0] ≝ | tll_base: ∀ends_flag: trace_ends_with_ret. ∀start_status: Status. ∀end_status: Status. trace_label_label_pre cost_labels ends_flag start_status end_status → current_instruction_is_labelled cost_labels start_status → trace_label_label cost_labels ends_flag start_status end_status with trace_label_label_pre: trace_ends_with_ret → Status → Status → Type[0] ≝ | tllp_base_not_return: ∀start_status: Status. let final_status ≝ execute 1 start_status in current_instruction start_status ≠ (RealInstruction … (RET [[relative]])) → ¬ (is_jump_p (current_instruction start_status)) → current_instruction_is_labelled cost_labels final_status → trace_label_label_pre cost_labels doesnt_end_with_ret start_status final_status | tllp_base_return: ∀start_status: Status. let final_status ≝ execute 1 start_status in current_instruction start_status = (RealInstruction … (RET [[relative]])) → trace_label_label_pre cost_labels ends_with_ret start_status final_status | tllp_base_jump: ∀start_status: Status. let final_status ≝ execute 1 start_status in is_jump_p (current_instruction start_status) → current_instruction_is_labelled cost_labels final_status → trace_label_label_pre cost_labels doesnt_end_with_ret start_status final_status | tllp_step_call: ∀end_flag: trace_ends_with_ret. ∀status_pre_fun_call: Status. ∀status_after_fun_call: Status. ∀status_final: Status. let status_start_fun_call ≝ execute 1 status_pre_fun_call in is_call_p (current_instruction status_pre_fun_call) → next_instruction_properly_relates_program_counters status_pre_fun_call status_after_fun_call → trace_label_return cost_labels status_start_fun_call status_after_fun_call → trace_label_label_pre cost_labels end_flag status_after_fun_call status_final → trace_label_label_pre cost_labels end_flag status_pre_fun_call status_final | tllp_step_default: ∀end_flag: trace_ends_with_ret. ∀status_pre: Status. ∀status_end: Status. let status_init ≝ execute 1 status_pre in trace_label_label_pre cost_labels end_flag status_init status_end → ¬ (is_call_p (current_instruction status_pre)) → ¬ (is_jump_p (current_instruction status_pre)) → (current_instruction status_pre) ≠ (RealInstruction … (RET [[relative]])) → ¬ (current_instruction_is_labelled cost_labels status_init) → trace_label_label_pre cost_labels end_flag status_pre status_end. *) (* XXX: not us definition compute_simple_path0 (s:Status) (is_labelled s) (well_balanced s) (labelled_p p): trace_lab_ret ≝ …. lemma simple_path_ok: ∀st: Status.∀H. let p ≝ compute_simple_path0 st H in ∀n. nth_path n p = execute n st ∧ (execute' (path_length p) st = 〈st',τ〉 → τ = cost_trace p) ]. *) let rec compute_max_trace_label_label_cost (cost_labels: BitVectorTrie costlabel 16) (trace_ends_flag: trace_ends_with_ret) (start_status: Status) (final_status: Status) (the_trace: trace_label_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status) on the_trace: nat ≝ match the_trace with [ tll_base ends_flag initial final given_trace labelled_proof ⇒ compute_max_trace_any_label_cost … given_trace ] and compute_max_trace_any_label_cost (cost_labels: BitVectorTrie costlabel 16) (trace_ends_flag: trace_ends_with_ret) (start_status: Status) (final_status: Status) (the_trace: trace_any_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status) on the_trace: nat ≝ match the_trace with [ tal_base_not_return the_status _ _ _ _ ⇒ current_instruction_cost the_status | tal_base_return the_status _ _ _ ⇒ current_instruction_cost the_status | tal_step_call end_flag pre_fun_call start_fun_call after_fun_call final _ _ _ call_trace final_trace ⇒ let current_instruction_cost ≝ current_instruction_cost pre_fun_call in let call_trace_cost ≝ compute_max_trace_label_return_cost … call_trace in let final_trace_cost ≝ compute_max_trace_any_label_cost cost_labels end_flag … final_trace in call_trace_cost + current_instruction_cost + final_trace_cost | tal_step_default end_flag status_pre status_init status_end _ tail_trace _ _ ⇒ let current_instruction_cost ≝ current_instruction_cost status_pre in let tail_trace_cost ≝ compute_max_trace_any_label_cost cost_labels end_flag status_init status_end tail_trace in current_instruction_cost + tail_trace_cost ] and compute_max_trace_label_return_cost (cost_labels: BitVectorTrie costlabel 16) (start_status: Status) (final_status: Status) (the_trace: trace_label_return (ASM_abstract_status cost_labels) start_status final_status) on the_trace: nat ≝ match the_trace with [ tlr_base before after trace_to_lift ⇒ compute_max_trace_label_label_cost … trace_to_lift | tlr_step initial labelled final labelled_trace ret_trace ⇒ let labelled_cost ≝ compute_max_trace_label_label_cost … labelled_trace in let return_cost ≝ compute_max_trace_label_return_cost … ret_trace in labelled_cost + return_cost ]. (*Useless now? (* To be moved *) lemma pred_minus_1: ∀m, n: nat. ∀proof: n < m. pred (m - n) = m - n - 1. #m #n cases m [ #proof cases(lt_to_not_zero … proof) | #m' #proof normalize in ⊢ (???%); cases n [ normalize // | #n' normalize cases(m' - n') [ % | #Sm_n' normalize // ] ] ] qed. lemma succ_m_plus_one: ∀m: nat. S m = m + 1. // qed.*) include alias "arithmetics/nat.ma". (* lemma minus_m_n_minus_minus_plus_1: ∀m, n: nat. ∀proof: n < m. m - n = (m - n - 1) + 1. /3 by lt_plus_to_minus_r, plus_minus/ qed. *) lemma plus_minus_rearrangement_1: ∀m, n, o: nat. ∀proof_n_m: n ≤ m. ∀proof_m_o: m ≤ o. (m - n) + (o - m) = o - n. #m #n #o #H1 #H2 lapply (minus_to_plus … H1 (refl …)) #K1 >K1 lapply (minus_to_plus … H2 (refl …)) #K2 >K2 /2 by plus_minus/ qed. lemma plus_minus_rearrangement_2: ∀m, n, o: nat. n ≤ m → o ≤ n → (m - n) + (n - o) = m - o. /2 by plus_minus_rearrangement_1/ qed. axiom current_instruction_cost_non_zero: ∀s: Status. current_instruction_cost s > 0. lemma le_monotonic_plus: ∀m, n, o: nat. m ≤ n → m + o ≤ n + o. #m #n #o #hyp /2/ qed. lemma le_plus_to_minus_l: ∀m, n, o: nat. n ≤ o → m ≤ o - n → m + n ≤ o. #m #n #o #le_hyp #hyp generalize in match (le_monotonic_plus m (o - n) n hyp); generalize in match (plus_minus_m_m o n le_hyp); #assm 0 → n ≤ m. (* #m #n #o #eq_hyp le_monot *) let rec compute_max_trace_label_label_cost_is_ok (cost_labels: BitVectorTrie costlabel 16) (trace_ends_flag: trace_ends_with_ret) (start_status: Status) (final_status: Status) (the_trace: trace_label_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status) on the_trace: And (compute_max_trace_label_label_cost cost_labels trace_ends_flag start_status final_status the_trace = (clock … final_status) - (clock … start_status)) (clock … start_status ≤ clock … final_status) ≝ ? and compute_max_trace_any_label_cost_is_ok (cost_labels: BitVectorTrie costlabel 16) (trace_ends_flag: trace_ends_with_ret) (start_status: Status) (final_status: Status) (the_trace: trace_any_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status) on the_trace: (compute_max_trace_any_label_cost cost_labels trace_ends_flag start_status final_status the_trace = (clock … final_status) - (clock … start_status)) ∧ (clock … start_status ≤ clock … final_status) ≝ ? and compute_max_trace_label_return_cost_is_ok (cost_labels: BitVectorTrie costlabel 16) (start_status: Status) (final_status: Status) (the_trace: trace_label_return (ASM_abstract_status cost_labels) start_status final_status) on the_trace: (compute_max_trace_label_return_cost cost_labels start_status final_status the_trace = (clock … final_status) - (clock … start_status)) ∧ (clock … start_status ≤ clock … final_status) ≝ ?. [1: cases the_trace #ends_flag #start_status #end_status #any_label_trace #is_costed normalize @compute_max_trace_any_label_cost_is_ok |2: cases the_trace [1: #start_status #final_status #is_next #is_not_return #is_costed change with (current_instruction_cost start_status) in ⊢ (?(??%?)?); cases(is_next) @conj [1: cases(execute_1 start_status) #the_status #assm whd in match eject; normalize nodelta >assm % |2: cases(execute_1 start_status) #the_status #assm whd in match eject; normalize nodelta @(minus_m_n_gt_o_O_le_n_m ? ? (current_instruction_cost start_status)) [1: assumption |2: @(current_instruction_cost_non_zero start_status) ] ] |2: #start_status #final_status #is_next #is_return change with (current_instruction_cost start_status) in ⊢ (?(??%?)?); cases(is_next) @conj [1: cases(execute_1 start_status) #the_status #assm whd in match eject; normalize nodelta >assm % |2: cases(execute_1 start_status) #the_status #assm whd in match eject; normalize nodelta @(minus_m_n_gt_o_O_le_n_m ? ? (current_instruction_cost start_status)) [1: assumption |2: @current_instruction_cost_non_zero ] ] |3: #end_flag #status_pre_fun_call #status_start_fun_call #status_after_fun_call #status_final #is_next #is_call #is_after_return #call_trace #final_trace change with ( let current_instruction_cost ≝ current_instruction_cost status_pre_fun_call in let call_trace_cost ≝ compute_max_trace_label_return_cost … call_trace in let final_trace_cost ≝ compute_max_trace_any_label_cost cost_labels end_flag … final_trace in call_trace_cost + current_instruction_cost + final_trace_cost) in ⊢ (?(??%?)?); normalize nodelta; @conj [1: cases(compute_max_trace_label_return_cost_is_ok cost_labels status_start_fun_call status_after_fun_call call_trace) #trace_label_return_eq #trace_label_return_le cases(compute_max_trace_any_label_cost_is_ok cost_labels end_flag status_after_fun_call status_final final_trace) #trace_any_label_eq #trace_any_label_le >trace_label_return_eq >trace_any_label_eq >commutative_plus cut ((clock (BitVectorTrie Byte 16) status_after_fun_call -clock (BitVectorTrie Byte 16) status_start_fun_call +current_instruction_cost status_pre_fun_call +(clock (BitVectorTrie Byte 16) status_final -clock (BitVectorTrie Byte 16) status_after_fun_call) = ((clock (BitVectorTrie Byte 16) status_after_fun_call - clock (BitVectorTrie Byte 16) status_start_fun_call) + (clock (BitVectorTrie Byte 16) status_final - clock (BitVectorTrie Byte 16) status_after_fun_call)) + current_instruction_cost status_pre_fun_call)) [ 1: cases daemon (* XXX: complete me *) | 2: #eq_hyp >eq_hyp >plus_minus_rearrangement_1 [1: |2, 3: assumption ] ] |2: cases daemon (* XXX: complete me *) ] |4: #end_flag #status_pre #status_init #status_end #is_next #trace_any_label #is_other #is_not_costed cases daemon (* XXX: complete me *) ] |3: cases the_trace [1: #status_before #status_after #trace_to_lift normalize @compute_max_trace_label_label_cost_is_ok |2: #status_initial #status_labelled #status_final #labelled_trace #ret_trace normalize cases(compute_max_trace_label_label_cost_is_ok cost_labels doesnt_end_with_ret status_initial status_labelled labelled_trace); #label_label_trace_equality #label_label_trace_leq cases(compute_max_trace_label_return_cost_is_ok cost_labels status_labelled status_final ret_trace); #label_return_trace_equality #label_return_trace_leq >label_label_trace_equality >label_return_trace_equality @conj [2: @(transitive_le (clock (BitVectorTrie (Vector bool 8) 16) status_initial) (clock (BitVectorTrie Byte 16) status_labelled) (clock (BitVectorTrie (Vector bool 8) 16) status_final)) |1: @plus_minus_rearrangement_1 ] assumption ] ] qed. (* XXX: work below here: *) let rec compute_paid_trace_any_label (cost_labels: BitVectorTrie costlabel 16) (trace_ends_flag: trace_ends_with_ret) (start_status: Status) (final_status: Status) (the_trace: trace_any_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status) on the_trace: nat ≝ match the_trace with [ tal_base_not_return the_status _ _ _ _ ⇒ current_instruction_cost the_status | tal_base_return the_status _ _ _ ⇒ current_instruction_cost the_status | tal_step_call end_flag pre_fun_call start_fun_call after_fun_call final _ _ _ call_trace final_trace ⇒ let current_instruction_cost ≝ current_instruction_cost pre_fun_call in let final_trace_cost ≝ compute_paid_trace_any_label cost_labels end_flag … final_trace in current_instruction_cost + final_trace_cost | tal_step_default end_flag status_pre status_init status_end _ tail_trace _ _ ⇒ let current_instruction_cost ≝ current_instruction_cost status_pre in let tail_trace_cost ≝ compute_paid_trace_any_label cost_labels end_flag status_init status_end tail_trace in current_instruction_cost + tail_trace_cost ]. definition compute_paid_trace_label_label ≝ λcost_labels: BitVectorTrie costlabel 16. λtrace_ends_flag: trace_ends_with_ret. λstart_status: Status. λfinal_status: Status. λthe_trace: trace_label_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status. match the_trace with [ tll_base ends_flag initial final given_trace labelled_proof ⇒ compute_paid_trace_any_label … given_trace ]. let rec compute_trace_label_label_cost_using_paid (cost_labels: BitVectorTrie costlabel 16) (trace_ends_flag: trace_ends_with_ret) (start_status: Status) (final_status: Status) (the_trace: trace_label_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status) on the_trace: nat ≝ match the_trace with [ tll_base ends_flag initial final given_trace labelled_proof ⇒ compute_paid_trace_label_label cost_labels … the_trace + compute_trace_any_label_cost_using_paid … given_trace ] and compute_trace_any_label_cost_using_paid (cost_labels: BitVectorTrie costlabel 16) (trace_ends_flag: trace_ends_with_ret) (start_status: Status) (final_status: Status) (the_trace: trace_any_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status) on the_trace: nat ≝ match the_trace with [ tal_base_not_return the_status _ _ _ _ ⇒ 0 | tal_base_return the_status _ _ _ ⇒ 0 | tal_step_call end_flag pre_fun_call start_fun_call after_fun_call final _ _ _ call_trace final_trace ⇒ let call_trace_cost ≝ compute_trace_label_return_cost_using_paid … call_trace in let final_trace_cost ≝ compute_trace_any_label_cost_using_paid cost_labels end_flag … final_trace in call_trace_cost + final_trace_cost | tal_step_default end_flag status_pre status_init status_end _ tail_trace _ _ ⇒ compute_trace_any_label_cost_using_paid cost_labels end_flag status_init status_end tail_trace ] and compute_trace_label_return_cost_using_paid (cost_labels: BitVectorTrie costlabel 16) (start_status: Status) (final_status: Status) (the_trace: trace_label_return (ASM_abstract_status cost_labels) start_status final_status) on the_trace: nat ≝ match the_trace with [ tlr_base before after trace_to_lift ⇒ compute_trace_label_label_cost_using_paid … trace_to_lift | tlr_step initial labelled final labelled_trace ret_trace ⇒ let labelled_cost ≝ compute_trace_label_label_cost_using_paid … labelled_trace in let return_cost ≝ compute_trace_label_return_cost_using_paid … ret_trace in labelled_cost + return_cost ]. let rec compute_trace_label_label_cost_using_paid_ok (cost_labels: BitVectorTrie costlabel 16) (trace_ends_flag: trace_ends_with_ret) (start_status: Status) (final_status: Status) (the_trace: trace_label_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status) on the_trace: compute_trace_label_label_cost_using_paid cost_labels … the_trace = compute_max_trace_label_label_cost … the_trace ≝ ? and compute_trace_any_label_cost_using_paid_ok (cost_labels: BitVectorTrie costlabel 16) (trace_ends_flag: trace_ends_with_ret) (start_status: Status) (final_status: Status) (the_trace: trace_any_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status) on the_trace: compute_paid_trace_any_label cost_labels trace_ends_flag … the_trace +compute_trace_any_label_cost_using_paid cost_labels trace_ends_flag … the_trace =compute_max_trace_any_label_cost cost_labels trace_ends_flag … the_trace ≝ ? and compute_trace_label_return_cost_using_paid_ok (cost_labels: BitVectorTrie costlabel 16) (start_status: Status) (final_status: Status) (the_trace: trace_label_return (ASM_abstract_status cost_labels) start_status final_status) on the_trace: compute_trace_label_return_cost_using_paid cost_labels … the_trace = compute_max_trace_label_return_cost cost_labels … the_trace ≝ ?. [ cases the_trace #endsf #ss #es #tr #H normalize @compute_trace_any_label_cost_using_paid_ok | cases the_trace [ #ss #fs #H1 #H2 #H3 whd in ⊢ (??(?%%)%); commutative_plus in ⊢ (???(?%?)); >commutative_plus in ⊢ (??(??%)?); >associative_plus >associative_plus in ⊢ (???%); @eq_f2 try % associative_plus @eq_f2 [ % | @compute_trace_any_label_cost_using_paid_ok ] ] | cases the_trace [ #sb #sa #tr normalize @compute_trace_label_label_cost_using_paid_ok | #si #sl #sf #tr1 #tr2 normalize @eq_f2 [ @compute_trace_label_label_cost_using_paid_ok | @compute_trace_label_return_cost_using_paid_ok ]]] qed. (* let rec compute_paid_trace_label_return (cost_labels: BitVectorTrie costlabel 16) (start_status: Status) (final_status: Status) (the_trace: trace_label_return (ASM_abstract_status cost_labels) start_status final_status) on the_trace: nat ≝ match the_trace with [ tlr_base before after trace_to_lift ⇒ compute_paid_trace_label_label … trace_to_lift | tlr_step initial labelled final labelled_trace ret_trace ⇒ let labelled_cost ≝ compute_paid_trace_label_label … labelled_trace in let return_cost ≝ compute_paid_trace_label_return … ret_trace in labelled_cost + return_cost ]. *) let rec compute_cost_trace_label_label (cost_labels: BitVectorTrie costlabel 16) (trace_ends_flag: trace_ends_with_ret) (start_status: Status) (final_status: Status) (the_trace: trace_label_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status) on the_trace: list (Σk:costlabel. ∃pc. lookup_opt … pc cost_labels = Some … k) ≝ match the_trace with [ tll_base ends_flag initial final given_trace labelled_proof ⇒ let pc ≝ program_counter … initial in let label ≝ match lookup_opt … pc cost_labels return λx. match x with [None ⇒ False | Some _ ⇒ True] → costlabel with [ None ⇒ λabs. ⊥ | Some l ⇒ λ_. l ] labelled_proof in (dp … label ?)::compute_cost_trace_any_label … given_trace ] and compute_cost_trace_any_label (cost_labels: BitVectorTrie costlabel 16) (trace_ends_flag: trace_ends_with_ret) (start_status: Status) (final_status: Status) (the_trace: trace_any_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status) on the_trace: list (Σk:costlabel. ∃pc. lookup_opt … pc cost_labels = Some … k) ≝ match the_trace with [ tal_base_not_return the_status _ _ _ _ ⇒ [] | tal_base_return the_status _ _ _ ⇒ [] | tal_step_call end_flag pre_fun_call start_fun_call after_fun_call final _ _ _ call_trace final_trace ⇒ let call_cost_trace ≝ compute_cost_trace_label_return … call_trace in let final_cost_trace ≝ compute_cost_trace_any_label cost_labels end_flag … final_trace in call_cost_trace @ final_cost_trace | tal_step_default end_flag status_pre status_init status_end _ tail_trace _ _ ⇒ compute_cost_trace_any_label cost_labels end_flag status_init status_end tail_trace ] and compute_cost_trace_label_return (cost_labels: BitVectorTrie costlabel 16) (start_status: Status) (final_status: Status) (the_trace: trace_label_return (ASM_abstract_status cost_labels) start_status final_status) on the_trace: list (Σk:costlabel. ∃pc. lookup_opt … pc cost_labels = Some … k) ≝ match the_trace with [ tlr_base before after trace_to_lift ⇒ compute_cost_trace_label_label … trace_to_lift | tlr_step initial labelled final labelled_trace ret_trace ⇒ let labelled_cost ≝ compute_cost_trace_label_label … labelled_trace in let return_cost ≝ compute_cost_trace_label_return … ret_trace in labelled_cost @ return_cost ]. [ %{pc} whd in match label; generalize in match labelled_proof; whd in ⊢ (% → ?); cases (lookup_opt costlabel … pc cost_labels) normalize [ #abs cases abs | // ] | // ] qed. (* ??????????????????????? *) axiom block_cost_static_dynamic_ok: ∀cost_labels: BitVectorTrie costlabel 16. ∀trace_ends_flag. ∀start_status: Status. ∀final_status: Status. ∀the_trace: trace_label_label (ASM_abstract_status cost_labels) trace_ends_flag start_status final_status. let mem ≝ code_memory … start_status in let pc ≝ program_counter … start_status in let program_size ≝ 2^16 in block_cost mem cost_labels pc program_size = compute_paid_trace_label_label cost_labels trace_ends_flag start_status final_status the_trace. (* (* To be moved elsewhere*) lemma le_S_n_m_to_le_n_m: ∀n,m. S n ≤ m → n ≤ m. #n #m #H change with (pred (S n) ≤ m) @(transitive_le ? (S n)) // qed. *) include "arithmetics/bigops.ma". (* This shoudl go in bigops! *) theorem bigop_sum_rev: ∀k1,k2,p,B,nil.∀op:Aop B nil.∀f:nat→B. \big[op,nil]_{ibigop_sum >commutative_plus @same_bigop #i @leb_elim normalize [2,4: // | #H1 #H2 (lookup_lookup_present … k_pres) % #abs destruct (abs) qed. lemma ltb_rect: ∀P:Type[0].∀n,m. (n < m → P) → (¬ n < m → P) → P. #P #n #m lapply (refl … (ltb n m)) cases (ltb n m) in ⊢ (???% → %); #E #H1 #H2 [ @H1 @leb_true_to_le @E | @H2 @leb_false_to_not_le @E ] qed. lemma same_ltb_rect: ∀P,n,m,H1,H2,n',m',H1',H2'. ltb n m = ltb n' m' → (∀x,y. H1 x = H1' y) → (∀x,y. H2 x = H2' y) → ltb_rect P n m H1 H2 = ltb_rect P n' m' H1' H2'. #P #n #m #H1 #H2 #n' #m' #H1' #H2' #E #K1 #K2 whd in ⊢ (??%?); cut (∀xxx,yyy,xxx',yyy'. match ltb n m return λx:bool. eq bool (ltb n m) x → (lt n m → P) → (Not (lt n m) → P) → P with [ true ⇒ λE0:eq bool (ltb n m) true. λH10:lt n m → P. λH20:Not (lt n m) → P. H10 (xxx E0) | false ⇒ λE0:eq bool (ltb n m) false. λH10:lt n m → P. λH20:Not (lt n m) → P. H20 (yyy E0)] (refl … (ltb n m)) H1 H2 = match ltb n' m' return λx:bool. eq bool (ltb n' m') x → (lt n' m' → P) → (Not (lt n' m') → P) → P with [ true ⇒ λE0:eq bool (ltb n' m') true. λH10:lt n' m' → P. λH20:Not (lt n' m') → P. H10 (xxx' E0) | false ⇒ λE0:eq bool (ltb n' m') false. λH10:lt n' m' → P. λH20:Not (lt n' m') → P. H20 (yyy' E0)] (refl … (ltb n' m')) H1' H2' ) [2: #X @X] >E cases (ltb n' m') #xxx #yyy #xxx' #yyy' normalize [ @K1 | @K2 ] qed. definition tech_cost_of_label: ∀cost_labels: BitVectorTrie costlabel 16. ∀cost_map: identifier_map CostTag nat. ∀codom_dom: (∀pc,k. lookup_opt … pc cost_labels = Some … k → present … cost_map k). list (Σk:costlabel.∃pc. lookup_opt costlabel 16 pc cost_labels = Some ? k) → nat → nat ≝ λcost_labels,cost_map,codom_dom,ctrace,i. ltb_rect ? i (|ctrace|) (λH. lookup_present ?? cost_map (nth_safe ? i ctrace H) ?) (λ_.0). @tech_cost_of_label0 @codom_dom qed. lemma shift_nth_safe: ∀T,i,l2,l1,K1,K2. nth_safe T i l1 K1 = nth_safe T (i+|l2|) (l2@l1) K2. #T #i #l2 elim l2 normalize [ #l1 #K1 le_to_leb_true [//] change with (? < ?) >append_length >commutative_plus @monotonic_lt_plus_r // | #H1 change with (¬ i < ?) in H1; >not_le_to_leb_false [//] >append_length >commutative_plus % #H2 change with (? < ?) in H2; @(absurd ?? H1) @(lt_plus_to_lt_r … H2) ] | #E whd in match tech_cost_of_label; normalize nodelta @same_ltb_rect [@E] [ #K1 #K2 lapply (shift_nth_safe ? i l1 l2 K1 K2) #H check eq_f lapply (eq_f ?? (lookup_present CostTag nat cost_Map) ?? H) cut (decide_bool (ltb i (|l2|)) = decide_bool (ltb (i+(|l1|)) (|l1@l2|))) whd in match tech_cost_of_label; normalize nodelta let rec compute_max_trace_label_return_cost_ok_with_trace (cost_labels: BitVectorTrie costlabel 16) (cost_map: identifier_map CostTag nat) (initial: Status) (final: Status) (trace: trace_label_return (ASM_abstract_status cost_labels) initial final) (codom_dom: (∀pc,k. lookup_opt … pc cost_labels = Some … k → present … cost_map k)) (dom_codom:(∀k. ∀k_pres:present … cost_map k. ∃pc. lookup_opt … pc cost_labels = Some … k → block_cost (code_memory … initial) cost_labels pc 2^16 = lookup_present … k_pres)) on trace: let ctrace ≝ compute_cost_trace_label_return … trace in compute_max_trace_label_return_cost … trace = Σ_{i < |ctrace|} (tech_cost_of_label cost_labels cost_map codom_dom ctrace i) ≝ ?. lapply (refl … initial) cases trace in ⊢ (??%? → %); [ #sb #sa #tr #_ whd in ⊢ (let ctrace ≝ % in ??%?); | #si #sl #sf #tr1 #tr2 #E whd in ⊢ (let ctrace ≝ % in ??%?); change with (let ctrace ≝ ? in ? = bigop (|? @ ?|) ?????) >append_length in ⊢ (let ctrace ≝ ? in ???(?%?????)); change with (?=?) >bigop_sum_rev >commutative_plus @eq_f2 [ >compute_max_trace_label_return_cost_ok_with_trace -compute_max_trace_label_return_cost_ok_with_trace [3: @cost_map |2: @codom_dom |4:lapply (code_memory_ro_label_label … tr1) >E #E2