Library mcertikos.proc.VMXIntroGenPassthrough
This file provide the contextual refinement proof between MAL layer and MPTIntro layer
Section Refinement.
Context `{real_params: RealParams}.
Context `{multi_oracle_prop: MultiOracleProp}.
Section WITHMEM.
Context `{Hstencil: Stencil}.
Context `{Hmem: Mem.MemoryModelX}.
Context `{Hmwd: UseMemWithData mem}.
Global Instance cpu_id_range: high_level_invariant_impl_CPU_ID_range.
Proof.
econstructor; eauto.
intros.
inv H.
assumption.
Qed.
Global Instance: (LoadStoreProp (hflatmem_store:= flatmem_store) (lflatmem_store:= flatmem_store)).
Proof.
accessor_prop_tac.
- eapply flatmem_store_exists; eauto.
- eapply flatmem_store_match; eauto.
Qed.
Lemma passthrough_correct:
sim (crel HDATA LDATA) vmxintro_passthrough vmcsinit.
Proof.
sim_oplus.
- apply fload_sim.
- apply fstore_sim.
- apply vmxinfo_get_sim.
- apply palloc_sim.
- apply setPT_sim.
- apply ptRead_sim.
- apply ptResv_sim.
- apply shared_mem_status_sim.
- apply offer_shared_mem_sim.
- apply biglow_thread_wakeup_sim.
- apply biglow_thread_yield_sim.
- apply biglow_thread_sleep_sim.
- apply biglow_sched_init_sim.
- apply uctx_get_sim.
- apply uctx_set_sim.
- apply biglow_proc_create_sim.
- apply container_get_nchildren_sim.
- apply container_get_quota_sim.
- apply container_get_usage_sim.
- apply container_can_consume_sim.
- apply get_CPU_ID_sim.
- apply get_curid_sim.
- apply acquire_lock_SC_sim.
- apply release_lock_SC_sim.
- apply get_sync_chan_busy_sim.
- apply set_sync_chan_busy_sim.
- apply ipc_send_body_sim.
- apply ipc_receive_body_sim.
- apply ept_invalidate_mappings_sim.
- apply ept_add_mapping_sim.
- apply rdmsr_sim.
- apply wrmsr_sim.
- apply vmcs_readz_sim.
- apply vmcs_writez_sim.
- apply vmptrld_sim.
- apply vmx_set_intercept_intwin_sim.
- apply vmx_set_desc1_sim.
- apply vmx_set_desc2_sim.
- apply vmx_inject_event_sim.
- apply vmx_set_tsc_offset_sim.
- apply vmx_get_tsc_offset_sim.
- apply vmx_get_exit_reason_sim.
- apply vmx_get_exit_fault_addr_sim.
- apply vmx_check_pending_event_sim.
- apply vmx_check_int_shadow_sim.
- apply vmcs_set_defaults_sim.
- apply cli_sim.
- apply sti_sim.
- apply serial_intr_disable_sim.
- apply serial_intr_enable_sim.
- apply serial_putc_sim.
- apply cons_buf_read_sim.
- apply hostin_sim.
- apply hostout_sim.
- apply proc_create_postinit_sim.
- apply trap_info_get_sim.
- apply trap_info_ret_sim.
- apply proc_start_user_sim.
intros; inv H; auto.
- apply proc_exit_user_sim.
- apply proc_start_user_sim2.
intros; inv H; auto.
- apply proc_exit_user_sim2.
- layer_sim_simpl.
+ eapply load_correct3.
+ eapply store_correct3.
Qed.
End WITHMEM.
End Refinement.
Context `{real_params: RealParams}.
Context `{multi_oracle_prop: MultiOracleProp}.
Section WITHMEM.
Context `{Hstencil: Stencil}.
Context `{Hmem: Mem.MemoryModelX}.
Context `{Hmwd: UseMemWithData mem}.
Global Instance cpu_id_range: high_level_invariant_impl_CPU_ID_range.
Proof.
econstructor; eauto.
intros.
inv H.
assumption.
Qed.
Global Instance: (LoadStoreProp (hflatmem_store:= flatmem_store) (lflatmem_store:= flatmem_store)).
Proof.
accessor_prop_tac.
- eapply flatmem_store_exists; eauto.
- eapply flatmem_store_match; eauto.
Qed.
Lemma passthrough_correct:
sim (crel HDATA LDATA) vmxintro_passthrough vmcsinit.
Proof.
sim_oplus.
- apply fload_sim.
- apply fstore_sim.
- apply vmxinfo_get_sim.
- apply palloc_sim.
- apply setPT_sim.
- apply ptRead_sim.
- apply ptResv_sim.
- apply shared_mem_status_sim.
- apply offer_shared_mem_sim.
- apply biglow_thread_wakeup_sim.
- apply biglow_thread_yield_sim.
- apply biglow_thread_sleep_sim.
- apply biglow_sched_init_sim.
- apply uctx_get_sim.
- apply uctx_set_sim.
- apply biglow_proc_create_sim.
- apply container_get_nchildren_sim.
- apply container_get_quota_sim.
- apply container_get_usage_sim.
- apply container_can_consume_sim.
- apply get_CPU_ID_sim.
- apply get_curid_sim.
- apply acquire_lock_SC_sim.
- apply release_lock_SC_sim.
- apply get_sync_chan_busy_sim.
- apply set_sync_chan_busy_sim.
- apply ipc_send_body_sim.
- apply ipc_receive_body_sim.
- apply ept_invalidate_mappings_sim.
- apply ept_add_mapping_sim.
- apply rdmsr_sim.
- apply wrmsr_sim.
- apply vmcs_readz_sim.
- apply vmcs_writez_sim.
- apply vmptrld_sim.
- apply vmx_set_intercept_intwin_sim.
- apply vmx_set_desc1_sim.
- apply vmx_set_desc2_sim.
- apply vmx_inject_event_sim.
- apply vmx_set_tsc_offset_sim.
- apply vmx_get_tsc_offset_sim.
- apply vmx_get_exit_reason_sim.
- apply vmx_get_exit_fault_addr_sim.
- apply vmx_check_pending_event_sim.
- apply vmx_check_int_shadow_sim.
- apply vmcs_set_defaults_sim.
- apply cli_sim.
- apply sti_sim.
- apply serial_intr_disable_sim.
- apply serial_intr_enable_sim.
- apply serial_putc_sim.
- apply cons_buf_read_sim.
- apply hostin_sim.
- apply hostout_sim.
- apply proc_create_postinit_sim.
- apply trap_info_get_sim.
- apply trap_info_ret_sim.
- apply proc_start_user_sim.
intros; inv H; auto.
- apply proc_exit_user_sim.
- apply proc_start_user_sim2.
intros; inv H; auto.
- apply proc_exit_user_sim2.
- layer_sim_simpl.
+ eapply load_correct3.
+ eapply store_correct3.
Qed.
End WITHMEM.
End Refinement.