include "compiler.ma". include "ASM/Interpret2.ma". include "Clight/Cexec.ma". include "Clight/labelSimulation.ma". theorem correct : ∀input_program. ∀object_code,costlabel_map,labelled,cost_map. compile input_program = OK ? 〈〈object_code,costlabel_map〉,❬labelled,cost_map❭〉 → not_wrong … (exec_inf … clight_fullexec input_program) → sim_with_labels (exec_inf … clight_fullexec input_program) (exec_inf … clight_fullexec labelled) ∧ True (* TODO *). #input_program #object_code #costlabel_map #labelled #cost_map #COMPILE #NOT_WRONG cases (bind_inversion ????? COMPILE) -COMPILE * * #init_cost #labelled' #rtlabs_program * #FRONTEND #COMPILE cases (bind_inversion ????? COMPILE) -COMPILE * #object_code' #costlabel_map' * #ASSEMBLER #COMPILE whd in COMPILE:(??%%); destruct cases (bind_inversion ????? FRONTEND) -FRONTEND #cminor_program * #CMINOR #FRONTEND whd in FRONTEND:(??%%); destruct % [ (* Needs switch removal too, now @labelling_sim @NOT_WRONG *) cases daemon | @I ] qed. include "Clight/Clight_abstract.ma". include "common/Measurable.ma". (* We could restrict this function to identifiers that are function names in the program and lift it (like the lift_cost_map_back_to_front function), but let's go with the easier notion of having a total map and ignore all the extra stuff. *) definition stack_cost_T ≝ ident → nat. (* Again, in principle we can use the fact that the state must be a call and also show some relevant invariants about call states, but we go with the easier option of returning zero elsewhere. *) definition Clight_stack_cost : stack_cost_T → cl_genv → ∀s:Clight_state. match Clight_classify s with [ cl_call ⇒ True | _ ⇒ False ] → nat ≝ λcost,ge,s,H. match s with [ Callstate v _ _ _ _ ⇒ match v with [ Vptr p ⇒ match symbol_for_block … ge (pblock p) with [ Some id ⇒ cost id | _ ⇒ 0 ] | _ ⇒ 0 ] | _ ⇒ 0 ]. definition Clight_pcs : preclassified_system ≝ mk_preclassified_system clight_fullexec (λ_.Clight_labelled) (λ_.Clight_classify) Clight_stack_cost. (* From measurable on Clight, we will end up with an RTLabs flat trace where we know that there are some m' and n' such that the prefix in Clight matches the prefix in RTLabs given by m', the next n steps in Clight are equivalent to the n' steps in RTLabs, and we have a suitable "will_return" for RTLabs for those n' steps so that we can build a corresponding structured trace. "Equivalent" here means, in particular, that the observables will be the same, and those observables will include the stack space costs. *) axiom observables_8051 : object_code → nat → nat → option ((list trace) × (list trace)). definition in_execution_prefix : execution_prefix Clight_state → costlabel → Prop ≝ λx,l. Exists … (λtrs. Exists … (λev. ev = EVcost l) (\fst trs)) x. let rec foldl_exists_aux (A,B:Type[0]) (l,l':list B) (f:A → ∀b:B. Exists … (λx.x=b) l → A) (a:A) on l' : (∀b. Exists … (λx.x=b) l' → Exists … (λx.x=b) l) → A ≝ match l' return λl'. (∀b. Exists … (λx.x=b) l' → Exists … (λx.x=b) l) → A with [ nil ⇒ λ_. a | cons h t ⇒ λH. foldl_exists_aux A B l t f (f a h (H …)) ? ]. [ %1 % | #b #H' @H %2 @H' ] qed. definition foldl_exists : ∀A,B:Type[0]. ∀l:list B. (A → ∀b:B. Exists … (λx. x = b ) l → A) → A → A ≝ λA,B,l,f,a. foldl_exists_aux A B l l f a (λb,H. H). lemma Exists_lift : ∀A,P,Q,l. (∀x. P x → Q x) → Exists A P l → Exists A Q l. #A #P #Q #l elim l [ // | #h #t #IH #H * [ #H' %1 @H @H' | #H' %2 @IH /2/ ] ] qed. definition measure_clock : ∀x:execution_prefix Clight_state. ((Σl:costlabel.in_execution_prefix x l)→ℕ) → nat ≝ λx,costmap. foldl_exists … x (λclock,trs,H. foldl_exists … (\fst trs) (λclock,ev. match ev return λev. Exists … (λx. x=ev) ? → nat with [ EVcost l ⇒ λH'. clock + costmap «l,?» | _ ⇒ λ_. clock ]) clock) 0. whd @(Exists_lift … H) * #tr1 #s1 #E destruct @(Exists_lift … H') #ev1 #E @E qed. definition clight_clock_after : ∀p:clight_program. nat → ((Σl:costlabel.in_clight_program p l)→ℕ) → option nat ≝ λp,n,costmap. let x ≝ exec_inf … clight_fullexec p in match split_trace … x n with [ Some traces ⇒ Some ? (measure_clock (\fst traces) (λl.costmap «l,?»)) | None ⇒ None ? ]. cases daemon qed. axiom initial_8051_status : ∀oc. Status oc. definition simulates ≝ λstack_cost : stack_cost_T. λ stack_bound, labelled, object_code, cost_map. let initial_status ≝ initial_8051_status (load_code_memory object_code) in ∀m1,m2. measurable Clight_pcs labelled m1 m2 stack_cost stack_bound → ∀c1,c2. clight_clock_after labelled m1 cost_map = Some ? c1 → clight_clock_after labelled m2 cost_map = Some ? c2 → ∃n1,n2. observables clight_fullexec labelled m1 m2 = observables_8051 object_code n1 n2 ∧ c2 - c1 = clock … (execute n2 ? initial_status) - clock … (execute n1 ? initial_status). axiom compile' : clight_program → res (object_code × costlabel_map × (𝚺labelled:clight_program. ((Σl:costlabel.in_clight_program labelled l)→ℕ)) × stack_cost_T × nat). theorem correct' : ∀input_program. not_wrong … (exec_inf … clight_fullexec input_program) → ∀object_code,costlabel_map,labelled,cost_map,stack_cost,stack_bound. compile' input_program = OK ? 〈〈〈object_code,costlabel_map〉,❬labelled,cost_map❭〉,stack_cost,stack_bound〉 → sim_with_labels (exec_inf … clight_fullexec input_program) (exec_inf … clight_fullexec labelled) ∧ simulates stack_cost stack_bound labelled object_code cost_map. (* start of old simulates (* [nth_state_of_with_stack state stack_cost stack_bound exec n] returns [Some s] iff after [n] steps of [exec] we have reached [s] without exceeding the [stack_bound] according to the [stack_cost] function. *) axiom nth_state_of_with_stack : ∀state. (state → nat) → nat → execution state io_out io_in → nat → option state. axiom nth_state_of : ∀state. execution state io_out io_in → nat → option state. let cl_trace ≝ exec_inf … clight_fullexec labelled in let asm_trace ≝ exec_inf … ASM_fullexec object_code in not_wrong ? cl_trace → ∀n,s. nth_state_of_with_stack ? stack_cost stack_bound cl_trace n = Some ? s → 𝚺m,s'. nth_state_of ? asm_trace m = Some ? s' ∧ s ≃ s' *) (* TODO ∀input_program. ! 〈object_code,costlabel_map,labelled,cost_map〉 ← compile input_program exec_inf … clight_fullexec input_program ≃l exec_inf … clight_fullexec labelled ∧ exec_inf … clight_fullexec labelled ≈ exec_inf … ASM_fullexec object_code (* Should we be lifting labels in some way here? *) ∧ ∀i,f : clight_status. Clight_labelled i → Clight_labelled f → ∀mx,time. let trace ≝ exec_inf_aux … clight_fullexec labelled i in will_return O O mx time f trace → mx < max_allowed_stack → ∃!i',f'. i ≃ i' ∧ f ≃ f' ∧ i' 8051~> f' ∧ time = clock f' - clock i'. ∀s,flat. let ge ≝ (globalenvs … labelled) in subtrace_of (exec_inf … RTLabs_fullexec labelled) flat → RTLabs_cost s = true → ∀WR : will_return ge 0 s flat. let structured_trace_rtlabs ≝ make_label_return' ge 0 s flat ??? WR in let labels_rtlabs ≝ flat_label_trace … flat WR in ∃!initial,final,structured_trace_asm. structured_trace_rtlabs ≈ structured_trace_asm ∧ clock … code_memory … final = clock … code_memory … initial + (Σ_{i < |labels_rtlabs|} (cost_map (match nth i labels_rtlabs with [ Some k ⇒ k | None ⇒ 0 ])). What is ≃l? Must show that "labelled" does everything that "input_program" does, without getting lost in some non-terminating loop part way. *)