Library mcertikos.mm.ContainerGen


This file provide the contextual refinement proof between MBoot layer and MALInit layer
Require Import Coqlib.
Require Import Errors.
Require Import AST.
Require Import Integers.
Require Import Floats.
Require Import Op.
Require Import Asm.
Require Import Events.
Require Import Globalenvs.
Require Import Smallstep.
Require Import Values.
Require Import Memory.
Require Import Maps.
Require Import CommonTactic.
Require Import AuxLemma.
Require Import FlatMemory.
Require Import AuxStateDataType.
Require Import Constant.
Require Import GlobIdent.
Require Import RealParams.
Require Import LoadStoreSem1.
Require Import AsmImplLemma.
Require Import GenSem.
Require Import RefinementTactic.
Require Import PrimSemantics.

Require Import liblayers.logic.PTreeModules.
Require Import liblayers.logic.LayerLogicImpl.
Require Import liblayers.compcertx.Stencil.
Require Import liblayers.compcertx.MakeProgram.
Require Import liblayers.compat.CompatLayers.
Require Import liblayers.compat.CompatGenSem.
Require Import compcert.cfrontend.Ctypes.
Require Import LayerCalculusLemma.

Require Import AbstractDataType.
Require Import DeviceStateDataType.

Require Import MContainer.
Require Import MContainerIntro.


Require Import ProofIrrelevance.
Require Import ContainerGenSpec.

Require Import CalTicketLock.

Definition of the refinement relation

Section Refinement.

  Local Open Scope string_scope.
  Local Open Scope error_monad_scope.
  Local Open Scope Z_scope.

  Context `{real_params: RealParams}.
  Context `{multi_oracle_prop: MultiOracleProp}.

  Notation HDATA := RData.
  Notation LDATA := RData.

  Notation HDATAOps := (cdata (cdata_ops := mcontainer_data_ops) HDATA).
  Notation LDATAOps := (cdata (cdata_ops := mcontainerintro_data_ops) LDATA).

  Section WITHMEM.

    Context `{Hstencil: Stencil}.
    Context `{Hmem: Mem.MemoryModelX}.
    Context `{Hmwd: UseMemWithData mem}.

Definition the refinement relation: relate_RData + match_RData

    Record relate_RData (f:meminj) (hadt: HDATA) (ladt: LDATA) :=
      mkrelate_RData {
          flatmem_re: FlatMem.flatmem_inj (HP hadt) (HP ladt);
          MM_re: MM hadt = MM ladt;
          MMSize_re: MMSize hadt = MMSize ladt;
          vmxinfo_re: vmxinfo hadt = vmxinfo ladt;
          CR3_re: CR3 hadt = CR3 ladt;
          AC_re: AC hadt = AC ladt;
          ikern_re: ikern hadt = ikern ladt;
          pg_re: pg hadt = pg ladt;
          ihost_re: ihost hadt = ihost ladt;
          ti_fst_re: (fst (ti hadt)) = (fst (ti ladt));
          ti_snd_re: val_inject f (snd (ti hadt)) (snd (ti ladt));
          init_re: init hadt = init ladt;

          buffer_re: buffer hadt = buffer ladt;

          com1_re: com1 hadt = com1 ladt;
          console_re: console hadt = console ladt;
          console_concrete_re: console_concrete hadt = console_concrete ladt;
          ioapic_re: ioapic hadt = ioapic ladt;
          lapic_re: lapic hadt = lapic ladt;
          intr_flag_re: intr_flag hadt = intr_flag ladt;
          curr_intr_num_re: curr_intr_num hadt = curr_intr_num ladt;
          in_intr_re: in_intr hadt = in_intr ladt;
          drv_serial_re: drv_serial hadt = drv_serial ladt;

          CPU_ID_re: CPU_ID hadt = CPU_ID ladt;
          cid_re: cid hadt = cid ladt;
          multi_oracle_re: multi_oracle hadt = multi_oracle ladt;
          multi_log_re: multi_log hadt = multi_log ladt;
          lock_re : lock hadt = lock ladt

        }.

    Inductive match_RData: stencilHDATAmemmeminjProp :=
    | MATCH_RDATA: habd m f s, match_RData s habd m f.

    Local Hint Resolve MATCH_RDATA.

    Global Instance rel_ops: CompatRelOps HDATAOps LDATAOps :=
      {
        relate_AbData s f d1 d2 := relate_RData f d1 d2;
        match_AbData s d1 m f := match_RData s d1 m f;
        new_glbl := nil
      }.

Properties of relations

    Section Rel_Property.

Prove that after taking one step, the refinement relation still holds
      Lemma relate_incr:
         abd abd´ f ,
          relate_RData f abd abd´
          → inject_incr f
          → relate_RData abd abd´.
      Proof.
        inversion 1; subst; intros; inv H; constructor; eauto.
      Qed.

      Global Instance rel_prf: CompatRel HDATAOps LDATAOps.
      Proof.
        constructor; intros; simpl; trivial.
        eapply relate_incr; eauto.
      Qed.

    End Rel_Property.

    Section OneStep_Forward_Relation.

      Section container_init_ref.

        Lemma container_init_spec_ref:
          compatsim (crel HDATA LDATA) (gensem container_init_spec)
                    container_init_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          functional inversion H1.
          inv .
          inv H10; inv H12.
          refine_split´.
          - econstructor; eauto.
            unfold container_init_spec.
            inv match_related.
            rewrite <- pg_re0, <- ikern_re0, <- ihost_re0, <- init_re0, <- in_intr_re0.
            rewrite <- ioapic_re0, <- lapic_re0.
            rewrite H0, H2, H3, H4, H5, H6, H7.
            reflexivity.
            inv match_related.
            constructor.
            rewrite <- ikern_re0; auto.
            rewrite <- ihost_re0; auto.
          - constructor.
          - constructor.
            eauto.
            inv match_related.
            constructor; simpl; auto.
            rewrite multi_log_re0; auto.
            rewrite lock_re0; auto.
            inv match_match.
            constructor.
            eauto.
            eauto.
            eauto.
            eauto.
            eauto.
          - eauto.
        Qed.

      End container_init_ref.

      Section container_split_ref.

        Lemma container_split_spec_ref:
          compatsim (crel HDATA LDATA) (gensem container_split_spec) container_split_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          functional inversion H1; subst.
          inv .
          inv H10; inv H12.
          inv H10; inv H13.
          assert (H9: init d1 = true).
          { inv Hhigh.
            inv match_related.
            rewrite init_re0.
            unfold c in H6.
            rewrite AC_re0 in H6.
            case_eq (init d2); intros; subst; auto.
            eapply init_false_empty_container in H.
            rewrite H in H6.
            rewrite ZMap.gi in H6.
            simpl in H6.
            inv H6. }
          refine_split´.
          + econstructor; eauto.
            unfold c in H6.
            inv Hhigh´.
            inv valid_container.
            generalize (cvalid_id _ H6).
            generalize (cvalid_quota _ H6).
            generalize (cvalid_usage _ H6); intros.
            unfold container_split_low_spec.
            inv match_related.
            rewrite <- ikern_re0, <- ihost_re0, <- init_re0.
            rewrite <- AC_re0.
            unfold c in ×.
            rewrite H6, H7, H8, H9; auto.
            rewrite zle_lt_true; eauto.
            rewrite zle_true; eauto.
            rewrite zle_le_true; eauto.
            simpl.
            rewrite H5.
            rewrite zle_lt_true; eauto.
            rewrite zlt_true; eauto.
            rewrite zle_le_true; eauto.
            inv match_related.
            constructor.
            rewrite <- ikern_re0; auto.
            rewrite <- ihost_re0; auto.
          + unfold c in *; unfold cur in ×.
            simpl in ×.
            rewrite H0.
            rewrite Int.repr_unsigned; constructor.
          + constructor; eauto.
            inv match_related.
            constructor; simpl; auto.
            unfold cur; unfold child.
            rewrite H6; reflexivity.
            constructor.
          + eauto.
        Qed.

      End container_split_ref.

      Section container_can_consume_sim.

        Lemma container_can_consume_spec_sim:
           id,
            sim (crel HDATA LDATA) (id gensem container_can_consume_spec)
                      (id gensem container_can_consume_intro_spec).
        Proof.
          intros. layer_sim_simpl. compatsim_simpl (@match_AbData). intros.
          functional inversion H2.
          - inv .
            inv H7; inv H9.
            inv H7; inv H10.
            refine_split´.
            + repeat (econstructor; eauto).
              unfold container_can_consume_intro_spec.
              inv match_related.
              rewrite <- ikern_re0, <- ihost_re0.
              rewrite <- AC_re0.
              rewrite H0, H1, H3, H4; auto.
              inv Hhigh´.
              inv valid_container.
              generalize (cvalid_id _ H0).
              generalize (cvalid_quota _ H0).
              generalize (cvalid_usage _ H0).
              intros.
              rewrite zle_lt_true; try omega.
              rewrite zle_true; try omega.
              rewrite zle_le_true; try omega.
              rewrite H.
              reflexivity.
            + constructor.
            + constructor; eauto.
            + eauto.
          - inv .
            inv H7; inv H9.
            inv H7; inv H10.
            refine_split´.
            + repeat (econstructor; eauto).
              unfold container_can_consume_intro_spec.
              inv match_related.
              rewrite <- ikern_re0, <- ihost_re0.
              rewrite <- AC_re0.
              rewrite H0, H1, H3, H4; auto.
              inv Hhigh´.
              inv valid_container.
              generalize (cvalid_id _ H0).
              generalize (cvalid_quota _ H0).
              generalize (cvalid_usage _ H0).
              intros.
              rewrite zle_lt_true; try omega.
              rewrite zle_true; try omega.
              rewrite zle_le_true; try omega.
              rewrite H.
              reflexivity.
            + constructor.
            + constructor; eauto.
            + eauto.
        Qed.

      End container_can_consume_sim.

      Section container_alloc_sim.

        Lemma container_alloc_spec_sim:
           id,
            sim (crel HDATA LDATA) (id gensem container_alloc_spec)
                      (id gensem container_alloc_intro_spec).
        Proof.
          intros. layer_sim_simpl. compatsim_simpl (@match_AbData). intros.
          functional inversion H1.
          - inv .
            inv H9; inv H11.
            refine_split´.
            + repeat (econstructor; eauto).
              unfold container_alloc_intro_spec.
              inv match_related.
              rewrite <- ikern_re0, <- ihost_re0, <- init_re0.
              rewrite <- AC_re0.
              unfold c in *; unfold cur in *; clear c cur.
              rewrite H3, H4, H5; auto.
              inv Hhigh´.
              inv valid_container.
              generalize (cvalid_id _ H2).
              generalize (cvalid_quota _ H2).
              generalize (cvalid_usage _ H2).
              intros.
              rewrite zle_lt_true; try omega.
              rewrite zle_true; try omega.
              rewrite zle_le_true; try omega.
              rewrite H2, H6.
              rewrite H0.
              reflexivity.
            + constructor.
            + econstructor; eauto.
              inv match_related.
              econstructor; simpl; eauto.
              unfold cur.
              unfold c.
              rewrite H0.
              rewrite <- H2.
              unfold c.
              reflexivity.
              econstructor.
            + auto.
          - inv .
            inv H9; inv H11.
            refine_split´.
            + repeat (econstructor; eauto).
              unfold container_alloc_intro_spec.
              inv match_related.
              rewrite <- ikern_re0, <- ihost_re0, <- init_re0.
              rewrite <- AC_re0.
              unfold c in *; unfold cur in *; clear c cur.
              rewrite H3, H4, H5; auto.
              inv Hhigh´.
              inv valid_container.
              generalize (cvalid_id _ H2).
              generalize (cvalid_quota _ H2).
              generalize (cvalid_usage _ H2).
              intros.
              rewrite zle_lt_true; try omega.
              rewrite zle_true; try omega.
              rewrite zle_le_true; try omega.
              rewrite H2, H6.
              rewrite H0.
              reflexivity.
            + constructor.
            + econstructor; eauto.
            + auto.
        Qed.

      End container_alloc_sim.

      Section container_get_quota_sim.

        Lemma container_get_quota_spec_sim:
           id,
            sim (crel HDATA LDATA) (id gensem container_get_quota_spec)
                      (id gensem container_get_quota_intro_spec).
        Proof.
          intros. layer_sim_simpl. compatsim_simpl (@match_AbData). intros.
          functional inversion H2; subst.
          inv .
          inv H6; inv H8.
          refine_split´.
          + repeat (econstructor; eauto).
            unfold container_get_quota_intro_spec.
            inv match_related.
            inv Hhigh´.
            inv valid_container.
            exploit cvalid_id; eauto; intros.
            rewrite zle_lt_true; eauto.
            rewrite <- ikern_re0, <- ihost_re0.
            rewrite <- AC_re0.
            fold c.
            rewrite H0, H1, H3; eauto.
            rewrite H; eauto.
            reflexivity.
          + constructor.
          + constructor; eauto.
          + auto.
        Qed.

      End container_get_quota_sim.

      Section container_get_usage_sim.

        Lemma container_get_usage_spec_sim:
           id,
            sim (crel HDATA LDATA) (id gensem container_get_usage_spec)
                (id gensem container_get_usage_intro_spec).
        Proof.
          intros. layer_sim_simpl. compatsim_simpl (@match_AbData). intros.
          functional inversion H2; subst.
          inv .
          inv H6; inv H8.
          refine_split´.
          + repeat (econstructor; eauto).
            unfold container_get_usage_intro_spec.
            inv match_related.
            inv Hhigh´.
            inv valid_container.
            exploit cvalid_id; eauto; intros.
            rewrite zle_lt_true; eauto.
            rewrite <- ikern_re0, <- ihost_re0.
            rewrite <- AC_re0.
            fold c.
            rewrite H0, H1, H3; eauto.
            rewrite H; eauto.
            reflexivity.
          + constructor.
          + constructor; eauto.
          + auto.
        Qed.

      End container_get_usage_sim.

      Section container_get_parent_sim.

        Lemma container_get_parent_spec_sim:
           id,
            sim (crel HDATA LDATA) (id gensem container_get_parent_spec)
                (id gensem container_get_parent_intro_spec).
        Proof.
          intros. layer_sim_simpl. compatsim_simpl (@match_AbData). intros.
          functional inversion H2; subst.
          inv .
          inv H6; inv H8.
          refine_split´.
          + repeat (econstructor; eauto).
            unfold container_get_parent_intro_spec.
            inv match_related.
            inv Hhigh´.
            inv valid_container.
            exploit cvalid_id; eauto; intros.
            rewrite zle_lt_true; eauto.
            rewrite <- ikern_re0, <- ihost_re0.
            rewrite <- AC_re0.
            fold c.
            rewrite H0, H1, H3; eauto.
            rewrite H; eauto.
            reflexivity.
          + constructor.
          + constructor; eauto.
          + auto.
        Qed.

      End container_get_parent_sim.

      Section container_get_nchildren_sim.

        Lemma container_get_nchildren_spec_sim:
           id,
            sim (crel HDATA LDATA) (id gensem container_get_nchildren_spec)
                (id gensem container_get_nchildren_intro_spec).
        Proof.
          intros. layer_sim_simpl. compatsim_simpl (@match_AbData). intros.
          functional inversion H2; subst.
          inv .
          inv H6; inv H8.
          refine_split´.
          + repeat (econstructor; eauto).
            unfold container_get_nchildren_intro_spec.
            inv match_related.
            inv Hhigh´.
            inv valid_container.
            exploit cvalid_id; eauto; intros.
            rewrite zle_lt_true; eauto.
            rewrite <- ikern_re0, <- ihost_re0.
            rewrite <- AC_re0.
            fold c.
            rewrite H0, H1, H3; eauto.
            rewrite H; eauto.
            reflexivity.
          + constructor.
          + constructor; eauto.
          + auto.
        Qed.

      End container_get_nchildren_sim.

      Ltac pattern2_refinement_simpl:=
        pattern2_refinement_simpl´ (@relate_AbData).

      Global Instance: (LoadStoreProp (hflatmem_store:= flatmem_store´) (lflatmem_store:= flatmem_store´)).
      Proof.
        accessor_prop_tac.
        - eapply flatmem_store´_exists; eauto.
      Qed.

      Lemma passthrough_correct:
        sim (crel HDATA LDATA) mcontainer_passthrough mcontainerintro.
      Proof.
        sim_oplus.
        - apply fload´_sim.
        - apply fstore´_sim.
        - apply page_copy´_sim.
        - apply page_copy_back´_sim.

        - apply vmxinfo_get_sim.
        - apply setPG_sim.
        - apply setCR3_sim.
        - apply get_size_sim.
        - apply is_mm_usable_sim.
        - apply get_mm_s_sim.
        - apply get_mm_l_sim.
        - apply container_can_consume_spec_sim.
        - apply container_alloc_spec_sim.
        - apply container_get_parent_spec_sim.
        - apply container_get_nchildren_spec_sim.
        - apply container_get_quota_spec_sim.
        - apply container_get_usage_spec_sim.
        - apply get_CPU_ID_sim.
        - apply get_curid_sim.
        - apply set_curid_sim.
        - apply set_curid_init_sim.
        - apply (release_lock_sim (valid_arg_imply:= Shared2ID0_imply)).
        -
          eapply acquire_lock_sim0; eauto.
          intros. inv H; trivial; try inv H0.
        - 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 trapin_sim.
        - apply trapout´_sim.
        - apply hostin_sim.
        - apply hostout´_sim.
        - apply proc_create_postinit_sim.
        - apply trap_info_get_sim.
        - apply trap_info_ret_sim.
        - layer_sim_simpl.
          + eapply load_correct1.
          + eapply store_correct1.
      Qed.

    End OneStep_Forward_Relation.

  End WITHMEM.

End Refinement.