Library mcertikos.multithread.lowrefins.E2PBThreadGen


Require Import Coqlib.
Require Import Maps.
Require Import ASTExtra.
Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Memory.
Require Import Events.
Require Import Stacklayout.
Require Import Globalenvs.
Require Import AsmX.
Require Import Smallstep.
Require Import AuxStateDataType.
Require Import Constant.
Require Import GlobIdent.
Require Import FlatMemory.
Require Import CommonTactic.
Require Import AuxLemma.
Require Import RealParams.
Require Import PrimSemantics.
Require Import PrimTMSemantics.
Require Import LAsm.
Require Import LoadStoreSem3.
Require Import XOmega.

Require Import liblayers.logic.PTreeModules.
Require Import liblayers.logic.LayerLogicImpl.
Require Import liblayers.compat.CompatLayers.
Require Import liblayers.compat.CompatGenSem.
Require Import I64Layer.

Require Import AlgebraicMem.

Require Import Decision.
Require Import FutureTactic.

Require Import GlobalOracleProp.
Require Import AbstractDataType.
Require Import DeviceStateDataType.
Require Import SingleConfiguration.
Require Import ObjPHBFlatMem.
Require Import PHBThread.
Require Import AsmE2L.

Require Import AuxSingleAbstractDataType.
Require Import SingleOracleImpl.
Require Import E2PBThreadComposeData.
Require Import LoadStoreSemPHB.
Require Import ObjPHBThreadDEV.
Require Import ObjPHBThreadINTELVIRT.
Require Import ObjPHBThreadIPC.
Require Import ObjPHBThreadReplayFunction.
Require Import ObjPHBThreadSCHED.
Require Import ObjPHBThreadVMM.
Require Import ObjInterruptManagement.
Require Import ObjConsole.
Require Import ObjInterruptController.
Require Import ObjSerialDevice.
Require Import OracleInstances.
Require Import ObjVMMFun.
Require Import ObjVMMGetSet.
Require Import ObjVMMDef.
Require Import ObjCV.
Require Import ObjCPU.
Require Import ObjVMX.
Require Import ObjVMCS.
Require Import ObjProc.
Require Import ObjBigThread.
Require Import ObjContainer.
Require Import ObjMultiprocessor.
Require Import ObjThread.
Require Import BigLogThreadConfigFunction.
Require Import BigLogThreadConfigFunctionProp.

Require Import StencilImpl.

Require Import E2PBThreadGenDef.

Require Import E2PBThreadGenLemma.

Require Import E2PBThreadGenLemmaPTRESV.
Require Import E2PBThreadGenLemmaSCHEDINIT.
Require Import E2PBThreadGenLemmaHASEVENT.
Require Import E2PBThreadGenLemmaINITMATCH.

Section PHBTHREADABSREL.

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

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

  Context `{s_oracle_ops : SingleOracleOps}.
  Context `{s_threads_ops: ThreadsConfigurationOps}.

  Local Instance s_oracle_prf : SingleOracle := SingleOracleImpl.s_oracle_prf.

  Context `{s_threads_prf: !ThreadsConfiguration (s_oracle := s_oracle_prf)}.
  Context `{s_oracle_basic_prop_prf: !SingleOracleBasicProp (single_oracle_prf := s_oracle_prf)
                                      (s_threads_ops := s_threads_ops)}.

  Local Instance s_oracle_prop_prf : SingleOracleProp := SingleOracleImpl.s_oracle_prop_prf.

  Context `{Hstencil : Stencil (stencil := stencil)}.

  Local Instance s_config : SingleConfiguration mem := s_config.

  Context `{s_oracle_basic_link_prop_prf: !SingleOracleBasicLinkProp (single_oracle_prf := s_oracle_prf)
                                           (s_threads_ops := s_threads_ops)}.

  Local Instance s_oracle_link_prop_prf : SingleOracleLinkProp := SingleOracleImpl.s_oracle_link_prop_prf.

  Local Instance s_link_config : SingleLinkConfiguration mem := s_link_config.

  Context `{acc_def: !AccessorsDefined (mem := mem) (stencil := stencil) (stencil_ops := stencil_ops)
                      (memory_model_ops := memory_model_ops)
                      (phbthread (s_config := s_config) L64)}.

  Section ABS_REL_INSTANCE.

    Context `{e2pbthread_gen_prop: !E2PBThreadGenProp s_oracle_prf}.

    Program Instance abs_rel : AbstractRel (spl_data := single_data) (mem := mem) (spl_mem_ops := memory_model_ops)
                                           (s_link_config := s_link_config) (stencil_ops := @stencil_ops)
                                           (real_params_ops := real_params_ops) (LH:= phbthread) :=
      {
        match_EData_RData ed l rd := kctxt_pool, relate_RData kctxt_pool (update init_shared_adt l) ed rd
      }.

    Opaque full_thread_list.

    Next Obligation.
      apply abs_rel_init_match_aux.
    Qed.


    Next Obligation.
      rename Hdata into kctxt_pool.
       kctxt_pool.
      inv H; constructor; try (simpl in *; intros; auto); try (rewrite zmap_set_eq; eauto).
      generalize (per_data_re i); auto.
      constructor.
      intros.
      inv sh_mem_inv_re.
      case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l))
                               (cid (update init_shared_adt l)))); intros; subst.
      - rewrite ZMap.gss in H0.
        rewrite ZMap.gso in H1; eauto.
      - rewrite ZMap.gso in H0; auto.
        case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                 (cid (update init_shared_adt l)))); intros; subst.
        + rewrite ZMap.gss in H1; eauto.
        + rewrite ZMap.gso in H1; eauto.
    Qed.

    Next Obligation.
      rename H1 into kctxt_pool.
      inversion acc_def; subst.
      unfold acc_exec_load in H.
      unfold acc_exec_load_strong in H; simpl.
      destruct acc_def.
      simpl in H.
      unfold PHBThread.single_exec_loadex in H.
      unfold single_exec_loadex in H.
      assert (Hflag: AbstractDataType.init a = init (uRData l)
                     AbstractDataType.ihost a = ihost d
                     AbstractDataType.pg a = pg (uRData l)
                     AbstractDataType.ikern a = ikern d).
      { generalize (per_data_re kctxt_pool (uRData l)
                                (ZMap.set (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                          (Some d) dp) a H3 (proc_id (uRData l))).
        simpl; rewrite ZMap.gss; intros.
        inv H1; simpl; try (rewrite zeq_true in *); eauto. }
      destruct Hflag as (Hflag_init & Hflag_ihost & Hflag_pg & Hflag_ikern); simpl in ×.
      unfold exec_loadex; simpl; unfold exec_loadex3.
      subdestruct.

      -
        simpl in *; a; rewrite Hflag_ihost, Hflag_ikern.
        unfold Asm.exec_load in ×.
        unfold Mem.loadv in *; rewrite Hdestruct3 in H; rewrite Hdestruct3.
        subdestruct; unfold Mem.load in Hdestruct4.
        simpl in Hdestruct4; lift_unfold.
        rewrite Hdestruct4; split; try (inv H; simpl; auto).
        split; auto; split; auto.
         kctxt_pool; auto.

      -
        rewrite Hflag_ihost, Hflag_ikern, Hflag_pg.
        unfold HostAccess2.exec_host_load2; unfold HostAccessPHB.single_exec_host_load in H.
        Opaque PDX Zdivide_dec zle Z.mul Z.add Z.sub Z.div.
        simpl.
        simpl in H.
        assert (Htemp: AbstractDataType.PT a = PT d
                       ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a) =
                       ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))).
        { inv H3.
          generalize (per_data_re (proc_id (uRData l))).
          intros.
          simpl in H1; rewrite ZMap.gss in H1.
          inv H1; simpl in ×.
          rewrite Hdestruct in PT_re.
          rewrite Hdestruct in sh_cid_re.
          rewrite <- sh_cid_re.
          rewrite zeq_true in PT_re; auto. }
        destruct Htemp as (Hflag_PT & Hflag_cid).
        simpl; subdestruct; subst.

        + rewrite Hflag_init.
          rename e into PT_cond.
          rewrite PT_cond.
          assert (Hptp_re: ZMap.get (AbstractDataType.PT a) (AbstractDataType.ptpool a) = ZMap.get (PT d) (ptpool d)).
          { inv H3.
            generalize (per_data_re (proc_id (uRData l))).
            intros.
            simpl in H1; rewrite ZMap.gss in H1.
            inv H1.
            rewrite <- Hflag_cid in ptp_re.
            rewrite Hflag_PT.
            rewrite <- PT_cond.
            rewrite <- Hflag_cid; eauto.
            rewrite Hdestruct in ptp_re; auto. }
          rewrite <- Hflag_PT.
          rewrite Hptp_re.
          rewrite Hdestruct7, Hdestruct8.
          unfold FlatLoadStoreSem.exec_flatmem_load.
          unfold FlatLoadStoreSemPHB.single_exec_flatmem_load in H; simpl; inv H.
           a.
          split; eauto.
          remember (FlatMem.load TY (AbstractDataType.HP a) (PTADDR v (Int.unsigned i))) as v1.
          remember (FlatMem.load TY (HP ) (PTADDR v (Int.unsigned i))) as v2.
          assert (loadv_eq: v1 = v2).
          { inv H3.
            generalize (per_data_re (ZMap.get (CPU_ID (update init_shared_adt l))
                                              (cid (update init_shared_adt l)))); intros.
            rewrite ZMap.gss in H.
            inv H.
            unfold PTADDR.
            symmetry.
            eapply flatmem_re; eauto. }
          unfold nextinstr_nf; simpl.
          unfold nextinstr; simpl; repeat f_equal; auto.

          try (split; auto).
          split; auto; kctxt_pool.
          rewrite Hflag_PT.
          rewrite <- PT_cond; auto.

        + rewrite Hflag_init.
          rename e into PT_cond.
          rewrite PT_cond.
          assert (Hptp_re: ZMap.get (AbstractDataType.PT a) (AbstractDataType.ptpool a) = ZMap.get (PT d) (ptpool d)).
          { inv H3.
            generalize (per_data_re (proc_id (uRData l))).
            intros.
            simpl in H1; rewrite ZMap.gss in H1.
            inv H1.
            rewrite <- Hflag_cid in ptp_re.
            rewrite Hflag_PT.
            rewrite <- PT_cond.
            rewrite <- Hflag_cid; eauto.
            rewrite Hdestruct in ptp_re; auto. }
          rewrite <- Hflag_PT.
          rewrite Hptp_re.
          rewrite Hdestruct7, Hdestruct8.
          unfold FlatLoadStoreSem.exec_flatmem_load.
          unfold FlatLoadStoreSemPHB.single_exec_flatmem_load in H; simpl; inv H.
           a.
          split; eauto.
          remember (FlatMem.load TY (AbstractDataType.HP a) (PTADDR v (Int.unsigned i))) as v1.
          remember (FlatMem.load TY (HP ) (PTADDR v (Int.unsigned i))) as v2.
          assert (loadv_eq: v1 = v2).
          { inv H3.
            generalize (per_data_re (ZMap.get (CPU_ID (update init_shared_adt l))
                                              (cid (update init_shared_adt l)))); intros.
            rewrite ZMap.gss in H.
            inv H.
            unfold PTADDR.
            symmetry.
            eapply flatmem_re; eauto. }
          unfold nextinstr_nf; simpl.
          unfold nextinstr; simpl; repeat f_equal; auto.

          try (split; auto).
          split; auto; kctxt_pool.
          rewrite Hflag_PT.
          rewrite <- PT_cond; auto.

        +
          unfold PageFault.exec_pagefault.
          unfold PageFaultPHB.single_exec_pagefault in H.
          simpl; subdestruct.
          rewrite Hflag_init.
          rewrite Hflag_PT.
          rewrite <- e.
          assert (Hptp_re: ZMap.get (AbstractDataType.PT a) (AbstractDataType.ptpool a) = ZMap.get (PT d) (ptpool d)).
          { inv H3.
            generalize (per_data_re (proc_id (uRData l))).
            intros.
            simpl in H1; rewrite ZMap.gss in H1.
            inv H1.
            rewrite <- Hflag_cid in ptp_re.
            rewrite Hflag_PT.
            rewrite <- e.
            rewrite <- Hflag_cid; eauto.
            rewrite Hdestruct in ptp_re; auto. }
          rewrite e.
          rewrite <- Hflag_PT.
          rewrite Hptp_re.
          rewrite Hdestruct7, Hdestruct8.
          inv H; unfold trapinfo_set; unfold single_trapinfo_set; simpl.
           (a {ti : (i, rs RA)}); split; auto; split; auto; split; auto.
           kctxt_pool.
          inv H3.
          rewrite Hflag_PT; rewrite <- e.
          constructor; simpl; auto.
          intros.
          generalize (per_data_re i0); intros.
          case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros; subst.
          { rewrite ZMap.gss in H; rewrite ZMap.gss.
            inv H; constructor; simpl; auto.
            - rewrite zeq_true; eauto.
            - constructor; simpl; intros.
              inv inv_re.
              apply dirty_page_per_thread_inv; auto.
          }
          { rewrite ZMap.gso in H; auto; rewrite ZMap.gso; auto.
            destruct (ZMap.get i0 dp); auto.
            - inv H; simpl; constructor; simpl; auto.
              + rewrite zeq_false; auto.
                rewrite zeq_false in ti_re; auto.
              + constructor.
                inv inv_re.
                intros.
                apply dirty_page_per_thread_inv; auto.
            - constructor. }
          { inv sh_shared_inv_re.
            constructor; simpl; auto. }
          { inv sh_mem_inv_re.
            constructor; simpl; intros.
            case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l))
                                      (cid (update init_shared_adt l)))); intros; subst.
            - rewrite ZMap.gss in H1.
              rewrite ZMap.gso in H3; eauto.
              inv H1; simpl in H4.
              eauto.
              eapply pperm_disjoint; eauto.
              rewrite ZMap.gss; auto.
              rewrite ZMap.gso; auto.
            - rewrite ZMap.gso in H1; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                 (cid (update init_shared_adt l)))); intros; subst.
              + rewrite ZMap.gss in H3; eauto.
                inv H3.
                simpl.
                eapply pperm_disjoint; eauto.
                rewrite ZMap.gso; eauto.
                rewrite ZMap.gss; eauto.
              + rewrite ZMap.gso in H3; eauto.
                eapply pperm_disjoint in H; eauto.
                rewrite ZMap.gso; eauto.
                rewrite ZMap.gso; eauto. }

      -
        simpl in *; a; rewrite Hflag_ihost, Hflag_ikern.
        unfold Asm.exec_load in ×.
        unfold Mem.loadv in *; rewrite Hdestruct3 in H; rewrite Hdestruct3.
        subdestruct; unfold Mem.load in Hdestruct4.
        simpl in Hdestruct4; lift_unfold.
        rewrite Hdestruct4; split; try (inv H; simpl; auto).
        split; auto; split; auto.
         kctxt_pool; auto.

      -
        rewrite Hflag_ihost.
        rewrite Hflag_init.
        assert (Htemp: ds´ = update init_shared_adt l d = m = ).
        { unfold GuestAccessIntelPHB.single_exec_guest_intel_load in H.
          unfold GuestAccessIntelDefPHB.single_exec_guest_intel_accessor,
          GuestAccessIntelPHB.single_load_accessor in H.
          subdestruct.
          subst.
          unfold FlatLoadStoreSemPHB.single_exec_flatmem_load in H; inversion H.
          subst; auto. }
        destruct Htemp as (Htemp_a & Htemp_b & Htemp_c); subst; a.
        split; simpl; eauto.
        unfold GuestAccessIntelPHB.single_exec_guest_intel_load in H.
        unfold GuestAccessIntelDefPHB.single_exec_guest_intel_accessor,
        GuestAccessIntelPHB.single_load_accessor in H.
        unfold GuestAccessIntel2.exec_guest_intel_load2.
        unfold GuestAccessIntelDef2.exec_guest_intel_accessor2.
        unfold GuestAccessIntel2.load_accessor2.
        subdestruct.
        assert (Hcid_eq: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))
                         = ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)).
        { inv H3.
          rewrite Hdestruct in sh_cid_re.
          auto. }
        rewrite Hcid_eq in ×.
        assert (Hept_eq: ept = ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.ept a)).
        { inv H3.
          generalize (per_data_re (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))).
          intros.
          rewrite ZMap.gss in H1.
          inv H1.
          simpl in e.
          rewrite Hdestruct in cid_re.
          rewrite <- cid_re in ept_re.
          rewrite e in ept_re.
          rewrite zeq_true in ept_re.
          auto. }
        rewrite <- Hept_eq.
        simpl in ×.
        rewrite Hdestruct6, Hdestruct7, Hdestruct8, Hdestruct9.
        unfold FlatLoadStoreSemPHB.single_exec_flatmem_load in H.
        inv H.
        unfold FlatLoadStoreSem.exec_flatmem_load.
        simpl.
        remember (FlatMem.load TY (AbstractDataType.HP a) (EPTADDR (hpa / 4096) (Int.unsigned i))) as v1.
        remember (FlatMem.load TY (HP ) (EPTADDR (hpa / 4096) (Int.unsigned i))) as v2.
        assert (loadv_eq: v1 = v2).
        { inv H3.
          generalize (per_data_re (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros.
          rewrite ZMap.gss in H.
          inv H.
          unfold EPTADDR.
          symmetry.
          eapply flatmem_re; eauto. }
        rewrite loadv_eq.
        reflexivity.

      -
        rewrite Hflag_ihost.
        rewrite Hflag_init.
        assert (Htemp: ds´ = update init_shared_adt l d = m = ).
        { unfold GuestAccessIntelPHB.single_exec_guest_intel_load in H.
          unfold GuestAccessIntelDefPHB.single_exec_guest_intel_accessor,
          GuestAccessIntelPHB.single_load_accessor in H.
          subdestruct.
          subst.
          unfold FlatLoadStoreSemPHB.single_exec_flatmem_load in H; inversion H.
          subst; auto. }
        destruct Htemp as (Htemp_a & Htemp_b & Htemp_c); subst; a.
        split; simpl; eauto.
        unfold GuestAccessIntelPHB.single_exec_guest_intel_load in H.
        unfold GuestAccessIntelDefPHB.single_exec_guest_intel_accessor,
        GuestAccessIntelPHB.single_load_accessor in H.
        unfold GuestAccessIntel2.exec_guest_intel_load2.
        unfold GuestAccessIntelDef2.exec_guest_intel_accessor2.
        unfold GuestAccessIntel2.load_accessor2.
        subdestruct.
        assert (Hcid_eq: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))
                         = ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)).
        { inv H3.
          rewrite Hdestruct in sh_cid_re.
          auto. }
        rewrite Hcid_eq in ×.
        assert (Hept_eq: ept = ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.ept a)).
        { inv H3.
          generalize (per_data_re (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))).
          intros.
          rewrite ZMap.gss in H1.
          inv H1.
          simpl in e.
          rewrite Hdestruct in cid_re.
          rewrite <- cid_re in ept_re.
          rewrite e in ept_re.
          rewrite zeq_true in ept_re.
          auto. }
        rewrite <- Hept_eq.
        simpl in ×.
        rewrite Hdestruct6, Hdestruct7, Hdestruct8, Hdestruct9.
        unfold FlatLoadStoreSemPHB.single_exec_flatmem_load in H.
        inv H.
        unfold FlatLoadStoreSem.exec_flatmem_load.
        simpl.
        remember (FlatMem.load TY (AbstractDataType.HP a) (EPTADDR (hpa / 4096) (Int.unsigned i))) as v1.
        remember (FlatMem.load TY (HP ) (EPTADDR (hpa / 4096) (Int.unsigned i))) as v2.
        assert (loadv_eq: v1 = v2).
        { inv H3.
          generalize (per_data_re (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros.
          rewrite ZMap.gss in H.
          inv H.
          unfold EPTADDR.
          symmetry.
          eapply flatmem_re; eauto. }
        rewrite loadv_eq.
        reflexivity.

      -
        simpl in *; a; rewrite Hflag_ihost, Hflag_ikern.
        unfold Asm.exec_load in ×.
        unfold Mem.loadv in *; rewrite Hdestruct2 in H; rewrite Hdestruct2.
        subdestruct; unfold Mem.load in Hdestruct3.
        simpl in Hdestruct3; lift_unfold.
        rewrite Hdestruct3; split; try (inv H; simpl; auto).
        split; auto; split; auto.
         kctxt_pool; auto.
    Fail idtac.
    Qed.
    Next Obligation.
      rename H1 into kctxt_pool.
      unfold single_exec_storeex in H.
      unfold acc_exec_store_strong in H; simpl.
      destruct acc_def.
      simpl in H.
      unfold PHBThread.single_exec_storeex in H.
      unfold single_exec_storeex in H; simpl in ×.
      assert (Hflag: AbstractDataType.init a = init (uRData l)
                     AbstractDataType.ihost a = ihost d
                     AbstractDataType.pg a = pg (uRData l)
                     AbstractDataType.ikern a = ikern d).
      { generalize (per_data_re kctxt_pool (uRData l)
                                (ZMap.set (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                          (Some d) dp) a H3 (proc_id (uRData l))).
        simpl; rewrite ZMap.gss; intros.
        inv H1; simpl; try (rewrite zeq_true in *); eauto. }
      destruct Hflag as (Hflag_init & Hflag_ihost & Hflag_pg & Hflag_ikern); simpl in ×.
      unfold exec_storeex; simpl; unfold exec_storeex3.
      subdestruct.

      -
        simpl in *; a; rewrite Hflag_ihost, Hflag_ikern.
        unfold Asm.exec_store in ×.
        unfold Mem.storev in ×. rewrite Hdestruct3 in H; rewrite Hdestruct3.
        subdestruct; unfold Mem.store in Hdestruct4.
        simpl in Hdestruct4; lift_unfold.
        destruct Hdestruct4; destruct m0; simpl in H4.
        rewrite H1; simpl; inv H; simpl; split; auto.
        simpl in H8; inv H8; split; auto; split; auto.
         kctxt_pool; auto.

      -
        rewrite Hflag_ihost, Hflag_ikern, Hflag_pg.
        unfold HostAccess2.exec_host_store2; unfold HostAccessPHB.single_exec_host_store in H.
        Opaque PDX Zdivide_dec zle Z.mul Z.add Z.sub Z.div.
        simpl.
        simpl in H.
        assert (Htemp: AbstractDataType.PT a = PT d
                       ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a) =
                       ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))).
        { inv H3.
          generalize (per_data_re (proc_id (uRData l))).
          intros.
          simpl in H1; rewrite ZMap.gss in H1.
          inv H1; simpl in ×.
          rewrite Hdestruct in PT_re.
          rewrite Hdestruct in sh_cid_re.
          rewrite <- sh_cid_re.
          rewrite zeq_true in PT_re; auto. }
        destruct Htemp as (Hflag_PT & Hflag_cid).
        simpl; subdestruct; subst.

        + rewrite Hflag_init.
          rename e into PT_cond.
          rewrite PT_cond.
          assert (Hptp_re: ZMap.get (AbstractDataType.PT a) (AbstractDataType.ptpool a) = ZMap.get (PT d) (ptpool d)).
          { inv H3.
            generalize (per_data_re (proc_id (uRData l))).
            intros.
            simpl in H1; rewrite ZMap.gss in H1.
            inv H1.
            rewrite <- Hflag_cid in ptp_re.
            rewrite Hflag_PT.
            rewrite <- PT_cond.
            rewrite <- Hflag_cid; eauto.
            rewrite Hdestruct in ptp_re; auto. }
          rewrite <- Hflag_PT.
          rewrite Hptp_re.
          rewrite Hdestruct7, Hdestruct8.
          unfold FlatLoadStoreSem.exec_flatmem_store.
          unfold FlatLoadStoreSemPHB.single_exec_flatmem_store in H; simpl; inv H.
          subdestruct; eauto.
          unfold single_flatmem_store in Hdestruct9; subdestruct.
          unfold flatmem_store.
          simpl in Hdestruct9; inv Hdestruct9.
          assert (Hpperm_re: pperm_inject (pperm d) (AbstractDataType.pperm a)).
          { inv H3; eauto.
            generalize (per_data_re (proc_id (uRData l))); intros.
            simpl; rewrite ZMap.gss in H.
            inv H; auto. }
          unfold pperm_inject in Hpperm_re.
          generalize (Hpperm_re (PageI (PTADDR v (Int.unsigned i)))); intros.
          simpl in Hdestruct10.
          rewrite Hdestruct10 in H.
          clear Hpperm_re.
          subdestruct; try contradiction.
          clear Hdestruct4.
          inv H4.
          rewrite <- Hflag_cid in PT_cond.
           (a {HP: FlatMem.store TY (AbstractDataType.HP a) (PTADDR v (Int.unsigned i)) (rs rd)}).
          split; eauto.
          split; eauto.
          split; eauto.
           kctxt_pool.
          generalize H3; intros Hrelate.
          simpl.
          inv Hrelate; simpl in ×.
          constructor; simpl; auto.
          intros.
          generalize (per_data_re i0).
          intros.
          rewrite Hflag_PT; rewrite <- PT_cond.
          rewrite Hflag_cid.
          case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros; subst.
          { rewrite ZMap.gss in H1.
            rewrite ZMap.gss.
            inv H1.
            constructor; simpl; auto.
            { eapply abs_rel_flatmem_store_mine_related_prop; eauto. }
            { inv inv_re.
              constructor; simpl.
              assert ( i0 : ZIndexed.t, ZMap.get i0 (pperm d) = PGUndef
                                                adr: ZIndexed.t,
                                                 ZMap.get adr (HP d) =
                                                 ZMap.get adr (FlatMem.free_page i0 (HP d))).
              { intros; eapply dirty_page_per_thread_inv; eauto. }
              assert (FlatMem.store TY (HP d) (PTADDR v (Int.unsigned i)) (rs rd) =
                      FlatMem.store TY (HP d) (PTADDR v (Int.unsigned i)) (rs rd)) by reflexivity.
              assert (PTADDR v (Int.unsigned i) mod 4096 4096 - size_chunk TY).
              { eapply PTADDR_mod_lt; auto. }
              generalize (single_dirty_ppage´_store_unmapped´
                            (pperm d) (HP d) H1 _ Hdestruct10 (rs rd)
                            (FlatMem.store TY (HP d) (PTADDR v (Int.unsigned i)) (rs rd)) TY H5 H6).
              intros.
              eapply H7; auto. } }
          { rewrite ZMap.gso in H1; auto.
            rewrite ZMap.gso; auto.
            subdestruct; auto.
            inv H1.
            per_data_auto.
            - assert (ZMap.get (PageI (PTADDR v (Int.unsigned i))) (pperm p) = PGUndef).
              { inv sh_mem_inv_re.
                generalize (pperm_disjoint (ZMap.get (CPU_ID (update init_shared_adt l))
                                                     (cid (update init_shared_adt l)))
                                           i0).
                rewrite ZMap.gss.
                rewrite ZMap.gso; auto.
                intros Htemp.
                exploit Htemp; eauto.
                rewrite Hdestruct10.
                intro contra; inv contra. }
              eapply abs_rel_flatmem_store_other_related_prop; eauto.
            - unfold relate_AC_per_pd in *; simpl.
              rewrite Hdestruct4; rewrite Hdestruct4 in AC_re; auto.
            - auto.
            - rewrite Hdestruct4; rewrite Hdestruct4 in PT_re; auto.
            - unfold relate_SyncChanPool_per_pd; unfold relate_SyncChanPool_per_pd in syncchpool_re.
              rewrite Hdestruct4; rewrite Hdestruct4 in syncchpool_re; auto.
            - auto.
            - constructor.
              inv inv_re; eauto.
            - constructor. }
          { inv sh_shared_inv_re.
            constructor; simpl; [ | auto | auto | auto | auto | auto | auto | auto | auto | auto | auto].
            intros.
            simpl in ×.
            assert ( i0 : ZIndexed.t, ZMap.get i0 (AbstractDataType.pperm a) = PGUndef
                                             adr: ZIndexed.t,
                                              ZMap.get adr (AbstractDataType.HP a) =
                                              ZMap.get adr (FlatMem.free_page i0 (AbstractDataType.HP a))).
            { intros; eapply dirty_page_shared_inv; eauto. }
            assert (FlatMem.store TY (AbstractDataType.HP a) (PTADDR v (Int.unsigned i)) (rs rd) =
                    FlatMem.store TY (AbstractDataType.HP a) (PTADDR v (Int.unsigned i)) (rs rd)) by reflexivity.
            assert (PTADDR v (Int.unsigned i) mod 4096 4096 - size_chunk TY).
            { eapply PTADDR_mod_lt; auto. }
            assert (ZMap.get (PageI (PTADDR v (Int.unsigned i))) (AbstractDataType.pperm a) = PGAlloc).
            { generalize (per_data_re (ZMap.get (CPU_ID (update init_shared_adt l))
                                                  (cid (update init_shared_adt l)))); intros Htemp.
              rewrite ZMap.gss in Htemp.
              inv Htemp.
              generalize (pperm_re (PageI (PTADDR v (Int.unsigned i)))); intros Htemp´.
              rewrite Hdestruct10 in Htemp´.
              repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              subdestruct; inv Htemp´.
              auto. }
            generalize (single_dirty_ppage´_store_unmapped´
                          (AbstractDataType.pperm a) (AbstractDataType.HP a) H4 _ H7 (rs rd)
                          (FlatMem.store TY (AbstractDataType.HP a) (PTADDR v (Int.unsigned i)) (rs rd)) TY H5 H6).
            intros.
            eapply H8; auto. }
          { inv sh_mem_inv_re.
             rewrite Hflag_PT.
             rewrite <- PT_cond.
             rewrite Hdestruct in sh_cid_re.
             rewrite <- sh_cid_re.
             constructor; simpl; intros.
             case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
            - rewrite ZMap.gss in H4.
              rewrite ZMap.gso in H5; eauto.
              inv H4; simpl in H6.
              eauto.
              eapply pperm_disjoint; eauto.
              rewrite ZMap.gss; auto.
              rewrite ZMap.gso; auto.
            - rewrite ZMap.gso in H4; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              + rewrite ZMap.gss in H5; eauto.
                inv H5.
                simpl.
                eapply pperm_disjoint; eauto.
                rewrite ZMap.gso; eauto.
                rewrite ZMap.gss; eauto.
              + rewrite ZMap.gso in H5; eauto.
                eapply pperm_disjoint in H1; eauto.
                rewrite ZMap.gso; eauto.
                rewrite ZMap.gso; eauto. }

        + rewrite Hflag_init.
          rename e into PT_cond.
          rewrite PT_cond.
          assert (Hptp_re: ZMap.get (AbstractDataType.PT a) (AbstractDataType.ptpool a) = ZMap.get (PT d) (ptpool d)).
          { inv H3.
            generalize (per_data_re (proc_id (uRData l))).
            intros.
            simpl in H1; rewrite ZMap.gss in H1.
            inv H1.
            rewrite <- Hflag_cid in ptp_re.
            rewrite Hflag_PT.
            rewrite <- PT_cond.
            rewrite <- Hflag_cid; eauto.
            rewrite Hdestruct in ptp_re; auto. }
          rewrite <- Hflag_PT.
          rewrite Hptp_re.
          rewrite Hdestruct7, Hdestruct8.
          unfold FlatLoadStoreSem.exec_flatmem_store.
          unfold FlatLoadStoreSemPHB.single_exec_flatmem_store in H; simpl; inv H.
          subdestruct; eauto.
          unfold single_flatmem_store in Hdestruct9; subdestruct.
          unfold flatmem_store.
          simpl in Hdestruct9; inv Hdestruct9.
          assert (Hpperm_re: pperm_inject (pperm d) (AbstractDataType.pperm a)).
          { inv H3; eauto.
            generalize (per_data_re (proc_id (uRData l))); intros.
            simpl; rewrite ZMap.gss in H.
            inv H; auto. }
          unfold pperm_inject in Hpperm_re.
          generalize (Hpperm_re (PageI (PTADDR v (Int.unsigned i)))); intros.
          simpl in Hdestruct10.
          rewrite Hdestruct10 in H.
          clear Hpperm_re.
          subdestruct; try contradiction.
          clear Hdestruct4.
          inv H4.
          rewrite <- Hflag_cid in PT_cond.
           (a {HP: FlatMem.store TY (AbstractDataType.HP a) (PTADDR v (Int.unsigned i)) (rs rd)}).
          split; eauto.
          split; eauto.
          split; eauto.
           kctxt_pool.
          generalize H3; intros Hrelate.
          simpl.
          inv Hrelate; simpl in ×.
          constructor; simpl; auto.
          intros.
          generalize (per_data_re i0).
          intros.
          rewrite Hflag_PT; rewrite <- PT_cond.
          rewrite Hflag_cid.
          case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros; subst.
          { rewrite ZMap.gss in H1.
            rewrite ZMap.gss.
            inv H1.
            constructor; simpl; auto.
            { eapply abs_rel_flatmem_store_mine_related_prop; eauto. }
            { inv inv_re.
              constructor; simpl.
              assert ( i0 : ZIndexed.t, ZMap.get i0 (pperm d) = PGUndef
                                                adr: ZIndexed.t,
                                                 ZMap.get adr (HP d) =
                                                 ZMap.get adr (FlatMem.free_page i0 (HP d))).
              { intros; eapply dirty_page_per_thread_inv; eauto. }
              assert (FlatMem.store TY (HP d) (PTADDR v (Int.unsigned i)) (rs rd) =
                      FlatMem.store TY (HP d) (PTADDR v (Int.unsigned i)) (rs rd)) by reflexivity.
              assert (PTADDR v (Int.unsigned i) mod 4096 4096 - size_chunk TY).
              { eapply PTADDR_mod_lt; auto. }
              generalize (single_dirty_ppage´_store_unmapped´
                            (pperm d) (HP d) H1 _ Hdestruct10 (rs rd)
                            (FlatMem.store TY (HP d) (PTADDR v (Int.unsigned i)) (rs rd)) TY H5 H6).
              intros.
              eapply H7; auto. } }
          { rewrite ZMap.gso in H1; auto.
            rewrite ZMap.gso; auto.
            subdestruct; auto.
            inv H1.
            per_data_auto.
            - assert (ZMap.get (PageI (PTADDR v (Int.unsigned i))) (pperm p) = PGUndef).
              { inv sh_mem_inv_re.
                generalize (pperm_disjoint (ZMap.get (CPU_ID (update init_shared_adt l))
                                                     (cid (update init_shared_adt l)))
                                           i0).
                rewrite ZMap.gss.
                rewrite ZMap.gso; auto.
                intros Htemp.
                exploit Htemp; eauto.
                rewrite Hdestruct10.
                intro contra; inv contra. }
              eapply abs_rel_flatmem_store_other_related_prop; eauto.
            - unfold relate_AC_per_pd in *; simpl.
              rewrite Hdestruct4; rewrite Hdestruct4 in AC_re; auto.
            - auto.
            - rewrite Hdestruct4; rewrite Hdestruct4 in PT_re; auto.
            - unfold relate_SyncChanPool_per_pd; unfold relate_SyncChanPool_per_pd in syncchpool_re.
              rewrite Hdestruct4; rewrite Hdestruct4 in syncchpool_re; auto.
            - auto.
            - constructor.
              inv inv_re; eauto.
            - constructor. }
          { inv sh_shared_inv_re.
            constructor; simpl; [ | auto | auto | auto | auto | auto | auto | auto | auto | auto | auto].
            intros.
            simpl in ×.
            assert ( i0 : ZIndexed.t, ZMap.get i0 (AbstractDataType.pperm a) = PGUndef
                                             adr: ZIndexed.t,
                                              ZMap.get adr (AbstractDataType.HP a) =
                                              ZMap.get adr (FlatMem.free_page i0 (AbstractDataType.HP a))).
            { intros; eapply dirty_page_shared_inv; eauto. }
            assert (FlatMem.store TY (AbstractDataType.HP a) (PTADDR v (Int.unsigned i)) (rs rd) =
                    FlatMem.store TY (AbstractDataType.HP a) (PTADDR v (Int.unsigned i)) (rs rd)) by reflexivity.
            assert (PTADDR v (Int.unsigned i) mod 4096 4096 - size_chunk TY).
            { eapply PTADDR_mod_lt; auto. }
            assert (ZMap.get (PageI (PTADDR v (Int.unsigned i))) (AbstractDataType.pperm a) = PGAlloc).
            { generalize (per_data_re (ZMap.get (CPU_ID (update init_shared_adt l))
                                                  (cid (update init_shared_adt l)))); intros Htemp.
              rewrite ZMap.gss in Htemp.
              inv Htemp.
              generalize (pperm_re (PageI (PTADDR v (Int.unsigned i)))); intros Htemp´.
              rewrite Hdestruct10 in Htemp´.
              repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              subdestruct; inv Htemp´.
              auto. }
            generalize (single_dirty_ppage´_store_unmapped´
                          (AbstractDataType.pperm a) (AbstractDataType.HP a) H4 _ H7 (rs rd)
                          (FlatMem.store TY (AbstractDataType.HP a) (PTADDR v (Int.unsigned i)) (rs rd)) TY H5 H6).
            intros.
            eapply H8; auto. }
          { inv sh_mem_inv_re.
            rewrite Hflag_PT.
            rewrite <- PT_cond.
            rewrite Hdestruct in sh_cid_re.
            rewrite <- sh_cid_re.
            constructor; simpl; intros.
            case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l))
                                      (cid (update init_shared_adt l)))); intros; subst.
            - rewrite ZMap.gss in H4.
              rewrite ZMap.gso in H5; eauto.
              inv H4; simpl in H6.
              eauto.
              eapply pperm_disjoint; eauto.
              rewrite ZMap.gss; auto.
              rewrite ZMap.gso; auto.
            - rewrite ZMap.gso in H4; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              + rewrite ZMap.gss in H5; eauto.
                inv H5.
                simpl.
                eapply pperm_disjoint; eauto.
                rewrite ZMap.gso; eauto.
                rewrite ZMap.gss; eauto.
              + rewrite ZMap.gso in H5; eauto.
                eapply pperm_disjoint in H1; eauto.
                rewrite ZMap.gso; eauto.
                rewrite ZMap.gso; eauto. }

        +
          unfold PageFault.exec_pagefault.
          unfold PageFaultPHB.single_exec_pagefault in H.
          simpl; subdestruct.
          rewrite Hflag_init.
          rewrite Hflag_PT.
          rewrite <- e.
          assert (Hptp_re: ZMap.get (AbstractDataType.PT a) (AbstractDataType.ptpool a) = ZMap.get (PT d) (ptpool d)).
          { inv H3.
            generalize (per_data_re (proc_id (uRData l))).
            intros.
            simpl in H1; rewrite ZMap.gss in H1.
            inv H1.
            rewrite <- Hflag_cid in ptp_re.
            rewrite Hflag_PT.
            rewrite <- e.
            rewrite <- Hflag_cid; eauto.
            rewrite Hdestruct in ptp_re; auto. }
          rewrite e.
          rewrite <- Hflag_PT.
          rewrite Hptp_re.
          rewrite Hdestruct7, Hdestruct8.
          inv H; unfold trapinfo_set; unfold single_trapinfo_set; simpl.
           (a {ti : (i, rs RA)}); split; auto; split; auto; split; auto.
           kctxt_pool.
          inv H3.
          rewrite Hflag_PT; rewrite <- e.
          constructor; simpl; auto.
          intros.
          generalize (per_data_re i0); intros.
          case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros; subst.
          { rewrite ZMap.gss in H; rewrite ZMap.gss.
            inv H; constructor; simpl; auto.
            - rewrite zeq_true; eauto.
            - constructor; simpl; intros.
              inv inv_re.
              apply dirty_page_per_thread_inv; auto.
          }
          { rewrite ZMap.gso in H; auto; rewrite ZMap.gso; auto.
            destruct (ZMap.get i0 dp); auto.
            - inv H; simpl; constructor; simpl; auto.
              + rewrite zeq_false; auto.
                rewrite zeq_false in ti_re; auto.
              + constructor.
                inv inv_re.
                intros.
                apply dirty_page_per_thread_inv; auto.
            - constructor. }
          { inv sh_shared_inv_re.
            constructor; simpl; auto. }
          { inv sh_mem_inv_re.
            constructor; simpl; intros.
            case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l))
                                      (cid (update init_shared_adt l)))); intros; subst.
            - rewrite ZMap.gss in H1.
              rewrite ZMap.gso in H3; eauto.
              inv H1; simpl in H4.
              eauto.
              eapply pperm_disjoint; eauto.
              rewrite ZMap.gss; auto.
              rewrite ZMap.gso; auto.
            - rewrite ZMap.gso in H1; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              + rewrite ZMap.gss in H3; eauto.
                inv H3.
                simpl.
                eapply pperm_disjoint; eauto.
                rewrite ZMap.gso; eauto.
                rewrite ZMap.gss; eauto.
              + rewrite ZMap.gso in H3; eauto.
                eapply pperm_disjoint in H; eauto.
                rewrite ZMap.gso; eauto.
                rewrite ZMap.gso; eauto. }

      -
        simpl in *; a; rewrite Hflag_ihost, Hflag_ikern.
        unfold Asm.exec_store in ×.
        unfold Mem.storev in ×. rewrite Hdestruct3 in H; rewrite Hdestruct3.
        subdestruct; unfold Mem.store in Hdestruct4.
        simpl in Hdestruct4; lift_unfold.
        destruct Hdestruct4; destruct m0; simpl in H4.
        rewrite H1; simpl; inv H; simpl; split; auto.
        simpl in H8; inv H8; split; auto; split; auto.
         kctxt_pool; auto.

      -
        rewrite Hflag_init, Hflag_ihost.
        assert (Htemp: ds´ = update init_shared_adt l m = ).
        { unfold GuestAccessIntelPHB.single_exec_guest_intel_store in H.
          unfold GuestAccessIntelDefPHB.single_exec_guest_intel_accessor, GuestAccessIntelPHB.single_store_accessor in H.
          subdestruct; subst.
          lift_unfold; unfold FlatLoadStoreSemPHB.single_exec_flatmem_store in H; inversion H.
          subdestruct; inv H4; subst.
          inv H.
          unfold single_flatmem_store in Hdestruct10.
          subdestruct.
          inv Hdestruct10; subst; auto. }
        destruct Htemp as (Htemp_a & Htemp_b); subst.
        assert (Hcid_eq: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))
                         = ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)).
        { inv H3.
          rewrite Hdestruct in sh_cid_re.
          auto. }
        rewrite Hcid_eq in ×.
        set (cid := ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)).
        fold cid in H3, H2.
        unfold GuestAccessIntelPHB.single_exec_guest_intel_store in H.
        unfold GuestAccessIntelDefPHB.single_exec_guest_intel_accessor,
        GuestAccessIntelPHB.single_store_accessor in H.
        simpl in H.
        subdestruct.
        assert (Hept_eq: ept d = ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.ept a)).
        { inv H3.
          generalize (per_data_re (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))).
          intros.
          rewrite ZMap.gss in H1.
          inv H1.
          simpl in e.
          rewrite Hdestruct in cid_re.
          rewrite <- cid_re in ept_re.
          rewrite e in ept_re.
          rewrite zeq_true in ept_re.
          auto. }
        unfold GuestAccessIntel2.exec_guest_intel_store2.
        unfold GuestAccessIntelDef2.exec_guest_intel_accessor2,
        GuestAccessIntel2.store_accessor2.
        simpl.
        rewrite <- Hept_eq.
        rewrite Hdestruct5, Hdestruct6, Hdestruct7, Hdestruct8, Hdestruct9.
        rewrite Hdestruct4.
        unfold FlatLoadStoreSemPHB.single_exec_flatmem_store in H.
        subdestruct.
        inv H.
        unfold FlatLoadStoreSem.exec_flatmem_store.
        unfold single_flatmem_store in Hdestruct10.
        subdestruct.
        inv Hdestruct10.
        unfold flatmem_store.
        simpl.
        simpl in Hdestruct11.
        assert (Hpperm: ZMap.get (PageI (EPTADDR (hpa / 4096) (Int.unsigned i))) (AbstractDataType.pperm a) = PGAlloc).
        { inv H3.
          generalize (per_data_re (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))).
          intros.
          rewrite ZMap.gss in H.
          inv H.
          unfold pperm_inject in pperm_re.
          generalize (pperm_re (PageI (EPTADDR (hpa / 4096) (Int.unsigned i)))); intros.
          rewrite Hdestruct11 in H.
          destruct (ZMap.get (PageI (EPTADDR (hpa / 4096) (Int.unsigned i))) (AbstractDataType.pperm a)); auto;
          try contradiction. }
        rewrite Hpperm.
        esplit; eauto.
        refine_split´; eauto.
        instantiate (1:= kctxt_pool).
        generalize H3; intros Hrelate.
        simpl.
        inv Hrelate; simpl in ×.
        constructor; simpl; auto.
        intros.
        generalize (per_data_re i0).
        intros.
        case_eq (zeq i0 cid); intros; subst.
        { rewrite ZMap.gss in H.
          rewrite ZMap.gss.
          inv H.
          constructor; simpl; auto.
          { eapply abs_rel_flatmem_store_mine_related_prop; eauto. }
          { inv inv_re.
            constructor; simpl in ×.
            intros.
            assert ( i0 : ZIndexed.t,
                      ZMap.get i0 (pperm d) = PGUndef
                       adr: ZIndexed.t,
                        ZMap.get adr (HP d) =
                        ZMap.get adr (FlatMem.free_page i0 (HP d))).
            { intros; eapply dirty_page_per_thread_inv; eauto. }
            assert ((FlatMem.store TY (HP d) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd)) =
                    (FlatMem.store TY (HP d) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd))) by reflexivity.
            assert ((EPTADDR (hpa / 4096) (Int.unsigned i)) mod 4096 4096 - size_chunk TY).
            { apply PTADDR_mod_lt; eauto. }
            generalize (single_dirty_ppage´_store_unmapped´ (pperm d) (HP d) H4 _ Hdestruct11 (rs rd)
                             (FlatMem.store TY (HP d) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd)) TY H5 H6).
            intros.
            eapply H7; auto. } }
        { rewrite ZMap.gso in H; auto.
          rewrite ZMap.gso; auto.
          subdestruct; auto.
          inv H.
          per_data_auto.
          - assert (ZMap.get (PageI (EPTADDR (hpa / 4096) (Int.unsigned i))) (pperm p) = PGUndef).
            { inv sh_mem_inv_re.
              generalize (pperm_disjoint cid i0).
              rewrite ZMap.gss.
              rewrite ZMap.gso; auto.
              intros Htemp.
              exploit Htemp; eauto.
              rewrite Hdestruct11.
              intro contra; inv contra. }
            eapply abs_rel_flatmem_store_other_related_prop; eauto.
          - unfold relate_AC_per_pd in *; simpl in ×.
            rewrite Hdestruct10; rewrite Hdestruct10 in AC_re; auto.
          - auto.
          - rewrite Hdestruct10; rewrite Hdestruct10 in PT_re; auto.
          - unfold relate_SyncChanPool_per_pd; unfold relate_SyncChanPool_per_pd in syncchpool_re.
            rewrite Hdestruct10; rewrite Hdestruct10 in syncchpool_re; auto.
          - auto.
          - constructor.
            inv inv_re.
            intros.
            apply dirty_page_per_thread_inv; auto.
          - constructor. }
        { inv sh_shared_inv_re.
          constructor; simpl; [ | auto | auto | auto | auto | auto | auto | auto | auto | auto | auto].
          intros.
          simpl in ×.
          assert ( i0 : ZIndexed.t, ZMap.get i0 (AbstractDataType.pperm a) = PGUndef
                                           adr: ZIndexed.t,
                                            ZMap.get adr (AbstractDataType.HP a) =
                                            ZMap.get adr (FlatMem.free_page i0 (AbstractDataType.HP a))).
          { intros; eapply dirty_page_shared_inv; eauto. }
          assert (FlatMem.store TY (AbstractDataType.HP a) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd) =
                  FlatMem.store TY (AbstractDataType.HP a) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd)) by reflexivity.
          assert ((EPTADDR (hpa / 4096)) (Int.unsigned i) mod 4096 4096 - size_chunk TY).
          { eapply PTADDR_mod_lt; auto. }
          assert (ZMap.get (PageI (EPTADDR (hpa / 4096) (Int.unsigned i))) (AbstractDataType.pperm a) = PGAlloc).
          { generalize (per_data_re cid); intros Htemp.
            rewrite ZMap.gss in Htemp.
            inv Htemp.
            generalize (pperm_re (PageI (EPTADDR (hpa /4096) (Int.unsigned i)))); intros Htemp´.
            rewrite Hdestruct11 in Htemp´.
            repeat match goal with
                   | [H: context[if _ then _ else _] |- _] ⇒ clear H
                   end.
            subdestruct; inv Htemp´.
            auto. }
          generalize (single_dirty_ppage´_store_unmapped´
                        (AbstractDataType.pperm a) (AbstractDataType.HP a) H1 _ H6 (rs rd)
                        (FlatMem.store TY (AbstractDataType.HP a) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd)) TY H4 H5).
          intros.
          eapply H7; auto. }
        { inv sh_mem_inv_re.
          constructor; simpl; intros.
          case_eq (zeq i0 cid); intros; subst.
          - rewrite ZMap.gss in H1.
            rewrite ZMap.gso in H4; eauto.
            inv H1; simpl in H5.
            eauto.
            eapply pperm_disjoint; eauto.
            rewrite ZMap.gss; auto.
            rewrite ZMap.gso; auto.
          - rewrite ZMap.gso in H1; auto.
            case_eq (zeq j cid); intros; subst.
            + rewrite ZMap.gss in H4; eauto.
              inv H4.
              simpl.
              eapply pperm_disjoint; eauto.
              rewrite ZMap.gso; eauto.
              rewrite ZMap.gss; eauto.
            + rewrite ZMap.gso in H4; eauto.
              eapply pperm_disjoint in H; eauto.
              rewrite ZMap.gso; eauto.
              rewrite ZMap.gso; eauto. }
      -
        rewrite Hflag_init, Hflag_ihost.
        assert (Htemp: ds´ = update init_shared_adt l m = ).
        { unfold GuestAccessIntelPHB.single_exec_guest_intel_store in H.
          unfold GuestAccessIntelDefPHB.single_exec_guest_intel_accessor, GuestAccessIntelPHB.single_store_accessor in H.
          subdestruct; subst.
          lift_unfold; unfold FlatLoadStoreSemPHB.single_exec_flatmem_store in H; inversion H.
          subdestruct; inv H4; subst.
          inv H.
          unfold single_flatmem_store in Hdestruct10.
          subdestruct.
          inv Hdestruct10; subst; auto. }
        destruct Htemp as (Htemp_a & Htemp_b); subst.
        assert (Hcid_eq: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))
                         = ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)).
        { inv H3.
          rewrite Hdestruct in sh_cid_re.
          auto. }
        rewrite Hcid_eq in ×.
        set (cid := ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)).
        fold cid in H3, H2.
        unfold GuestAccessIntelPHB.single_exec_guest_intel_store in H.
        unfold GuestAccessIntelDefPHB.single_exec_guest_intel_accessor,
        GuestAccessIntelPHB.single_store_accessor in H.
        simpl in H.
        subdestruct.
        assert (Hept_eq: ept d = ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.ept a)).
        { inv H3.
          generalize (per_data_re (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))).
          intros.
          rewrite ZMap.gss in H1.
          inv H1.
          simpl in e.
          rewrite Hdestruct in cid_re.
          rewrite <- cid_re in ept_re.
          rewrite e in ept_re.
          rewrite zeq_true in ept_re.
          auto. }
        unfold GuestAccessIntel2.exec_guest_intel_store2.
        unfold GuestAccessIntelDef2.exec_guest_intel_accessor2,
        GuestAccessIntel2.store_accessor2.
        simpl.
        rewrite <- Hept_eq.
        rewrite Hdestruct5, Hdestruct6, Hdestruct7, Hdestruct8, Hdestruct9.
        rewrite Hdestruct4.
        unfold FlatLoadStoreSemPHB.single_exec_flatmem_store in H.
        subdestruct.
        inv H.
        unfold FlatLoadStoreSem.exec_flatmem_store.
        unfold single_flatmem_store in Hdestruct10.
        subdestruct.
        inv Hdestruct10.
        unfold flatmem_store.
        simpl.
        simpl in Hdestruct11.
        assert (Hpperm: ZMap.get (PageI (EPTADDR (hpa / 4096) (Int.unsigned i))) (AbstractDataType.pperm a) = PGAlloc).
        { inv H3.
          generalize (per_data_re (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))).
          intros.
          rewrite ZMap.gss in H.
          inv H.
          unfold pperm_inject in pperm_re.
          generalize (pperm_re (PageI (EPTADDR (hpa / 4096) (Int.unsigned i)))); intros.
          rewrite Hdestruct11 in H.
          destruct (ZMap.get (PageI (EPTADDR (hpa / 4096) (Int.unsigned i))) (AbstractDataType.pperm a)); auto;
          try contradiction. }
        rewrite Hpperm.
        esplit; eauto.
        refine_split´; eauto.
        instantiate (1:= kctxt_pool).
        generalize H3; intros Hrelate.
        simpl.
        inv Hrelate; simpl in ×.
        constructor; simpl; auto.
        intros.
        generalize (per_data_re i0).
        intros.
        case_eq (zeq i0 cid); intros; subst.
        { rewrite ZMap.gss in H.
          rewrite ZMap.gss.
          inv H.
          constructor; simpl; auto.
          { eapply abs_rel_flatmem_store_mine_related_prop; eauto. }
          { inv inv_re.
            constructor; simpl in ×.
            intros.
            assert ( i0 : ZIndexed.t,
                      ZMap.get i0 (pperm d) = PGUndef
                       adr: ZIndexed.t,
                        ZMap.get adr (HP d) =
                        ZMap.get adr (FlatMem.free_page i0 (HP d))).
            { intros; eapply dirty_page_per_thread_inv; eauto. }
            assert ((FlatMem.store TY (HP d) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd)) =
                    (FlatMem.store TY (HP d) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd))) by reflexivity.
            assert ((EPTADDR (hpa / 4096) (Int.unsigned i)) mod 4096 4096 - size_chunk TY).
            { apply PTADDR_mod_lt; eauto. }
            generalize (single_dirty_ppage´_store_unmapped´ (pperm d) (HP d) H4 _ Hdestruct11 (rs rd)
                             (FlatMem.store TY (HP d) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd)) TY H5 H6).
            intros.
            eapply H7; auto. } }
        { rewrite ZMap.gso in H; auto.
          rewrite ZMap.gso; auto.
          subdestruct; auto.
          inv H.
          per_data_auto.
          - assert (ZMap.get (PageI (EPTADDR (hpa / 4096) (Int.unsigned i))) (pperm p) = PGUndef).
            { inv sh_mem_inv_re.
              generalize (pperm_disjoint cid i0).
              rewrite ZMap.gss.
              rewrite ZMap.gso; auto.
              intros Htemp.
              exploit Htemp; eauto.
              rewrite Hdestruct11.
              intro contra; inv contra. }
            eapply abs_rel_flatmem_store_other_related_prop; eauto.
          - unfold relate_AC_per_pd in *; simpl in ×.
            rewrite Hdestruct10; rewrite Hdestruct10 in AC_re; auto.
          - auto.
          - rewrite Hdestruct10; rewrite Hdestruct10 in PT_re; auto.
          - unfold relate_SyncChanPool_per_pd; unfold relate_SyncChanPool_per_pd in syncchpool_re.
            rewrite Hdestruct10; rewrite Hdestruct10 in syncchpool_re; auto.
          - auto.
          - constructor.
            inv inv_re.
            intros.
            apply dirty_page_per_thread_inv; auto.
          - constructor. }
        { inv sh_shared_inv_re.
          constructor; simpl; [ | auto | auto | auto | auto | auto | auto | auto | auto | auto | auto].
          intros.
          simpl in ×.
          assert ( i0 : ZIndexed.t, ZMap.get i0 (AbstractDataType.pperm a) = PGUndef
                                           adr: ZIndexed.t,
                                            ZMap.get adr (AbstractDataType.HP a) =
                                            ZMap.get adr (FlatMem.free_page i0 (AbstractDataType.HP a))).
          { intros; eapply dirty_page_shared_inv; eauto. }
          assert (FlatMem.store TY (AbstractDataType.HP a) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd) =
                  FlatMem.store TY (AbstractDataType.HP a) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd)) by reflexivity.
          assert ((EPTADDR (hpa / 4096)) (Int.unsigned i) mod 4096 4096 - size_chunk TY).
          { eapply PTADDR_mod_lt; auto. }
          assert (ZMap.get (PageI (EPTADDR (hpa / 4096) (Int.unsigned i))) (AbstractDataType.pperm a) = PGAlloc).
          { generalize (per_data_re cid); intros Htemp.
            rewrite ZMap.gss in Htemp.
            inv Htemp.
            generalize (pperm_re (PageI (EPTADDR (hpa /4096) (Int.unsigned i)))); intros Htemp´.
            rewrite Hdestruct11 in Htemp´.
            repeat match goal with
                   | [H: context[if _ then _ else _] |- _] ⇒ clear H
                   end.
            subdestruct; inv Htemp´.
            auto. }
          generalize (single_dirty_ppage´_store_unmapped´
                        (AbstractDataType.pperm a) (AbstractDataType.HP a) H1 _ H6 (rs rd)
                        (FlatMem.store TY (AbstractDataType.HP a) (EPTADDR (hpa / 4096) (Int.unsigned i)) (rs rd)) TY H4 H5).
          intros.
          eapply H7; auto. }
        { inv sh_mem_inv_re.
          constructor; simpl; intros.
          case_eq (zeq i0 cid); intros; subst.
          - rewrite ZMap.gss in H1.
            rewrite ZMap.gso in H4; eauto.
            inv H1; simpl in H5.
            eauto.
            eapply pperm_disjoint; eauto.
            rewrite ZMap.gss; auto.
            rewrite ZMap.gso; auto.
          - rewrite ZMap.gso in H1; auto.
            case_eq (zeq j cid); intros; subst.
            + rewrite ZMap.gss in H4; eauto.
              inv H4.
              simpl.
              eapply pperm_disjoint; eauto.
              rewrite ZMap.gso; eauto.
              rewrite ZMap.gss; eauto.
            + rewrite ZMap.gso in H4; eauto.
              eapply pperm_disjoint in H; eauto.
              rewrite ZMap.gso; eauto.
              rewrite ZMap.gso; eauto. }

      -
        simpl in *; a; rewrite Hflag_ihost, Hflag_ikern.
        unfold Asm.exec_store in ×.
        unfold Mem.storev in ×. rewrite Hdestruct2 in H; rewrite Hdestruct2.
        subdestruct; unfold Mem.store in Hdestruct3.
        simpl in Hdestruct3; lift_unfold.
        destruct Hdestruct3; destruct m0; simpl in H4.
        rewrite H1; simpl; inv H; simpl; split; auto.
        simpl in H8; inv H8; split; auto; split; auto.
         kctxt_pool; auto.
    Fail idtac.
    Qed.
    Next Obligation.
      rename H4 into source_kctxt.
      assert (cid_eq: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))
                      = ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)).
      { inv H19.
        unfold state_checks in H11.
        simpl in H11.
        unfold state_check in H11.
        rewrite H5 in H11.
        rewrite peq_true in H11.
        revert H11.
        revert sh_cid_re.
        clear.
        intros.
        unfold thread_yield_state_check in H11.
        subdestruct_one; try (inv H11; auto).
        simpl in ×.
        rewrite Hdestruct in sh_cid_re.
        rewrite sh_cid_re; auto. }
      generalize H11; intro Hstate_checks.
      unfold state_checks in H11.
      simpl in H11.
      unfold state_check in H11.
      rewrite H5 in H11.
      rewrite peq_true in H11.
      unfold thread_yield_state_check in H11.
      subdestruct; try inv H11; try contradiction.
      rewrite cid_eq in ×.
      simpl.
      simpl in Hdestruct.
      rewrite Hdestruct.
      simpl in H7, H8.
      assert (CPU_ID_eq: (CPU_ID (update init_shared_adt l)) = (AbstractDataType.CPU_ID a)).
      { inv H19; rewrite sh_CPU_ID_re; auto. }
      assert (big_log_eq: (big_log (update init_shared_adt l)) = (AbstractDataType.big_log a)).
      { inv H19; rewrite sh_big_log_re; auto. }
      assert (Htid_val: ti d = init_trap_info).
      { assert (fst (ti d) = Int.zero).
        { generalize Hdestruct8.
          clear.
          intros.
          unfold Int.eq in Hdestruct8.
          subdestruct.
          unfold Int.zero; unfold Int.zero in e.
          clear Hdestruct.
          eapply Int_unsigned_eq in e.
          symmetry; auto. }
        simpl in Hdestruct8.
        unfold init_trap_info.
        rewrite <- H4.
        destruct (ti d).
        subst; simpl.
        f_equal.
        simpl in Hdestruct12; rewrite Hdestruct12.
        unfold Vzero.
        simpl in Hdestruct13.
        rewrite <- (Int.eq_true Int.zero) in Hdestruct13.
        unfold Int.eq in Hdestruct13.
        clear H7 H8.
        subdestruct.
        clear Hdestruct15.
        apply Int_unsigned_eq in e.
        rewrite e; auto. }
      
      rewrite <- Hdestruct10 in Hdestruct6.
      rewrite <- cid_eq in Hdestruct14.
      rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct10, Hdestruct14, Hdestruct16 in H7, H8.
      rewrite <- cid_eq in Hdestruct11.
      rewrite Hdestruct17, Hdestruct18, Hdestruct19, Hdestruct20 in H7, H8.
      simpl in H7.
      simpl in ×.
      set (rs1 := (Pregmap.init Vundef)#ESP <- (rs#ESP)#EDI <- (rs#EDI)#ESI <- (rs#ESI)
                                       #EBX <- (rs#EBX)#EBP <- (rs#EBP)#RA <- (rs#RA)).
      set (rs´ := ZMap.get run_id (kctxt a)).

      rewrite ZMap.gss in H7.
      rewrite ZMap.gss in H8.

      assert (Hproc_id: Constant.TOTAL_CPU <
                        ZMap.get (CPU_ID (update init_shared_adt l))
                                 (cid (update init_shared_adt l)) < Constant.num_proc
                        ZMap.get (CPU_ID (update init_shared_adt l))
                                 (cid (update init_shared_adt l)) = main_thread).
      { eapply valid_thread_list.
        exploit all_cid_in_full_thread_list.
        simpl.
        instantiate (1 := l); intro; auto. }
      assert (Hrun_id_range: Constant.TOTAL_CPU < run_id < Constant.num_proc run_id = main_thread).
      { assert (Hnew_id:
                  ZMap.get (CPU_ID (update init_shared_adt (LEvent (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))
                                                                   (LogYield (Mem.nextblock m)) :: l)))
                           (cid (update init_shared_adt (LEvent (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))
                                                                (LogYield (Mem.nextblock m)) :: l))) = run_id).
        { simpl.
          rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct10, Hdestruct14, Hdestruct16.
          rewrite Hdestruct17, Hdestruct18, Hdestruct19, Hdestruct20.
          simpl; rewrite ZMap.gss; auto. }
        rewrite <- Hnew_id.
        eapply valid_thread_list.
        exploit all_cid_in_full_thread_list.
        simpl.
        instantiate (1 := LEvent (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))
                                 (LogYield (Mem.nextblock m)) :: l); intros; auto. }
      assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
      { rewrite main_thread_val.
        generalize current_CPU_ID_range.
        clear.
        intros.
        omega. }

      assert (Hcurid_cused: cused (ZMap.get (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))
                                            (AbstractDataType.AC a)) = true).
      { inv H19.
        inv sh_shared_inv_re.
        rewrite <- cid_eq.
        assert (In (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                   full_thread_list).
        { generalize all_cid_in_full_thread_list.
          simpl; intros Htemp; apply Htemp. }
        rewrite cid_eq in H4.
        rewrite cid_eq.
        generalize (container_used_inv (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)) H4).
        unfold B_GetContainerUsed´.
        rewrite Hdestruct10.
        intros Htemp ; rewrite <- Htemp.
        rewrite <- cid_eq.
        auto. }
      
      assert (Hdes_cused : cused (ZMap.get run_id (AbstractDataType.AC a)) = true).
      { subst; inv H19.
        assert (proc_id (uRData (LEvent (proc_id (uRData l)) (LogYield (Mem.nextblock m))::l)) = run_id).
        { simpl in ×.
          rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct10, Hdestruct14, Hdestruct16.
          unfold B_GetContainerUsed_opt.
          rewrite Hdestruct17.
          destruct msg; [ | inv Hdestruct18].
          simpl in Hdestruct18; rewrite Hdestruct18, Hdestruct19, Hdestruct20.
          simpl.
          rewrite ZMap.gss.
          auto. }
        inv sh_shared_inv_re.
        assert (In (proc_id (uRData (LEvent (proc_id (uRData l)) (LogYield (Mem.nextblock m))::l)))
                   full_thread_list).
        { apply all_cid_in_full_thread_list. }
        generalize (container_used_inv (proc_id (uRData (LEvent (proc_id (uRData l)) (LogYield (Mem.nextblock m))::l))) H4).
        unfold B_GetContainerUsed´.
        rewrite Hdestruct10.
        intros Htemp ; rewrite <- Htemp.
        auto. }

      unfold B_GetContainerUsed_opt in Hdestruct18.
      destruct msg; try (inv Hdestruct18; fail).

       (a {big_log: BigDef (TBEVENT (AbstractDataType.CPU_ID a)
                                          (TBSHARED (BTDYIELD (ZMap.get (AbstractDataType.CPU_ID a)
                                                                        (AbstractDataType.cid a))))
                                          :: (AbstractDataType.big_oracle a) (AbstractDataType.CPU_ID a) l0 ++ l0)}
                {kctxt: ZMap.set (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)) rs1 (AbstractDataType.kctxt a)}
                {cid : ZMap.set (AbstractDataType.CPU_ID a) run_id (AbstractDataType.cid a)}).

      remember (a {big_log: BigDef (TBEVENT (AbstractDataType.CPU_ID a)
                                            (TBSHARED (BTDYIELD (ZMap.get (AbstractDataType.CPU_ID a)
                                                                        (AbstractDataType.cid a))))
                                          :: (AbstractDataType.big_oracle a) (AbstractDataType.CPU_ID a) l0 ++ l0)}
                  {kctxt: ZMap.set (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)) rs1 (AbstractDataType.kctxt a)}
                  {cid : ZMap.set (AbstractDataType.CPU_ID a) run_id (AbstractDataType.cid a)}) as .
      assert (Hassumptions: ( v r, ZtoPreg v = Some rVal.has_type (des_rs´#r) AST.Tint)
                            ( v r, ZtoPreg v = Some r
                                                       val_inject (Mem.flat_inj (Mem.nextblock m)) (des_rs´#r) (des_rs´#r))
                            Hthread_sched_sem_hyp rs s m a des_rs´ ZMap.get run_id (kctxt a) = des_rs´
                            (ZMap.get run_id rsm = EAsmCommon.Running des_rs´
                              pc, initial_thread_pc (spl_data := single_data) ge run_id l = Some pc)).
      { assert (Hev_yield_eq: LEvent (proc_id (uRData l)) (LogYield (Mem.nextblock m)) =
                              LEvent (proc_id (uRData l)) (LogYield (Mem.nextblock m))) by reflexivity.
        assert (Hcommon_cond: Hthread_sched_trans_common_cond ge source_kctxt l
                                                               (LEvent (proc_id (uRData l)) (LogYield (Mem.nextblock m)))
                                                               dp d a s thread_yield b sg rsm rs des_rs´).
        { unfold Hthread_sched_trans_common_cond.
          refine_split; auto.
          - intros.
            simpl in H4.
            rewrite cid_eq in H4.
            eapply H0; eauto.
          - simpl; rewrite cid_eq; auto.
          - simpl; rewrite cid_eq; auto.
          - simpl.
            rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct10, Hdestruct14, Hdestruct16.
            unfold B_GetContainerUsed_opt.
            rewrite Hdestruct17, Hdestruct18, Hdestruct19, Hdestruct20.
            simpl.
            rewrite ZMap.gss; auto.
          - simpl.
            rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct10, Hdestruct14, Hdestruct16.
            unfold B_GetContainerUsed_opt.
            rewrite Hdestruct17, Hdestruct18, Hdestruct19, Hdestruct20.
            simpl.
            rewrite ZMap.gss; auto. }
        
        generalize (Hthread_sched_match_prop ge source_kctxt l dp d a s b sg
                                             rsm rs des_rs´ m (LEvent (proc_id (uRData l)) (LogYield (Mem.nextblock m)))
                                             Hstate_checks Hev_yield_eq Hcommon_cond).
        simpl.
        rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct10, Hdestruct14, Hdestruct16.
        unfold B_GetContainerUsed_opt.
        rewrite Hdestruct17, Hdestruct18, Hdestruct19, Hdestruct20.
        simpl.
        rewrite ZMap.gss; auto.
        intros (? & ? & ? & ?).
        repeat split; trivial.
        eapply Hthread_sched_sem_hyp_correct.
      }
      destruct Hassumptions as (N_TYPE & N_INJECT_NEUTRAL & Hthread_sched_sem & Htrans_kctxt & Hinit_exist_aux); subst.

      refine_split´; eauto.
      -
        replace ((Pos.to_nat (Mem.nextblock m) - Pos.to_nat (Mem.nextblock m))%nat) with O.
        simpl.
        assert ((@mem_lift_nextblock mem memory_model_ops Hmem
                                     (@AlgebraicMemImpl.algebraicmem mem memory_model_ops Hmem) m O) = m).
        { remember (mem_lift_nextblock m 0) as .
          symmetry in Heqm´.
          eapply mem_lift_nextblock_source_eq; eauto. }
        rewrite H4.
        set (rs2 := ((((((Pregmap.init Vundef) # ESP <- (rs ESP)) # EDI <- (rs EDI))
                         # ESI <- (rs ESI)) # EBX <- (rs EBX)) # EBP <- (rs EBP)) # RA <- (rs RA)).
        set ( := (a {big_log: BigDef (TBEVENT (AbstractDataType.CPU_ID a)
                                                (TBSHARED (BTDYIELD (ZMap.get (AbstractDataType.CPU_ID a)
                                                                              (AbstractDataType.cid a))))
                                                :: (AbstractDataType.big_oracle a) (AbstractDataType.CPU_ID a) l0 ++ l0)}
                      {kctxt: ZMap.set (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)) rs1
                                       (AbstractDataType.kctxt a)}
                      {cid : ZMap.set (AbstractDataType.CPU_ID a) run_id (AbstractDataType.cid a)})).
        generalize (primcall_threadschedule_sem_intro (prim_ident := thread_yield) big_thread_yield_spec s m rs
                                                      (ZMap.get run_id (kctxt a)) rs2 a b).
        intros Hsem_intro.
        exploit Hsem_intro; eauto.
        + unfold big_thread_yield_spec.
          clear H4.
          generalize H19; intro Hrelated_save.
          substx.
          inv H19.
          generalize (per_data_re (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))).
          intros priv_data_re.
          rewrite cid_eq in priv_data_re.
          rewrite H5 in priv_data_re.
          inv priv_data_re.
          unfold relate_AC_per_pd in ×.
          unfold proc_id in *; simpl in ×.
          rewrite cid_eq in ×.
          rewrite Hdestruct in ×.
          rewrite zeq_true in ×.
          rewrite <- pg_re, <- ikern_re, <- ihost_re, <- ipt_re, <- big_log_re, <- CPU_ID_re, <- lock_re.
          rewrite Hdestruct0, Hdestruct1, Hdestruct2, Hdestruct3, Hdestruct7, Hdestruct10.
          rewrite CPU_ID_re.
          rewrite Hcurid_cused.
          rewrite CPU_ID_re in Hdestruct14.
          rewrite sh_big_oracle_re in Hdestruct14.
          rewrite Hdestruct14.
          rewrite Hdestruct16.
          unfold IsCused.

          assert (Hz_cused: cused (ZMap.get z (AbstractDataType.AC a)) = true).
          { inv sh_shared_inv_re.
            assert (In z full_thread_list).
            { revert Hdestruct19; clear.
              intros.
              case_eq (in_dec zeq z full_thread_list); intros; subst; auto.
              rewrite H in Hdestruct19.
              inv Hdestruct19. }
            generalize (container_used_inv z H4).
            intros.
            unfold B_GetContainerUsed´ in H11.
            rewrite Hdestruct10 in H11.
            rewrite Hdestruct18 in H11.
            symmetry; auto. }
          rewrite Hz_cused.
          repeat f_equal.

        + clear.
          remember (Pos.to_nat (Mem.nextblock m)) as val.
          clear.
          induction val.
          simpl; auto.
          simpl; auto.

      - intros; simpl.
        simpl in H4, H11.
        unfold match_estate_kctxt; simpl in ×.
        intros.
        inv H11.
        + simpl.
          rewrite ZMap.gso in H13; eauto.
          destruct H13.
          unfold initial_thread_kctxt in H13.
          subdestruct.
          inv H13.
          simpl in Hdestruct9.
          subdestruct; inv Hdestruct9.
          unfold initial_regset_kctxt.
          simpl.
          rewrite ZMap.gss in H4; auto.
          case_eq (zeq (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
          × rewrite H6 in H11; inv H11.
          × rewrite ZMap.gso; auto.
            unfold match_estate_kctxt in H0.
            unfold match_estate_kctxt in H0.
            clear H13.
            eapply H0 in n0; eauto.
            unfold match_estate_regset.
            left.
            split; eauto.
            unfold initial_thread_kctxt.
            simpl in Hdestruct15.
            subdestruct; inv Hdestruct15.
            unfold initial_regset_kctxt.
            rewrite Hdestruct9.
            auto.
        + rewrite ZMap.gso in H13; eauto.
          case_eq (zeq (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
          × rewrite ZMap.gss; auto.
            unfold rs1.
            rewrite H6 in H13.
            substx.
            inv H13.
            inv H12.
            unfold PregtoZ in H14.
            subdestruct; try inv H14; reflexivity.
          × intros; subst.
            rewrite ZMap.gso; eauto.
            clear H11.
            unfold match_estate_regset in H0.
            unfold match_estate_kctxt in H0.
            eapply H0 in n0; eauto.
      - simpl.
        intros.
        unfold thread_init_invariant; intros; simpl.
        unfold initial_thread_kctxt; simpl; intros.
        unfold thread_init_pc_of.
        unfold thread_init_pc.
        simpl in H4.
        unfold thread_init_invariant in H9; simpl in ×.
        rewrite ZMap.gss in H4.
        case_eq (zeq i0 run_id).
        + intros; subst.
          rewrite ZMap.gss in H4.
          destruct H7.
          × destruct H7.
            unfold initial_thread_kctxt in H12.
            simpl in H12.
            subdestruct.
          × inv H4.
            unfold initial_regset_kctxt.
            assert (Hinit_exist: pc,
                       @initial_thread_pc (@single_data oracle_ops big_ops) mem memory_model_ops
                                          (@SingleOracleImpl.s_config
                                             s_oracle_ops s_threads_ops real_params_ops
                                             real_params oracle_ops6
                                             oracle_ops big_ops s_threads_prf
                                             s_oracle_basic_prop_prf mem memory_model_ops Hmem) fundef
                                          unit ge run_id l
                       = Some pc).
            { eapply Hinit_exist_aux; eauto. }
            destruct Hinit_exist as (pc´, Hinit_exist).
            rewrite Hinit_exist.
            intro contra; inv contra.
        + intros.
          rewrite ZMap.gso in H4; auto.
          eapply H9 in H4.
          unfold initial_thread_kctxt in H4.
          simpl; intro contra.
          elim H4; subdestruct; reflexivity.
      - simpl.
        instantiate (1:= ZMap.set (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)) rs1 (kctxt a)).
        generalize H19; intro Hrelated.
        inv H19.
        constructor; simpl; auto.
        +
          rewrite <- sh_CPU_ID_re.
          rewrite Hdestruct.
          repeat rewrite ZMap.gss.
          reflexivity.
        +
          rewrite sh_big_oracle_re, sh_CPU_ID_re; auto.
        + intros i0.
          generalize (per_data_re (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))).
          rewrite H5.
          intros.
          case_eq (@ZMap.get (option privData) i0 dp); intros.
          × inversion H4.
            unfold relate_AC_per_pd in ×.
            rewrite Hdestruct in ×.
            rewrite Hdestruct9 in ×.
            simpl in ×.
            rewrite cid_eq in ×.
            rewrite zeq_true in ×.
            unfold B_GetContainerUsed´ in AC_re.
            rename AC_re into AC_re´.
            assert (AC_re: AC d = ZMap.get (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))
                                           (AbstractDataType.AC a)).
            { destruct Hproc_id as [Hproc_id | Hproc_id]; auto.
              + assert (Hno_main: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) main_thread).
                { revert Hproc_id.
                  revert Hmain_thread.
                  rewrite cid_eq.
                  clear; intros.
                  omega. }
                rewrite <- cid_eq in AC_re´.
                rewrite zeq_false in AC_re´; try assumption.
                rewrite zlt_lt_true in AC_re´; try (rewrite cid_eq; assumption).
                rewrite cid_eq in AC_re´.
                rewrite Hdestruct10, Hdestruct11 in AC_re´.
                auto.
              + rewrite Hproc_id in AC_re´.
                rewrite zeq_true in AC_re´; try assumption.
                rewrite Hdestruct10 in AC_re´.
                rewrite <- Hproc_id in AC_re´.
                rewrite Hdestruct11 in AC_re´.
                auto. }
            clear AC_re´.

            case_eq (zeq i0 run_id).
            {
              intros; subst.
              rewrite H8 in H11.
              inv H11.
              generalize (per_data_re run_id); intros.
              rewrite H8 in H11.
              clear H12.
              inv H11.
              rewrite Hdestruct in ×.
              per_data_auto_simpl.
              -
                rewrite Hdestruct; auto.
              -
                rewrite Hdestruct.
                repeat rewrite ZMap.gss; reflexivity.
              -
                rewrite big_oracle_re, CPU_ID_re; auto.
              -
                simpl.
                rewrite ZMap.gss.
                rewrite zeq_true.
                simpl in ikern_re0.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in ikern_re0.
                  rewrite ikern_re0; reflexivity.
                + rewrite zeq_false in ikern_re0; auto.
                  rewrite ikern_re0.
                  rewrite <- ikern_re.
                  rewrite Hdestruct1.
                reflexivity.
              -
                simpl.
                rewrite ZMap.gss.
                rewrite zeq_true.
                simpl in ikern_re0.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in ihost_re0.
                  rewrite ihost_re0; reflexivity.
                + rewrite zeq_false in ihost_re0; auto.
                  rewrite ihost_re0.
                  rewrite <- ihost_re.
                  rewrite Hdestruct2.
                  reflexivity.
              -
                unfold relate_AC_per_pd; simpl.
                rewrite Hdestruct.
                unfold relate_AC_per_pd in AC_re0; rewrite Hdestruct in AC_re0.
                unfold B_GetContainerUsed´ in AC_re0.
                rewrite Hdestruct10 in AC_re0.
                assert (B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l)) l0
                        = B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l))
                                             (TBEVENT (CPU_ID (update init_shared_adt l))
                                                      (TBSHARED
                                                         (BTDYIELD
                                                            (ZMap.get (AbstractDataType.CPU_ID a)
                                                                      (AbstractDataType.cid a))))
                                                      :: big_oracle (update init_shared_adt l)
                                                      (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                { assert (B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l)) l0
                          = B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l))
                                                (big_oracle (update init_shared_adt l)
                                                            (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                  { eapply single_big_oracle_query_preserves_cused.
                    inv sh_shared_inv_re; auto. }
                  unfold B_GetContainerUsed in H11.
                  unfold B_GetContainerUsed.
                  simpl; auto. }
                rewrite <- H11; clear H11.
                auto.
              -
                simpl.
                rewrite ZMap.gss.
                rewrite zeq_true.
                simpl in ti_re0.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in ti_re0.
                  rewrite ti_re0; reflexivity.
                + rewrite zeq_false in ti_re0; auto.
                  rewrite ti_re0.
                  rewrite <- ti_re.
                  rewrite Htid_val; auto.
              -
                auto.
              -
                rewrite Hdestruct.
                rewrite ZMap.gss.
                rewrite zeq_true.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in PT_re0.
                  rewrite PT_re0; reflexivity.
                + rewrite zeq_false in PT_re0; auto.
                  rewrite PT_re0.
                  rewrite <- PT_re.
                  reflexivity.
              -
                rewrite Hdestruct.
                assumption.
              -
                rewrite ZMap.gss.
                rewrite zeq_true.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in ipt_re0.
                  rewrite ipt_re0; reflexivity.
                + rewrite zeq_false in ipt_re0; auto.
                  rewrite ipt_re0.
                  rewrite <- ipt_re.
                  rewrite Hdestruct3.
                  reflexivity.
              -
                rewrite ZMap.gss.
                rewrite zeq_true.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in intr_flag_re0.
                  rewrite intr_flag_re0; reflexivity.
                + rewrite zeq_false in intr_flag_re0; auto.
                  rewrite intr_flag_re0.
                  rewrite <- intr_flag_re.
                  rewrite Hdestruct5.
                  reflexivity.
              -
                rewrite ZMap.gss.
                rewrite zeq_true.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in in_intr_re0.
                  rewrite in_intr_re0; reflexivity.
                + rewrite zeq_false in in_intr_re0; auto.
                  rewrite in_intr_re0.
                  rewrite <- in_intr_re.
                  rewrite Hdestruct4.
                  reflexivity.
              - unfold relate_SyncChanPool_per_pd.
                simpl; rewrite Hdestruct; auto.
                rewrite zeq_true; auto.
              - unfold relate_uctxt_per_pd in *; simpl in ×.
                rewrite Hdestruct in ×.
                unfold B_GetContainerUsed´ in uctxt_re0.
                rewrite Hdestruct10 in uctxt_re0.
                assert (B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l)) l0
                        = B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l))
                                             (TBEVENT (CPU_ID (update init_shared_adt l))
                                                      (TBSHARED
                                                         (BTDYIELD
                                                            (ZMap.get (AbstractDataType.CPU_ID a)
                                                                      (AbstractDataType.cid a))))
                                                      :: big_oracle (update init_shared_adt l)
                                                      (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                { assert (B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l)) l0
                          = B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l))
                                                (big_oracle (update init_shared_adt l)
                                                            (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                  { eapply single_big_oracle_query_preserves_cused.
                    inv sh_shared_inv_re; auto. }
                  unfold B_GetContainerUsed in H11.
                  unfold B_GetContainerUsed.
                  simpl; auto. }
                rewrite <- H11; clear H11.
                auto.
              - auto.
              - auto.
              - auto.
              - constructor; simpl.
                inv inv_re0.
                auto. }
            {
              repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              intros; subst.
              generalize (per_data_re i0).
              rewrite H11.
              intros.
              inv H13.
              per_data_auto_simpl.
              -
                rewrite Hdestruct.
                repeat rewrite ZMap.gss; reflexivity.
              -
                rewrite big_oracle_re, CPU_ID_re; auto.
              -
                simpl.
                rewrite ZMap.gss.
                rewrite zeq_false; auto.
                case_eq (zeq i0 (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in ikern_re0.
                  rewrite ikern_re0.
                  rewrite <- ikern_re.
                  rewrite Hdestruct1.
                  reflexivity.
              + rewrite zeq_false in ikern_re0; auto.
            -
              simpl.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i0 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in ihost_re0.
                rewrite ihost_re0.
                rewrite <- ihost_re.
                rewrite Hdestruct2.
                reflexivity.
              + rewrite zeq_false in ihost_re0; auto.
            -
              unfold relate_AC_per_pd in *; simpl.
              rewrite Hdestruct in ×.
              unfold B_GetContainerUsed´ in AC_re0.
              rewrite Hdestruct10 in AC_re0.
              assert (B_GetContainerUsed i0 (CPU_ID (update init_shared_adt l)) l0
                      = B_GetContainerUsed i0 (CPU_ID (update init_shared_adt l))
                                           (TBEVENT (CPU_ID (update init_shared_adt l))
                                                    (TBSHARED
                                                       (BTDYIELD
                                                          (ZMap.get (AbstractDataType.CPU_ID a)
                                                                      (AbstractDataType.cid a))))
                                                    :: big_oracle (update init_shared_adt l)
                                                    (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
              { assert (B_GetContainerUsed i0 (CPU_ID (update init_shared_adt l)) l0
                        = B_GetContainerUsed i0 (CPU_ID (update init_shared_adt l))
                                              (big_oracle (update init_shared_adt l)
                                                          (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                { eapply single_big_oracle_query_preserves_cused.
                  inv sh_shared_inv_re; auto. }
                unfold B_GetContainerUsed in H11.
                unfold B_GetContainerUsed.
                simpl; auto. }
              rewrite <- H13; clear H13.
              auto.
            -
              simpl.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i0 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in ti_re0.
                rewrite ti_re0.
                rewrite <- ti_re.
                rewrite Htid_val.
                reflexivity.
              + rewrite zeq_false in ti_re0; auto.
            -
              auto.
            -
              rewrite Hdestruct.
              rewrite Hdestruct in PT_re0.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i0 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in PT_re0.
                rewrite PT_re0.
                rewrite <- PT_re.
                reflexivity.
              + rewrite zeq_false in PT_re0; auto.
            -
              rewrite Hdestruct.
              rewrite Hdestruct in ptp_re0.
              auto.
            -
              simpl.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i0 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in ipt_re0.
                rewrite ipt_re0.
                rewrite <- ipt_re.
                rewrite Hdestruct3.
                reflexivity.
              + rewrite zeq_false in ipt_re0; auto.
            -
              simpl.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i0 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in intr_flag_re0.
                rewrite intr_flag_re0.
                rewrite <- intr_flag_re.
                rewrite Hdestruct5.
                reflexivity.
              + rewrite zeq_false in intr_flag_re0; auto.
            -
              simpl.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i0 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in in_intr_re0.
                rewrite in_intr_re0.
                rewrite <- in_intr_re.
                rewrite Hdestruct4.
                reflexivity.
              + rewrite zeq_false in in_intr_re0; auto.
            -
              unfold relate_SyncChanPool_per_pd.
              simpl.
              rewrite Hdestruct.
              rewrite zeq_true; auto.
            - unfold relate_uctxt_per_pd in *; simpl in ×.
              rewrite Hdestruct in ×.
              unfold B_GetContainerUsed´ in uctxt_re0.
              rewrite Hdestruct10 in uctxt_re0.
              assert (B_GetContainerUsed i0 (CPU_ID (update init_shared_adt l)) l0
                      = B_GetContainerUsed i0 (CPU_ID (update init_shared_adt l))
                                           (TBEVENT (CPU_ID (update init_shared_adt l))
                                                    (TBSHARED
                                                       (BTDYIELD
                                                          (ZMap.get (AbstractDataType.CPU_ID a)
                                                                    (AbstractDataType.cid a))))
                                                    :: big_oracle (update init_shared_adt l)
                                                    (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
              { assert (B_GetContainerUsed i0 (CPU_ID (update init_shared_adt l)) l0
                        = B_GetContainerUsed i0 (CPU_ID (update init_shared_adt l))
                                             (big_oracle (update init_shared_adt l)
                                                         (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                { eapply single_big_oracle_query_preserves_cused.
                  inv sh_shared_inv_re; auto. }
                unfold B_GetContainerUsed in H11.
                unfold B_GetContainerUsed.
                simpl; auto. }
              rewrite <- H13; clear H13.
              auto.
            - auto.
            - auto.
            - auto.
            - constructor; simpl.
              inv inv_re0.
              intros.
              eapply dirty_page_per_thread_inv; auto. }
          × constructor.
        + inv sh_shared_inv_re.
          constructor; simpl; auto.
          × simpl.
            intros.
            rewrite Hdestruct in H4; inv H4.
          × simpl.
            rewrite zeq_true; intros.
            rewrite Hdestruct6 in syncchpool_inv.
            apply syncchpool_inv in H4.
            rewrite <- H4.
            unfold B_GetlastPush; rewrite Hdestruct10.
            symmetry.
            eapply single_big_oracle_query_preserves_B_GetlastPush_aux; eauto.
          × intros.
            simpl.
            generalize (container_used_inv tid H4).
            intros.
            unfold B_GetContainerUsed´ in H11.
            rewrite Hdestruct10 in H11.
            exploit single_big_oracle_query_preserves_cused; eauto.
          × intros.
            simpl.
            apply uctxt_used_inv; auto.
            unfold B_GetContainerUsed´.
            rewrite Hdestruct10.
            intros.
            exploit single_big_oracle_query_preserves_cused; eauto.
          × unfold valid_AT_log_type_B.
            intros.
            inv Hdef.
            destruct big_oracle_inv as (Htemp1 & Htemp2 & Htemp3 & Htemp4).
            unfold valid_AT_oracle_B in Htemp3.
            generalize (valid_AT_log_inv _ Hdestruct10); intros Hvalid_log.
            generalize (Htemp3 l0 (CPU_ID (update init_shared_adt l)) Hvalid_log); intros Hvalid_oracle_log.
            unfold valid_AT_log_B in ×.
            destruct Hvalid_oracle_log as (? & Hvalid_oracle_log).
            unfold B_CalAT_log_real in *; simpl.
            rewrite Hvalid_oracle_log; esplit; auto.
          × unfold valid_TCB_log_type_B.
            intros.
            inv Hdef.
            revert Hdestruct14.
            rewrite Hdestruct in sh_cid_re.
            rewrite sh_cid_re.
            rewrite sh_CPU_ID_re.
            unfold valid_TCB_log_B.
            unfold B_CalTCB_log_real.
            clear.
            intros; subst.
            subdestruct.
            esplit; auto.
          × intros.
            destruct big_oracle_inv as (Htemp1 & Htemp2 & Htemp3 & Htemp4).
            unfold valid_AT_oracle_B in Htemp3.
            unfold valid_AT_log_type_B in valid_AT_log_inv.
            generalize (valid_AT_log_inv l0 Hdestruct10); intros Hvalid_AT_log.
            generalize (Htemp3 _ (CPU_ID (update init_shared_adt l)) Hvalid_AT_log);
              intros Hvalid_AT_oracle_res.
            unfold valid_AT_log_type_B in valid_AT_log_inv.
            unfold valid_AT_log_B in valid_AT_log_inv.
            destruct Hvalid_AT_log as (at0 & Hvalid_AT_log).
            eapply (valid_B_Cal_AT_log_inv i0 l0 at0) in H4; eauto.
            Focus 2.
            rewrite <- sh_big_log_re; auto.
            inv H11.
            unfold B_CalAT_log_real in Hvalid_AT_log, H12.
            simpl in H12.
            case_eq (B_CalAT_log (big_oracle (update init_shared_adt l)
                                             (CPU_ID (update init_shared_adt l)) l0 ++ l0)
                                 (real_AT (ZMap.init ATUndef))); intros.
            rewrite <- sh_CPU_ID_re, <- sh_big_oracle_re in H12.
            rewrite H11 in H12.
            inv H12.
            eapply (B_CalAT_log_keeps_allocated_pages _ _ _ _ _ Hvalid_AT_log H11 i0 H4); eauto.
            rewrite <- sh_CPU_ID_re, <- sh_big_oracle_re in H12.
            rewrite H11 in H12.
            inv H12.
    Fail idtac.
    Qed.
    Next Obligation.
      rename H5 into source_kctxt.
      assert (cid_eq: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))
                      = ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)).
      { generalize H12; intro Hstate_checks.
        unfold state_checks in H12.
        simpl in H12.
        unfold state_check in H12.
        rewrite H6 in H12.
        rewrite peq_false in H12; [ | intro contra; inv contra].
        rewrite peq_true in H12.
        unfold thread_sleep_state_check in H12.
        simpl in H12.
        subdestruct; try inv H12; try contradiction.
        inv H20.
        rewrite Hdestruct in sh_cid_re.
        auto. }
      generalize H12; intro Hstate_checks.
      unfold state_checks in H12.
      simpl in H12.
      unfold state_check in H12.
      rewrite H6 in H12.
      rewrite peq_false in H12; [ | intro contra; inv contra].
      rewrite peq_true in H12.
      unfold thread_sleep_state_check in H12.
      simpl in H12.
      subdestruct; try inv H12; try contradiction.
      rewrite cid_eq in ×.
      unfold sync_chpool_check.
      rewrite peq_true.
      assert (CPU_ID_eq: (CPU_ID (update init_shared_adt l)) = (AbstractDataType.CPU_ID a)).
      { inv H20; rewrite sh_CPU_ID_re; auto. }
      assert (big_log_eq: (big_log (update init_shared_adt l)) = (AbstractDataType.big_log a)).
      { inv H20; rewrite sh_big_log_re; auto. }
      assert (Htid_val: ti d = init_trap_info).
      { assert (fst (ti d) = Int.zero).
        { generalize Hdestruct10.
          clear.
          intros.
          unfold Int.eq in Hdestruct10.
          subdestruct.
          unfold Int.zero; unfold Int.zero in e.
          clear Hdestruct.
          eapply Int_unsigned_eq in e.
          symmetry; auto. }
        simpl in Hdestruct17.
        unfold init_trap_info.
        rewrite <- H5.
        destruct (ti d).
        subst; simpl.
        f_equal.
        simpl in Hdestruct17; rewrite Hdestruct17.
        unfold Vzero.
        simpl in Hdestruct18.
        rewrite <- (Int.eq_true Int.zero) in Hdestruct18.
        unfold Int.eq in Hdestruct18.
        clear H8 H9.
        subdestruct.
        clear Hdestruct16.
        apply Int_unsigned_eq in e.
        rewrite e; auto. }
      
      assert (Hsyncchpool_eq : syncchpool d = AbstractDataType.syncchpool a).
      { inv H20.
        generalize (per_data_re (proc_id (uRData l))).
        simpl.
        rewrite cid_eq.
        rewrite H6.
        intros.
        inv H5.
        unfold relate_SyncChanPool_per_pd in ×.
        rewrite Hdestruct in syncchpool_re.
        rewrite <- Hdestruct12 in Hdestruct6.
        rewrite Hdestruct6 in syncchpool_re.
        simpl in syncchpool_re.
        rewrite cid_eq in syncchpool_re.
        rewrite zeq_true in syncchpool_re.
        assumption. }

      rewrite <- Hdestruct12 in Hdestruct6.
      rewrite <- cid_eq in Hdestruct19.
      simpl in H8, H9.
      unfold sync_chpool_check in H8, H9.
      rewrite peq_true in H8, H9.
      rewrite <- cid_eq in Hdestruct13.
      rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct8, Hdestruct15, Hdestruct9,
      Hdestruct12, Hdestruct14, Hdestruct19, Hdestruct23 in H8, H9.
      rewrite Hdestruct21, Hdestruct22, Hdestruct24 in H8, H9.
      rewrite zeq_true in H8, H9.
      simpl in H8, H9.

      set (rs1 := (Pregmap.init Vundef)#ESP <- (rs#ESP)#EDI <- (rs#EDI)#ESI <- (rs#ESI)
                                       #EBX <- (rs#EBX)#EBP <- (rs#EBP)#RA <- (rs#RA)).
      set (rs´ := ZMap.get run_id (kctxt a)).

      rewrite ZMap.gss in H8.
      rewrite ZMap.gss in H9.

      assert (Hproc_id: Constant.TOTAL_CPU <
                        ZMap.get (CPU_ID (update init_shared_adt l))
                                 (cid (update init_shared_adt l)) < Constant.num_proc
                        ZMap.get (CPU_ID (update init_shared_adt l))
                                 (cid (update init_shared_adt l)) = main_thread).
      { eapply valid_thread_list.
        exploit all_cid_in_full_thread_list.
        simpl.
        instantiate (1 := l); intro; auto. }

      assert (Hrun_id_range: Constant.TOTAL_CPU < run_id < Constant.num_proc run_id = main_thread).
      { assert (Hnew_id:
                  ZMap.get (CPU_ID (update init_shared_adt (LEvent (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))
                                                                   (LogSleep (Int.unsigned i0) (Mem.nextblock m)
                                                                             (Some (syncchpool d))) :: l)))
                           (cid (update init_shared_adt (LEvent (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))
                                                                (LogSleep (Int.unsigned i0) (Mem.nextblock m)
                                                                          (Some (syncchpool d))) :: l))) = run_id).
        { simpl.
          rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct8, Hdestruct15, Hdestruct9,
          Hdestruct12, Hdestruct14, Hdestruct19, Hdestruct23.
          rewrite Hdestruct21, Hdestruct22, Hdestruct24.
          rewrite zeq_true.
          simpl; rewrite ZMap.gss; auto. }
        rewrite <- Hnew_id.
        eapply valid_thread_list.
        exploit all_cid_in_full_thread_list.
        simpl.
        instantiate (1 := (LEvent (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))
                                  (LogSleep (Int.unsigned i0) (Mem.nextblock m) (Some (syncchpool d))) :: l)).
        intros.
        auto. }
      assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
      { rewrite main_thread_val.
        generalize current_CPU_ID_range.
        clear.
        intros.
        omega. }

      assert (Hcurid_cused: cused (ZMap.get (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))
                                            (AbstractDataType.AC a)) = true).
      { inv H20.
        inv sh_shared_inv_re.
        rewrite <- cid_eq.
        assert (In (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                   full_thread_list).
        { generalize all_cid_in_full_thread_list.
          simpl; intros Htemp; apply Htemp. }
        rewrite cid_eq in H5.
        rewrite cid_eq.
        generalize (container_used_inv (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)) H5).
        unfold B_GetContainerUsed´.
        rewrite Hdestruct12.
        intros Htemp ; rewrite <- Htemp.
        rewrite <- cid_eq.
        auto. }
      
      assert (Hdes_cused : cused (ZMap.get run_id (AbstractDataType.AC a)) = true).
      { subst; inv H20.
        assert (proc_id (uRData ((LEvent (proc_id (uRData l))
                                         (LogSleep (Int.unsigned i0) (Mem.nextblock m)
                                                   (sync_chpool_check thread_sleep (Lint i0 :: nil)
                                                                      (uRData l) d)))::l)) = run_id).
        { simpl in ×.
          unfold sync_chpool_check.
          rewrite peq_true.
          rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct8, Hdestruct15, Hdestruct9,
          Hdestruct12, Hdestruct14, Hdestruct19, Hdestruct23.
          rewrite Hdestruct21, Hdestruct22, Hdestruct24.
          rewrite zeq_true.
          simpl.
          rewrite ZMap.gss; auto. }
        inv sh_shared_inv_re.
        assert (In (proc_id (uRData ((LEvent (proc_id (uRData l))
                                             (LogSleep (Int.unsigned i0) (Mem.nextblock m)
                                                       (sync_chpool_check thread_sleep (Lint i0 :: nil)
                                                                          (uRData l) d)))::l)))
                            full_thread_list).
        { apply all_cid_in_full_thread_list. }
        generalize (container_used_inv
                      (proc_id (uRData ((LEvent (proc_id (uRData l))
                                                (LogSleep (Int.unsigned i0) (Mem.nextblock m)
                                                          (sync_chpool_check thread_sleep (Lint i0 :: nil)
                                                                             (uRData l) d)))::l))) H5).
        unfold B_GetContainerUsed´.
        rewrite Hdestruct12.
        intros Htemp ; rewrite <- Htemp.
        auto. }

       (a {big_log: BigDef (TBEVENT (AbstractDataType.CPU_ID a)
                                          (TBSHARED (BTDSLEEP (ZMap.get (AbstractDataType.CPU_ID a)
                                                                        (AbstractDataType.cid a))
                                                              (Int.unsigned i0) (AbstractDataType.syncchpool a)))
                                          :: (AbstractDataType.big_oracle a) (AbstractDataType.CPU_ID a) l0 ++ l0)}
                {kctxt: ZMap.set (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)) rs1 (AbstractDataType.kctxt a)}
                {cid : ZMap.set (AbstractDataType.CPU_ID a) run_id (AbstractDataType.cid a)}
                {lock : ZMap.set (slp_q_id (Int.unsigned i0) 0 + lock_TCB_range) LockFalse
                                 (ZMap.set (slp_q_id (Int.unsigned i0) 0 + ID_AT_range) LockFalse
                                           (AbstractDataType.lock a))}).

      remember (a {big_log: BigDef (TBEVENT (AbstractDataType.CPU_ID a)
                                            (TBSHARED (BTDSLEEP (ZMap.get (AbstractDataType.CPU_ID a)
                                                                          (AbstractDataType.cid a))
                                                                (Int.unsigned i0) (AbstractDataType.syncchpool a)))
                                            :: (AbstractDataType.big_oracle a) (AbstractDataType.CPU_ID a) l0 ++ l0)}
                  {kctxt: ZMap.set (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)) rs1 (AbstractDataType.kctxt a)}
                  {cid : ZMap.set (AbstractDataType.CPU_ID a) run_id (AbstractDataType.cid a)}
                  {lock : ZMap.set (slp_q_id (Int.unsigned i0) 0 + lock_TCB_range) LockFalse
                                   (ZMap.set (slp_q_id (Int.unsigned i0) 0 + ID_AT_range) LockFalse
                                             (AbstractDataType.lock a))}) as .
      assert (Hassumptions: ( v r, ZtoPreg v = Some rVal.has_type (des_rs´#r) AST.Tint)
                            ( v r, ZtoPreg v = Some r
                                                       val_inject (Mem.flat_inj (Mem.nextblock m)) (des_rs´#r) (des_rs´#r))
                            Hthread_trans_sem_hyp rs s m a des_rs´ ZMap.get run_id (kctxt a) = des_rs´
                            (ZMap.get run_id rsm = EAsmCommon.Running des_rs´
                              pc, initial_thread_pc (spl_data := single_data) ge run_id l = Some pc)).
      { assert (Hev_sleep_eq: LEvent (proc_id (uRData l))
                                     (LogSleep (Int.unsigned i0) (Mem.nextblock m)
                                               (sync_chpool_check thread_sleep (Lint i0 :: nil) (uRData l) d)) =
                              LEvent (proc_id (uRData l))
                                     (LogSleep (Int.unsigned i0) (Mem.nextblock m)
                                               (sync_chpool_check thread_sleep (Lint i0 :: nil) (uRData l) d))) by reflexivity.
        assert (Hcommon_cond: Hthread_sched_trans_common_cond ge source_kctxt l
                                                               (LEvent (proc_id (uRData l))
                                                                       (LogSleep (Int.unsigned i0) (Mem.nextblock m)
                                                                                 (sync_chpool_check
                                                                                    thread_sleep
                                                                                    (Lint i0 :: nil) (uRData l) d)))
                                                               dp d a s thread_sleep b sg rsm rs des_rs´).
        { unfold Hthread_sched_trans_common_cond.
          refine_split; auto.
          - intros.
            simpl in H5.
            rewrite cid_eq in H5.
            eapply H0; eauto.
          - simpl; rewrite cid_eq; auto.
          - simpl; rewrite cid_eq; auto.
          - simpl.
            unfold sync_chpool_check.
            rewrite peq_true.
            rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct8, Hdestruct15, Hdestruct9,
            Hdestruct12, Hdestruct14, Hdestruct19, Hdestruct23.
            rewrite Hdestruct21, Hdestruct22, Hdestruct24.
            rewrite zeq_true.
            simpl; rewrite ZMap.gss; auto.
          - simpl.
            unfold sync_chpool_check.
            rewrite peq_true.
            rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct8, Hdestruct15, Hdestruct9,
            Hdestruct12, Hdestruct14, Hdestruct19, Hdestruct23.
            rewrite Hdestruct21, Hdestruct22, Hdestruct24.
            rewrite zeq_true.
            simpl; rewrite ZMap.gss; auto. }
        generalize (Hthread_trans_match_prop ge source_kctxt l dp d a s b sg
                                             rsm rs des_rs´ m
                                             (LEvent (proc_id (uRData l))
                                                     (LogSleep (Int.unsigned i0) (Mem.nextblock m)
                                                               (sync_chpool_check thread_sleep (Lint i0 :: nil) (uRData l) d)))
                                             i0 H4 Hstate_checks Hev_sleep_eq Hcommon_cond).
        simpl.
        unfold sync_chpool_check.
        rewrite peq_true.
        rewrite Hdestruct, Hdestruct0, Hdestruct6, Hdestruct7, Hdestruct8, Hdestruct15, Hdestruct9,
        Hdestruct12, Hdestruct14, Hdestruct19, Hdestruct23.
        rewrite Hdestruct21, Hdestruct22, Hdestruct24.
        rewrite zeq_true.
        simpl.
        rewrite ZMap.gss; auto.
        intros (? & ? & ? & ?).
        repeat split; eauto.
        eapply Hthread_trans_sem_hyp_correct.
      }
      destruct Hassumptions as (N_TYPE & N_INJECT_NEUTRAL & Hthread_trans_sem & Htrans_kctxt & Hinit_exist_aux); subst.

      refine_split´; eauto.
      - replace ((Pos.to_nat (Mem.nextblock m) - Pos.to_nat (Mem.nextblock m))%nat) with O.
        simpl.
        assert ((@mem_lift_nextblock mem memory_model_ops Hmem
                                     (@AlgebraicMemImpl.algebraicmem mem memory_model_ops Hmem) m O) = m).
        { remember (mem_lift_nextblock m 0) as .
          symmetry in Heqm´.
          eapply mem_lift_nextblock_source_eq; eauto. }
        rewrite H5.
        set (rs2 := ((((((Pregmap.init Vundef) # ESP <- (rs ESP)) # EDI <- (rs EDI))
                          # ESI <- (rs ESI)) # EBX <- (rs EBX)) # EBP <- (rs EBP)) # RA <- (rs RA)).
        set ( := (a {big_log: BigDef (TBEVENT (AbstractDataType.CPU_ID a)
                                          (TBSHARED (BTDSLEEP (ZMap.get (AbstractDataType.CPU_ID a)
                                                                        (AbstractDataType.cid a))
                                                              (Int.unsigned i0) (AbstractDataType.syncchpool a)))
                                          :: (AbstractDataType.big_oracle a) (AbstractDataType.CPU_ID a) l0 ++ l0)}
                       {kctxt: ZMap.set (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)) rs1 (AbstractDataType.kctxt a)}
                       {cid : ZMap.set (AbstractDataType.CPU_ID a) run_id (AbstractDataType.cid a)}
                       {lock : ZMap.set (slp_q_id (Int.unsigned i0) 0 + lock_TCB_range) LockFalse
                                        (ZMap.set (slp_q_id (Int.unsigned i0) 0 + ID_AT_range) LockFalse
                                                  (AbstractDataType.lock a))})).
        set (sig := mksignature (AST.Tint::nil) None cc_default).
        generalize (primcall_threadtransfer_sem_intro big_thread_sleep_spec s m rs (ZMap.get run_id (kctxt a)) rs2
                                                       i0 a sig b).
        intros Hsem_intro.
        exploit Hsem_intro; eauto.
        + unfold big_thread_sleep_spec.
          clear H5.
          generalize H20; intro Hrelated_save.
          substx.
          inv H20.
          generalize (per_data_re (ZMap.get (CPU_ID (update init_shared_adt l))
                                  (cid (update init_shared_adt l)))).
          intros priv_data_re.
          rewrite cid_eq in priv_data_re.
          rewrite H6 in priv_data_re.
          inv priv_data_re.
          unfold relate_AC_per_pd in ×.
          unfold proc_id in *; simpl in ×.
          rewrite cid_eq in ×.
          rewrite Hdestruct in ×.
          rewrite zeq_true in ×.
          rewrite Hdestruct11 in ×.
          rewrite <- cid_eq.
          rewrite <- pg_re, <- ikern_re, <- ihost_re, <- ipt_re, <- big_log_re, <- CPU_ID_re, <- lock_re.

          rewrite Hdestruct0, Hdestruct1, Hdestruct2, Hdestruct3, Hdestruct7, Hdestruct8, Hdestruct12, Hdestruct9, Hdestruct15.
          rewrite <- cid_eq in Hcurid_cused.
          rewrite Hcurid_cused.
          rewrite <- big_oracle_re.
          rewrite <- Hsyncchpool_eq.
          rewrite <- cid_eq in Hdestruct19.
          rewrite Hdestruct19.
          rewrite Hdestruct24, Hdestruct14, Hdestruct21.
          rewrite zeq_true.
          repeat f_equal.
        +
          clear.
          remember (Pos.to_nat (Mem.nextblock m)) as val.
          clear.
          induction val.
          simpl; reflexivity.
          simpl; auto.
      - intros; simpl.
        rewrite <- cid_eq in H5, H12.
        rewrite Hdestruct, Hdestruct19 in H5, H12.
        rewrite Hdestruct21, Hdestruct22 in H5, H12.
        simpl in H5, H12.
        unfold match_estate_kctxt; simpl in ×.
        intros.
        inv H12.
        + simpl.
          rewrite ZMap.gso in H14; eauto.
          destruct H14.
          unfold initial_thread_kctxt in H14.
          subdestruct.
          inv H14.
          simpl in Hdestruct11.
          subdestruct; inv Hdestruct11.
          unfold initial_regset_kctxt.
          simpl in H5.
          rewrite <- cid_eq.
          simpl.
          rewrite ZMap.gss in H5; auto.
          case_eq (zeq (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
          × rewrite H7 in H12; inv H12.
          × rewrite ZMap.gso; auto.
            unfold match_estate_kctxt in H0.
            unfold match_estate_kctxt in H0.
            clear H14.
            eapply H0 in n0; eauto.
            unfold match_estate_regset.
            left.
            split; eauto.
            unfold initial_thread_kctxt.
            simpl in Hdestruct16.
            subdestruct.
            inv Hdestruct16.
            unfold initial_regset_kctxt.
            simpl.
            rewrite Hdestruct20.
            auto.
            auto.
            rewrite cid_eq; auto.
        + rewrite ZMap.gso in H14; eauto.
          case_eq (zeq (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
          × rewrite ZMap.gss; auto.
            unfold rs1.
            rewrite H7 in H14.
            substx.
            inv H14.
            inv H13.
            unfold PregtoZ in H15.
            subdestruct; try inv H15; reflexivity.
          × intros; subst.
            rewrite ZMap.gso; eauto.
            clear H12.
            unfold match_estate_regset in H0.
            unfold match_estate_kctxt in H0.
            eapply H0 in n0; eauto.
      - simpl.
        intros.
        unfold thread_init_invariant; intros; simpl.
        unfold initial_thread_kctxt; simpl; intros.
        unfold thread_init_pc_of.
        unfold thread_init_pc.
        simpl in H5.
        unfold thread_init_invariant in H10; simpl in ×.
        rewrite <- cid_eq in H5.
        rewrite Hdestruct, Hdestruct19 in H5.
        rewrite Hdestruct21, Hdestruct22, Hdestruct23 in H5.
        simpl in H5.
        rewrite ZMap.gss in H5.
        case_eq (zeq i1 run_id).
        + intros; subst.
          rewrite ZMap.gss in H5.
          destruct H8.
          × destruct H8.
            unfold initial_thread_kctxt in H13.
            simpl in H13.
            subdestruct.
          × inv H5.
            unfold initial_regset_kctxt.
            assert (Hinit_exist: pc,
                       @initial_thread_pc (@single_data oracle_ops big_ops) mem memory_model_ops
                                          (@SingleOracleImpl.s_config s_oracle_ops s_threads_ops
                                                                      real_params_ops real_params
                                                                      oracle_ops6 oracle_ops big_ops
                                                                      s_threads_prf s_oracle_basic_prop_prf
                                                                      mem memory_model_ops
                                                                      Hmem) fundef unit ge run_id l
                       = Some pc).
            { eapply Hinit_exist_aux; auto. }
            destruct Hinit_exist as (pc´, Hinit_exist).
            simpl in ×.
            rewrite Hinit_exist.
            intro contra; inv contra.
        + intros.
          rewrite ZMap.gso in H5; auto.
          eapply H10 in H5.
          unfold initial_thread_kctxt in H5.
          simpl; intro contra.
          elim H5; subdestruct; reflexivity.
      - simpl.
        instantiate (1:= ZMap.set (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)) rs1 (kctxt a)).
        rewrite <- cid_eq.
        rewrite Hdestruct, Hdestruct19.
        rewrite Hdestruct21, Hdestruct22, Hdestruct23.
        generalize H20; intro Hrelated.
        inv H20.
        constructor; simpl; auto.
        +
          rewrite sh_lock_re.
          auto.
        +
          rewrite <- sh_CPU_ID_re.
          rewrite Hdestruct.
          repeat rewrite ZMap.gss.
          reflexivity.
        +
          rewrite sh_big_oracle_re, sh_CPU_ID_re; auto.
          rewrite Hsyncchpool_eq.
          auto.
        + intros i1.
          generalize (per_data_re (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))).
          rewrite H6.
          intros.
          case_eq (@ZMap.get (option privData) i1 dp); intros.
          × inversion H5.
            unfold relate_AC_per_pd in ×.
            rewrite Hdestruct in ×.
            rewrite Hdestruct11 in ×.
            simpl in ×.
            rewrite cid_eq in ×.
            rewrite zeq_true in ×.
            unfold B_GetContainerUsed´ in AC_re.
            rename AC_re into AC_re´.
            assert (AC_re: AC d = ZMap.get (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))
                                           (AbstractDataType.AC a)).
            { destruct Hproc_id as [Hproc_id | Hproc_id]; auto.
              + assert (Hno_main: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) main_thread).
                { revert Hproc_id.
                  revert Hmain_thread.
                  rewrite cid_eq.
                  clear; intros.
                  omega. }
                rewrite <- cid_eq in AC_re´.
                rewrite zeq_false in AC_re´; try assumption.
                rewrite zlt_lt_true in AC_re´; try (rewrite cid_eq; assumption).
                rewrite cid_eq in AC_re´.
                rewrite Hdestruct12 in AC_re´.
                rewrite Hdestruct13 in AC_re´.
                auto.
              + rewrite Hproc_id in AC_re´.
                rewrite zeq_true in AC_re´; try assumption.
                rewrite Hdestruct12 in AC_re´.
                rewrite <- Hproc_id in AC_re´.
                rewrite Hdestruct13 in AC_re´.
                auto. }
            clear AC_re´.
            case_eq (zeq i1 run_id).
            {
              intros; subst.
              rewrite H9 in H12.
              inv H12.
              generalize (per_data_re run_id); intros.
              rewrite H9 in H12.
              clear H13.
              inv H12.
              rewrite Hdestruct in ×.
              per_data_auto_simpl.
              -
                rewrite Hdestruct; auto.
              -
                simpl.
                rewrite lock_re.
                auto.
              -
                rewrite Hdestruct.
                repeat rewrite ZMap.gss; reflexivity.
              -
                rewrite big_oracle_re, CPU_ID_re; auto.
                rewrite Hsyncchpool_eq.
                auto.
              -
                simpl.
                rewrite ZMap.gss.
                rewrite zeq_true.
                simpl in ikern_re0.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in ikern_re0.
                  rewrite ikern_re0; reflexivity.
                + rewrite zeq_false in ikern_re0; auto.
                  rewrite ikern_re0.
                  rewrite <- ikern_re.
                  rewrite Hdestruct1.
                reflexivity.
              -
                simpl.
                rewrite ZMap.gss.
                rewrite zeq_true.
                simpl in ikern_re0.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in ihost_re0.
                  rewrite ihost_re0; reflexivity.
                + rewrite zeq_false in ihost_re0; auto.
                  rewrite ihost_re0.
                  rewrite <- ihost_re.
                  rewrite Hdestruct2.
                  reflexivity.
              -
                unfold relate_AC_per_pd; simpl.
                rewrite Hdestruct.
                unfold relate_AC_per_pd in AC_re0; rewrite Hdestruct in AC_re0.
                unfold B_GetContainerUsed´ in AC_re0.
                rewrite Hdestruct12 in AC_re0.
                assert (B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l)) l0 =
                        B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l))
                                           (TBEVENT (CPU_ID (update init_shared_adt l))
                                                    (TBSHARED
                                                       (BTDSLEEP
                                                          (ZMap.get (AbstractDataType.CPU_ID a)
                                                                    (AbstractDataType.cid a)) (Int.unsigned i0)
                                                          (syncchpool d)))
                                                    :: big_oracle (update init_shared_adt l)
                                                    (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                { assert (B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l)) l0
                          = B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l))
                                                (big_oracle (update init_shared_adt l)
                                                            (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                  { eapply single_big_oracle_query_preserves_cused.
                    inv sh_shared_inv_re; auto. }
                  unfold B_GetContainerUsed in H12.
                  unfold B_GetContainerUsed.
                  simpl; auto. }
                rewrite <- H12; clear H12.
                auto.
              -
                simpl.
                rewrite ZMap.gss.
                rewrite zeq_true.
                simpl in ti_re0.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in ti_re0.
                  rewrite ti_re0; reflexivity.
                + rewrite zeq_false in ti_re0; auto.
                  rewrite ti_re0.
                  rewrite <- ti_re.
                  rewrite Htid_val; auto.
              -
                auto.
              -
                rewrite Hdestruct.
                rewrite ZMap.gss.
                rewrite zeq_true.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in PT_re0.
                  rewrite PT_re0; reflexivity.
                + rewrite zeq_false in PT_re0; auto.
                  rewrite PT_re0.
                  rewrite <- PT_re.
                  reflexivity.
              -
                rewrite Hdestruct.
                assumption.
              -
                rewrite ZMap.gss.
                rewrite zeq_true.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in ipt_re0.
                  rewrite ipt_re0; reflexivity.
                + rewrite zeq_false in ipt_re0; auto.
                  rewrite ipt_re0.
                  rewrite <- ipt_re.
                  rewrite Hdestruct3.
                  reflexivity.
              -
                rewrite ZMap.gss.
                rewrite zeq_true.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in intr_flag_re0.
                  rewrite intr_flag_re0; reflexivity.
                + rewrite zeq_false in intr_flag_re0; auto.
                  rewrite intr_flag_re0.
                  rewrite <- intr_flag_re.
                  rewrite Hdestruct5.
                  reflexivity.
              -
                rewrite ZMap.gss.
                rewrite zeq_true.
                case_eq (zeq run_id (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in in_intr_re0.
                  rewrite in_intr_re0; reflexivity.
                + rewrite zeq_false in in_intr_re0; auto.
                  rewrite in_intr_re0.
                  rewrite <- in_intr_re.
                  rewrite Hdestruct4.
                  reflexivity.
              - unfold relate_SyncChanPool_per_pd.
                simpl; rewrite Hdestruct; auto.
                rewrite zeq_true.
                auto.
              - unfold relate_uctxt_per_pd in *; simpl in ×.
                rewrite Hdestruct in ×.
                unfold B_GetContainerUsed´ in uctxt_re0.
                rewrite Hdestruct12 in uctxt_re0.
                assert (B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l)) l0
                        = B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l))
                                             (TBEVENT (CPU_ID (update init_shared_adt l))
                                                      (TBSHARED
                                                         (BTDSLEEP
                                                            (ZMap.get (AbstractDataType.CPU_ID a)
                                                                      (AbstractDataType.cid a)) (Int.unsigned i0)
                                                            (syncchpool d)))
                                                      :: big_oracle (update init_shared_adt l)
                                                      (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                { assert (B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l)) l0
                          = B_GetContainerUsed run_id (CPU_ID (update init_shared_adt l))
                                                (big_oracle (update init_shared_adt l)
                                                            (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                  { eapply single_big_oracle_query_preserves_cused.
                    inv sh_shared_inv_re; auto. }
                  unfold B_GetContainerUsed in H12.
                  unfold B_GetContainerUsed.
                  simpl; auto. }
                rewrite <- H12; clear H12.
                auto.
              - auto.
              - auto.
              - auto.
              - constructor; simpl.
                inv inv_re0.
                auto. }
            {
              repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              intros; subst.
              generalize (per_data_re i1).
              rewrite H12.
              intros.
              inv H14.
              per_data_auto_simpl.
              -
                rewrite sh_lock_re.
                auto.
              -
                rewrite Hdestruct.
                rewrite <- sh_CPU_ID_re.
                repeat rewrite ZMap.gss.
                reflexivity.
              -
                rewrite sh_big_oracle_re, sh_CPU_ID_re; auto.
                rewrite Hsyncchpool_eq.
                auto.
              -
                simpl.
                rewrite ZMap.gss.
                rewrite zeq_false; auto.
                case_eq (zeq i1 (proc_id (update init_shared_adt l))); intros; subst.
                + rewrite zeq_true in ikern_re0.
                  rewrite ikern_re0.
                  rewrite <- ikern_re.
                  rewrite Hdestruct1.
                  reflexivity.
              + rewrite zeq_false in ikern_re0; auto.
            -
              simpl.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i1 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in ihost_re0.
                rewrite ihost_re0.
                rewrite <- ihost_re.
                rewrite Hdestruct2.
                reflexivity.
              + rewrite zeq_false in ihost_re0; auto.
            -
              unfold relate_AC_per_pd in *; simpl.
              rewrite Hdestruct in ×.
              unfold B_GetContainerUsed´ in AC_re0.
              rewrite Hdestruct12 in AC_re0.
              assert (B_GetContainerUsed i1 (CPU_ID (update init_shared_adt l)) l0 =
                      B_GetContainerUsed i1 (CPU_ID (update init_shared_adt l))
                                         (TBEVENT (CPU_ID (update init_shared_adt l))
                                                  (TBSHARED
                                                     (BTDSLEEP
                                                        (ZMap.get (AbstractDataType.CPU_ID a)
                                                                  (AbstractDataType.cid a)) (Int.unsigned i0)
                                                        (syncchpool d)))
                                                  :: big_oracle (update init_shared_adt l)
                                                  (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
              { assert (B_GetContainerUsed i1 (CPU_ID (update init_shared_adt l)) l0
                        = B_GetContainerUsed i1 (CPU_ID (update init_shared_adt l))
                                              (big_oracle (update init_shared_adt l)
                                                          (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                { eapply single_big_oracle_query_preserves_cused.
                  inv sh_shared_inv_re; auto. }
                unfold B_GetContainerUsed in H12.
                unfold B_GetContainerUsed.
                simpl; auto. }
              rewrite <- H14; clear H14.
              auto.
            -
              simpl.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i1 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in ti_re0.
                rewrite ti_re0.
                rewrite <- ti_re.
                rewrite Htid_val.
                reflexivity.
              + rewrite zeq_false in ti_re0; auto.
            -
              auto.
            -
              rewrite Hdestruct.
              rewrite Hdestruct in PT_re0.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i1 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in PT_re0.
                rewrite PT_re0.
                rewrite <- PT_re.
                reflexivity.
              + rewrite zeq_false in PT_re0; auto.
            -
              rewrite Hdestruct.
              rewrite Hdestruct in ptp_re0.
              auto.
            -
              simpl.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i1 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in ipt_re0.
                rewrite ipt_re0.
                rewrite <- ipt_re.
                rewrite Hdestruct3.
                reflexivity.
              + rewrite zeq_false in ipt_re0; auto.
            -
              simpl.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i1 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in intr_flag_re0.
                rewrite intr_flag_re0.
                rewrite <- intr_flag_re.
                rewrite Hdestruct5.
                reflexivity.
              + rewrite zeq_false in intr_flag_re0; auto.
            -
              simpl.
              rewrite ZMap.gss.
              rewrite zeq_false; auto.
              case_eq (zeq i1 (proc_id (update init_shared_adt l))); intros; subst.
              + rewrite zeq_true in in_intr_re0.
                rewrite in_intr_re0.
                rewrite <- in_intr_re.
                rewrite Hdestruct4.
                reflexivity.
              + rewrite zeq_false in in_intr_re0; auto.
            -
              unfold relate_SyncChanPool_per_pd.
              simpl.
              rewrite Hdestruct.
              auto.
              rewrite zeq_true; auto.
            - unfold relate_uctxt_per_pd in *; simpl in ×.
              rewrite Hdestruct in ×.
              unfold B_GetContainerUsed´ in uctxt_re0.
              rewrite Hdestruct12 in uctxt_re0.
              assert (B_GetContainerUsed i1 (CPU_ID (update init_shared_adt l)) l0
                      = B_GetContainerUsed i1 (CPU_ID (update init_shared_adt l))
                                           (TBEVENT (CPU_ID (update init_shared_adt l))
                                                    (TBSHARED
                                                       (BTDSLEEP
                                                          (ZMap.get (AbstractDataType.CPU_ID a)
                                                                    (AbstractDataType.cid a)) (Int.unsigned i0)
                                                          (syncchpool d)))
                                                    :: big_oracle (update init_shared_adt l)
                                                    (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
              { assert (B_GetContainerUsed i1 (CPU_ID (update init_shared_adt l)) l0
                        = B_GetContainerUsed i1 (CPU_ID (update init_shared_adt l))
                                             (big_oracle (update init_shared_adt l)
                                                         (CPU_ID (update init_shared_adt l)) l0 ++ l0)).
                { eapply single_big_oracle_query_preserves_cused.
                  inv sh_shared_inv_re; auto. }
                unfold B_GetContainerUsed in H12.
                unfold B_GetContainerUsed.
                simpl; auto. }
              rewrite <- H14; clear H14.
              auto.
            - auto.
            - auto.
            - auto.
            - constructor; simpl.
              inv inv_re0.
              intros.
              eapply dirty_page_per_thread_inv; auto. }
          × constructor.
        + inv sh_shared_inv_re.
          constructor; simpl; auto.
          × simpl.
            rewrite Hdestruct; intro contra; inv contra.
          × simpl.
            intros.
            rewrite zeq_true.
            auto.
          × intros.
            simpl.
            generalize (container_used_inv tid).
            intros.
            unfold B_GetContainerUsed´ in H12.
            rewrite Hdestruct12 in H12.
            apply H12 in H5.
            exploit single_big_oracle_query_preserves_cused; eauto.
          × intros.
            apply uctxt_used_inv; auto.
            unfold B_GetContainerUsed´.
            rewrite Hdestruct12.
            rewrite <- H5.
            exploit single_big_oracle_query_preserves_cused; eauto.
          × unfold valid_AT_log_type_B.
            intros.
            inv Hdef.
            destruct big_oracle_inv as (Htemp1 & Htemp2 & Htemp3 & Htemp4).
            unfold valid_AT_oracle_B in Htemp3.
            generalize (valid_AT_log_inv _ Hdestruct12); intros Hvalid_log.
            generalize (Htemp3 l0 (CPU_ID (update init_shared_adt l)) Hvalid_log); intros Hvalid_oracle_log.
            unfold valid_AT_log_B in ×.
            destruct Hvalid_oracle_log as (? & Hvalid_oracle_log).
            unfold B_CalAT_log_real in *; simpl.
            rewrite Hvalid_oracle_log; esplit; auto.
          × unfold valid_TCB_log_type_B.
            intros.
            inv Hdef.
            revert Hdestruct19.
            rewrite Hdestruct in sh_cid_re.
            rewrite sh_cid_re.
            rewrite sh_CPU_ID_re.
            unfold valid_TCB_log_B.
            unfold B_CalTCB_log_real.
            clear.
            intros; subst.
            subdestruct.
            esplit; auto.
          × intros.
            destruct big_oracle_inv as (Htemp1 & Htemp2 & Htemp3 & Htemp4).
            unfold valid_AT_oracle_B in Htemp3.
            unfold valid_AT_log_type_B in valid_AT_log_inv.
            generalize (valid_AT_log_inv l0 Hdestruct12); intros Hvalid_AT_log.
            generalize (Htemp3 _ (CPU_ID (update init_shared_adt l)) Hvalid_AT_log);
              intros Hvalid_AT_oracle_res.
            unfold valid_AT_log_type_B in valid_AT_log_inv.
            unfold valid_AT_log_B in valid_AT_log_inv.
            destruct Hvalid_AT_log as (at0 & Hvalid_AT_log).
            eapply (valid_B_Cal_AT_log_inv i1 l0 at0) in H5; eauto.
            Focus 2.
            rewrite <- sh_big_log_re; auto.
            inv H12.
            unfold B_CalAT_log_real in Hvalid_AT_log, H13.
            simpl in H13.
            case_eq (B_CalAT_log (big_oracle (update init_shared_adt l)
                                             (CPU_ID (update init_shared_adt l)) l0 ++ l0)
                                 (real_AT (ZMap.init ATUndef))); intros.
            rewrite <- sh_CPU_ID_re, <- sh_big_oracle_re in H13.
            rewrite H12 in H13.
            inv H13.
            eapply (B_CalAT_log_keeps_allocated_pages _ _ _ _ _ Hvalid_AT_log H12 i1 H5); eauto.
            rewrite <- sh_CPU_ID_re, <- sh_big_oracle_re in H13.
            rewrite H12 in H13.
            inv H13.
    Fail idtac.
    Qed.
    Opaque PDX Zdivide_dec zle Z.mul Z.add Z.sub Z.div.

    Opaque compatlayer_ops.
    Opaque update.
    Opaque uRData.
    Opaque proc_id.

    Next Obligation.
      rename H1 into kctxt_pool.
      rename H3 into Hrel.
      destruct (ef_cases ef) as [Hcase|Hcase].

      - destruct ef; inv Hcase.
        inv H; inv H1.
        destruct H as [Hprims Hsem].
        inv_layer Hprims; has_event_remove BUILTIN_ENABLED.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          inv_monad H1; inv H1.
          Transparent proc_id update uRData.
          refine_split´; auto.
          × econstructor; [simpl | simpl; eauto].
             (gensem vmxinfo_get_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor.
            unfold vmxinfo_get_spec, ObjPHBThreadVMM.single_vmxinfo_get_spec in ×.
            simpl in *; subdestruct.
            inv Hrel.
            generalize (per_data_re (proc_id (uRData l))); intro Hrel_pd.
            simpl in Hrel_pd.
            rewrite H2 in Hrel_pd.
            inv Hrel_pd; simpl in *; rewrite zeq_true in ×.
            rewrite <- ikern_re, <- ihost_re.
            rewrite <- sh_vmxinfo_re.
            rewrite Hdestruct0, Hdestruct.
            inv H5; reflexivity.
          × instantiate (1:= kctxt_pool).
            apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          inv_monad H1.
          functional inversion H4; inv H3.
           (a {PT : Int.unsigned i}).
          Transparent proc_id update uRData.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem setPT_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold setPT_spec in ×.
            inv Hrel.
            generalize (per_data_re (proc_id (uRData l))); intros; simpl in ×.
            rewrite H2 in H1; inv H1; rewrite zeq_true in ×.
            rewrite <- ikern_re, <- ihost_re, <- ipt_re.
            rewrite H5, H6, H7, H9; reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in ×.
            generalize Hrel; intros Hrelate.
            inv Hrel; simpl in *; constructor; auto; intros.
            generalize (per_data_re i0).
            intros.
            case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))));
              intros; subst.
            { rewrite H2 in H1.
              rewrite ZMap.gss.
              inv H1.
              per_data_auto.
              - auto.
              - auto.
              - rewrite H8 in PT_re; rewrite H8.
                rewrite zeq_true in PT_re; rewrite zeq_true; auto.
              - auto.
              - auto.
              - constructor.
                inv inv_re; simpl.
                intros.
                eauto.
            }
            { rewrite ZMap.gso; auto.
              subdestruct; simpl; auto.
              inv H1.
              per_data_auto.
              - auto.
              - auto.
              - rewrite Hdestruct.
                rewrite Hdestruct in PT_re.
                rewrite zeq_false in PT_re; auto.
                rewrite zeq_false; auto; auto.
              - auto.
              - auto.
              - constructor.
                inv inv_re.
                intros; auto.
              - constructor. }
            { inv sh_shared_inv_re; constructor; auto. }
            { inv sh_mem_inv_re.
              constructor; simpl; intros.
              case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l))
                                        (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H3.
                rewrite ZMap.gso in H10; eauto.
                inv H3; simpl in H11.
                eauto.
              - rewrite ZMap.gso in H3; auto.
                case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                         (cid (update init_shared_adt l)))); intros; subst.
                + rewrite ZMap.gss in H10; eauto.
                  inv H10.
                  simpl.
                  eapply pperm_disjoint; eauto.
                + rewrite ZMap.gso in H10; eauto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          inv H8.
          inv_monad H1; inv H1.
           a.
          Transparent proc_id update uRData.
          refine_split; eauto; [ | eapply relate_RData_gss; eauto ].
          econstructor; [simpl | simpl; eauto].
           (gensem ptRead_spec); split; [reflexivity |].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor; simpl in ×.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel.
          revert H5; intros Hrelated.
          eapply ptRead_related in Hrelated; eauto.
          destruct Hrelated as (res´ & Hrelated_a & Hrelated_b).
          rewrite Hrelated_a; subst; auto.
          Opaque proc_id update uRData.


        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          inv H8; inv_monad H1; inv H1.
          unfold single_uctx_get_spec in H5; subdestruct; inv H5.
          Transparent proc_id update uRData.
          refine_split´; auto.
          × econstructor; [simpl | simpl; eauto].
             (gensem uctx_get_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor.
            unfold uctx_get_spec in *; simpl in *; subdestruct.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in Hper_data_rel; rewrite H2 in Hper_data_rel.
            inv Hper_data_rel.
            simpl in *; rewrite zeq_true in ×.
            rewritesb.
            rewrite e.
            rewrite Hdestruct1, Hdestruct2, Hdestruct3, Hdestruct4, Hdestruct5, Hdestruct6.
            rewrite <- e.
            unfold relate_uctxt_per_pd in uctxt_re.
            unfold B_GetContainerUsed´ in uctxt_re.
            rewrite Hdestruct0, Hdestruct10, Hdestruct11 in uctxt_re.
            assert (Hproc_id: Constant.TOTAL_CPU <
                         ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)) < Constant.num_proc
                              ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)) = main_thread).
            { eapply valid_thread_list.
              exploit all_cid_in_full_thread_list.
              simpl.
              instantiate (1 := l); intro; auto. }
            
            assert (ZMap.get (Int.unsigned i0) (uctxt ) =
                    (ZMap.get (Int.unsigned i0) (ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l))
                                                                    (cid (update init_shared_adt l)))
                                                          (AbstractDataType.uctxt a)))).
            { assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
              { rewrite main_thread_val.
                generalize current_CPU_ID_range.
                clear.
                intros.
                omega. }
              destruct Hproc_id as [Hproc_id | Hproc_id]; auto.
              + assert (Hno_main: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) main_thread).
                { revert Hproc_id.
                  revert Hmain_thread.
                  clear; intros.
                  omega. }
                rewrite zeq_false in uctxt_re; try assumption.
                rewrite zlt_lt_true in uctxt_re; try assumption.
                unfold uctxt_bieq in uctxt_re.
                destruct (ZMap.get 12 (uctxt )) eqn:?; try rewrite uctxt_re; eauto; try contradiction.

                unfold is_UCTXT_ptr in Hdestruct6.
                destruct uctxt_re as (uctxt_re & ?).
                generalize (uctxt_re (Int.unsigned i0)); intros.
                repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                       end.

                exploit H4; auto.
                subdestruct; try omega.
                subdestruct; try omega.

                unfold is_UCTXT_ptr in Hdestruct6.
                case_eq (Int.eq i2 (Int.repr 0)); intros; subst; rewrite H1 in uctxt_re.
                × repeat match goal with
                         | [H: context[if _ then _ else _] |- _] ⇒ clear H
                         end.
                  destruct uctxt_re as (uctxt_re & _).
                  eapply uctxt_re.
                  subdestruct; try omega.
                  subdestruct; try omega.
                × rewrite uctxt_re.
                  congruence.

              + rewrite Hproc_id in uctxt_re.
                rewrite zeq_true in uctxt_re; try assumption.
                unfold uctxt_bieq in uctxt_re.
                destruct (ZMap.get 12 (uctxt )) eqn:?; try rewrite uctxt_re; try rewrite Hproc_id; eauto; try contradiction.
                unfold is_UCTXT_ptr in Hdestruct6.
                destruct uctxt_re as (uctxt_re & ?).
                generalize (uctxt_re (Int.unsigned i0)); intros.
                repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                       end.
                exploit H4; auto.
                subdestruct; try omega.
                subdestruct; try omega.

                unfold is_UCTXT_ptr in Hdestruct6.
                case_eq (Int.eq i2 (Int.repr 0)); intros; subst; rewrite H1 in uctxt_re.
                × repeat match goal with
                         | [H: context[if _ then _ else _] |- _] ⇒ clear H
                         end.
                  destruct uctxt_re as (uctxt_re & _).
                  eapply uctxt_re.
                  subdestruct; try omega.
                  subdestruct; try omega.
                × rewrite uctxt_re.
                  congruence. }
            rewrite <- H1.
            rewrite Hdestruct12; auto.
          × instantiate (1:= kctxt_pool).
            apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          destruct args; inv H8.
          inv H9.
          unfold single_uctx_set_spec in H1; subdestruct; inv H1.
          Transparent proc_id update uRData.
          assert (Hproc_id: Constant.TOTAL_CPU <
                            ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)) < Constant.num_proc
                            ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)) = main_thread).
          { eapply valid_thread_list.
            exploit all_cid_in_full_thread_list.
            simpl.
            instantiate (1 := l); intro; auto. }
           (a {uctxt
                     : ZMap.set (Int.unsigned i)
                                (ZMap.set (Int.unsigned i0) (Vint (Int.repr (Int.unsigned i1)))
                                          (ZMap.get (Int.unsigned i) (AbstractDataType.uctxt a)))
                                (AbstractDataType.uctxt a)}).
          refine_split´; auto.
          × econstructor; [simpl | simpl; eauto].
             (gensem uctx_set_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor.
            unfold uctx_set_spec in *; simpl in *; subdestruct.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in Hper_data_rel; rewrite H2 in Hper_data_rel.
            inv Hper_data_rel.
            simpl in *; rewrite zeq_true in ×.
            rewritesb.
            rewrite e.
            rewrite Hdestruct1, Hdestruct2, Hdestruct3, Hdestruct4, Hdestruct5, Hdestruct6.
            rewrite <- e.
            reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in ×.
            generalize Hrel; intros Hrelate.
            inv Hrel; simpl in *; constructor; auto; intros.
            generalize (per_data_re i2).
            intros.
            case_eq (zeq i2 (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))));
              intros; subst.
            rewrite H2 in H1.
            { rewrite ZMap.gss.
              inv H1.
              per_data_auto.
              - auto.
              - auto.
              - rewrite zeq_true in PT_re; rewrite zeq_true; auto.
              - auto.
              - unfold relate_uctxt_per_pd in ×.
                simpl.
                unfold B_GetContainerUsed´ in ×.
                rewrite Hdestruct0, Hdestruct10, Hdestruct11 in ×.
                case_eq (zeq (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l))) main_thread); intros; subst.
                + rewrite H1 in uctxt_re.
                  repeat match goal with
                         | [H: context[if _ then _ else _] |- _] ⇒ clear H
                         end.
                  unfold uctxt_bieq in uctxt_re.
                  unfold uctxt_bieq.
                  destruct (ZMap.get 12 (uctxt d)) eqn:?; try contradiction.
                  × rewrite ZMap.gso.
                    rewrite Heqv.
                    unfold is_UCTXT_ptr in Hdestruct6.
                    destruct uctxt_re as (uctxt_re & ?).
                    generalize (uctxt_re (Int.unsigned i0)); intros.
                    repeat match goal with
                           | [H: context[if _ then _ else _] |- _] ⇒ clear H
                           end.
                    exploit H5; auto.
                    subdestruct; try omega.
                    subdestruct; try omega.

                    intros.
                    split.
                    { intros.
                      generalize (uctxt_re i2 H7 H8).
                      case_eq (zeq i2 (Int.unsigned i0)); intros; subst.
                      { rewrite ZMap.gss.
                        rewrite <- e.
                        rewrite ZMap.gss.
                        rewrite ZMap.gss.
                        auto. }
                      { rewrite ZMap.gso.
                        rewrite <- e.
                        rewrite ZMap.gss.
                        rewrite ZMap.gso.
                        auto.

                        assumption.
                        assumption. } }
                    { destruct H4 as (ueip_val & H4).
                       ueip_val.
                      rewrite <- e.
                      rewrite ZMap.gss.
                      rewrite ZMap.gso.
                      auto.
                      subdestruct; try omega. }

                    unfold is_UCTXT_ptr in Hdestruct6.
                    subdestruct; try omega.
                  × rewrite ZMap.gso.
                    rewrite Heqv.
                    case_eq (Int.eq i2 (Int.repr 0)); intros; auto.
                    { rewrite H4 in uctxt_re.
                      destruct uctxt_re as (uctxt_re & ?).
                      unfold is_UCTXT_ptr in Hdestruct6.
                      generalize (uctxt_re (Int.unsigned i0)); intros.
                      repeat match goal with
                             | [H: context[if _ then _ else _] |- _] ⇒ clear H
                             end.
                      exploit H6; auto.
                      subdestruct; try omega.
                      subdestruct; try omega.
                      intros.
                      split.
                      { intros.
                        generalize (uctxt_re i3 H8 H9).
                        case_eq (zeq i3 (Int.unsigned i0)); intros; subst.
                        { rewrite ZMap.gss.
                          rewrite <- e.
                          rewrite ZMap.gss.
                          rewrite ZMap.gss.
                          auto. }
                        { rewrite ZMap.gso.
                          rewrite <- e.
                          rewrite ZMap.gss.
                          rewrite ZMap.gso.
                          auto.

                          assumption.
                          assumption. } }
                      { destruct H5 as (ueip_val & H5).
                         ueip_val.
                        rewrite <- e.
                        rewrite ZMap.gss.
                        rewrite ZMap.gso.
                        auto.
                        subdestruct; try omega. } }
                    { rewrite <- e.
                      rewrite ZMap.gss.
                      rewrite H4 in uctxt_re.
                      rewrite uctxt_re.
                      reflexivity. }
                    
                    unfold is_UCTXT_ptr in Hdestruct6.
                    subdestruct.

                + rewrite H1 in uctxt_re.
                  case_eq (zlt_lt 8 (ZMap.get (CPU_ID (update init_shared_adt l))
                                              (cid (update init_shared_adt l))) 1024); intros; subst; try congruence.

                  rewrite H4 in uctxt_re.
                  repeat match goal with
                         | [H: context[if _ then _ else _] |- _] ⇒ clear H
                         end.
                  unfold uctxt_bieq in uctxt_re.
                  unfold uctxt_bieq.
                  destruct (ZMap.get 12 (uctxt d)) eqn:?; try contradiction.
                  × rewrite ZMap.gso.
                    rewrite Heqv.
                    unfold is_UCTXT_ptr in Hdestruct6.
                    destruct uctxt_re as (uctxt_re & ?).
                    generalize (uctxt_re (Int.unsigned i0)); intros.
                    repeat match goal with
                           | [H: context[if _ then _ else _] |- _] ⇒ clear H
                           end.
                    exploit H6; auto.
                    subdestruct; try omega.
                    subdestruct; try omega.

                    intros.
                    split.
                    { intros.
                      generalize (uctxt_re i2 H8 H9).
                      case_eq (zeq i2 (Int.unsigned i0)); intros; subst.
                      { rewrite ZMap.gss.
                        rewrite <- e.
                        rewrite ZMap.gss.
                        rewrite ZMap.gss.
                        auto. }
                      { rewrite ZMap.gso.
                        rewrite <- e.
                        rewrite ZMap.gss.
                        rewrite ZMap.gso.
                        auto.

                        assumption.
                        assumption. } }
                    { destruct H5 as (ueip_val & H5).
                       ueip_val.
                      rewrite <- e.
                      rewrite ZMap.gss.
                      rewrite ZMap.gso.
                      auto.
                      subdestruct; try omega. }

                    unfold is_UCTXT_ptr in Hdestruct6.
                    subdestruct; try omega.
                  × rewrite ZMap.gso.
                    rewrite Heqv.
                    case_eq (Int.eq i2 (Int.repr 0)); intros; auto.
                    { rewrite H5 in uctxt_re.
                      destruct uctxt_re as (uctxt_re & ?).
                      unfold is_UCTXT_ptr in Hdestruct6.
                      generalize (uctxt_re (Int.unsigned i0)); intros.
                      repeat match goal with
                             | [H: context[if _ then _ else _] |- _] ⇒ clear H
                             end.
                      exploit H7; auto.
                      subdestruct; try omega.
                      subdestruct; try omega.
                      intros.
                      split.
                      { intros.
                        generalize (uctxt_re i3 H9 H10).
                        case_eq (zeq i3 (Int.unsigned i0)); intros; subst.
                        { rewrite ZMap.gss.
                          rewrite <- e.
                          rewrite ZMap.gss.
                          rewrite ZMap.gss.
                          auto. }
                        { rewrite ZMap.gso.
                          rewrite <- e.
                          rewrite ZMap.gss.
                          rewrite ZMap.gso.
                          auto.

                          assumption.
                          assumption. } }
                      { destruct H6 as (ueip_val & H6).
                         ueip_val.
                        rewrite <- e.
                        rewrite ZMap.gss.
                        rewrite ZMap.gso.
                        auto.
                        subdestruct; try omega. } }
                    { rewrite <- e.
                      rewrite ZMap.gss.
                      rewrite H5 in uctxt_re.
                      rewrite uctxt_re.
                      reflexivity. }
                    
                    unfold is_UCTXT_ptr in Hdestruct6.
                    subdestruct.
              - auto.
              - inv inv_re; constructor; auto.
            }
            { rewrite ZMap.gso; auto.
              subdestruct; simpl; auto.
              inv H1.
              per_data_auto.
              - auto.
              - auto.
              - rewrite zeq_false in PT_re; auto.
                rewrite zeq_false; auto; auto.
              - auto.
              - unfold relate_uctxt_per_pd.
                simpl.
                rewrite ZMap.gso; auto.
                rewrite <- e.
                assumption.
              - auto.
              - constructor.
                inv inv_re.
                intros; auto.
              - constructor. }
            { inv sh_shared_inv_re.
              constructor; simpl; auto.
              rewrite Hdestruct0.
              intros contra; inv contra.
              - intros.
                case_eq (zeq tid (ZMap.get (CPU_ID (update init_shared_adt l))
                                           (cid (update init_shared_adt l)))).
                intros; subst.
                unfold B_GetContainerUsed´ in H1.
                rewrite Hdestruct10 in H1.
                rewrite Hdestruct11 in H1.
                inv H1.

                intros.
                rewrite <- e.
                rewrite ZMap.gso; auto. }
            { inv sh_mem_inv_re.
              constructor; simpl; intros.
              case_eq (zeq i2 (ZMap.get (CPU_ID (update init_shared_adt l))
                                        (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H3.
                rewrite ZMap.gso in H4; eauto.
                inv H3; simpl in H4.
                eauto.
              - rewrite ZMap.gso in H3; auto.
                case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                         (cid (update init_shared_adt l)))); intros; subst.
                + rewrite ZMap.gss in H4; eauto.
                  inv H4.
                  simpl.
                  eapply pperm_disjoint; eauto.
                + rewrite ZMap.gso in H4; eauto. }
            Opaque proc_id update uRData.

        +
          Transparent proc_id update uRData.
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9; inv H1.
          unfold ObjPHBThreadVMM.single_container_get_nchildren_spec in H4; subdestruct; inv H4.
          assert (Hproc_id: Constant.TOTAL_CPU <
                            ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)) < Constant.num_proc
                            ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)) = main_thread).
          { eapply valid_thread_list.
            exploit all_cid_in_full_thread_list.
            simpl.
            instantiate (1 := l); intro; auto. }
          assert (AC_eq: AC = ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                          (AbstractDataType.AC a)).
          { generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel.
            unfold relate_AC_per_pd in AC_re.
            unfold B_GetContainerUsed´ in AC_re.
            rewrite Hdestruct0, Hdestruct6, Hdestruct7 in AC_re.
            assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
            { rewrite main_thread_val.
              generalize current_CPU_ID_range.
              clear.
              intros.
              omega. }
            destruct Hproc_id as [Hproc_id | Hproc_id]; auto.
            + assert (Hno_main: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) main_thread).
              { revert Hproc_id.
                revert Hmain_thread.
                clear; intros.
                omega. }
              rewrite zeq_false in AC_re; try assumption.
              rewrite zlt_lt_true in AC_re; try assumption.
            + rewrite Hproc_id in AC_re.
              rewrite zeq_true in AC_re; try assumption.
              rewrite Hproc_id; auto. }
          assert (Hcused: cused (ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                          (AbstractDataType.AC a)) = true).
          { inv Hrel.
            inv sh_shared_inv_re.
            rewrite <- Hdestruct7.
            unfold B_GetContainerUsed´ in container_used_inv.
            rewrite Hdestruct6 in container_used_inv.
            exploit all_cid_in_full_thread_list.
            instantiate (1:= l); intros; auto.
            symmetry.
            apply container_used_inv; auto. }

          refine_split´; auto.
          × econstructor; [simpl | simpl; eauto].
             (gensem container_get_nchildren_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor.
            unfold container_get_nchildren_spec in *; simpl in *; subdestruct.
            rewrite e.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in Hper_data_rel; rewrite H2 in Hper_data_rel.
            inv Hper_data_rel; simpl in *; rewrite zeq_true in ×.

            rewrite <- ikern_re, <- ihost_re.
            rewrite Hdestruct1, Hdestruct2, Hcused.
            rewrite <- H6; simpl.
            rewrite AC_eq; reflexivity.
          × simpl.
            instantiate (1:= kctxt_pool).
            apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9; inv H1.
          unfold ObjPHBThreadVMM.single_container_get_quota_spec in H4; subdestruct; inv H4.
          Transparent proc_id update uRData.
          assert (Hproc_id: Constant.TOTAL_CPU <
                            ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)) < Constant.num_proc
                            ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)) = main_thread).
          { eapply valid_thread_list.
            exploit all_cid_in_full_thread_list.
            simpl.
            instantiate (1 := l); intro; auto. }
          assert (AC_eq: AC = ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                          (AbstractDataType.AC a)).
          { generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel.
            unfold relate_AC_per_pd in AC_re.
            unfold B_GetContainerUsed´ in AC_re.
            rewrite Hdestruct0, Hdestruct6, Hdestruct7 in AC_re.
            assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
            { rewrite main_thread_val.
              generalize current_CPU_ID_range.
              clear.
              intros.
              omega. }
            destruct Hproc_id as [Hproc_id | Hproc_id]; auto.
            + assert (Hno_main: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) main_thread).
              { revert Hproc_id.
                revert Hmain_thread.
                clear; intros.
                omega. }
              rewrite zeq_false in AC_re; try assumption.
              rewrite zlt_lt_true in AC_re; try assumption.
            + rewrite Hproc_id in AC_re.
              rewrite zeq_true in AC_re; try assumption.
              rewrite Hproc_id; auto. }
          assert (Hcused: cused (ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                          (AbstractDataType.AC a)) = true).
          { inv Hrel.
            inv sh_shared_inv_re.
            rewrite <- Hdestruct7.
            unfold B_GetContainerUsed´ in container_used_inv; simpl in ×.
            rewrite Hdestruct6 in container_used_inv.
            exploit all_cid_in_full_thread_list.
            instantiate (1:= l); intros; auto.
            symmetry.
            apply container_used_inv; auto. }

          refine_split´; auto.
          × econstructor; [simpl | simpl; eauto].
             (gensem container_get_quota_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor.
            unfold container_get_quota_spec in *; simpl in *; subdestruct.

            rewrite e.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in Hper_data_rel; rewrite H2 in Hper_data_rel.
            inv Hper_data_rel; simpl in *; rewrite zeq_true in ×.

            rewrite <- ikern_re, <- ihost_re.
            rewrite Hdestruct1, Hdestruct2, Hcused.
            rewrite <- H6; simpl.
            rewrite AC_eq; reflexivity.
          × simpl.
            instantiate (1:= kctxt_pool).
            apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9; inv H1.
          unfold ObjPHBThreadVMM.single_container_get_usage_spec in H4; subdestruct; inv H4.
          Transparent proc_id update uRData.
          assert (Hproc_id: Constant.TOTAL_CPU <
                            ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)) < Constant.num_proc
                            ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)) = main_thread).
          { eapply valid_thread_list.
            exploit all_cid_in_full_thread_list.
            simpl.
            instantiate (1 := l); intro; auto. }
          assert (AC_eq: AC = ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                          (AbstractDataType.AC a)).
          { generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel.
            unfold relate_AC_per_pd in AC_re.
            unfold B_GetContainerUsed´ in AC_re.
            rewrite Hdestruct0, Hdestruct6, Hdestruct7 in AC_re.
            assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
            { rewrite main_thread_val.
              generalize current_CPU_ID_range.
              clear.
              intros.
              omega. }
            destruct Hproc_id as [Hproc_id | Hproc_id]; auto.
            + assert (Hno_main: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) main_thread).
              { revert Hproc_id.
                revert Hmain_thread.
                clear; intros.
                omega. }
              rewrite zeq_false in AC_re; try assumption.
              rewrite zlt_lt_true in AC_re; try assumption.
            + rewrite Hproc_id in AC_re.
              rewrite zeq_true in AC_re; try assumption.
              rewrite Hproc_id; auto. }
          assert (Hcused: cused (ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                          (AbstractDataType.AC a)) = true).
          { inv Hrel.
            inv sh_shared_inv_re.
            rewrite <- Hdestruct7.
            unfold B_GetContainerUsed´ in container_used_inv; simpl in ×.
            rewrite Hdestruct6 in container_used_inv.
            exploit all_cid_in_full_thread_list.
            instantiate (1:= l); intros; auto.
            symmetry.
            apply container_used_inv; auto. }

          refine_split´; auto.
          × econstructor; [simpl | simpl; eauto].
             (gensem container_get_usage_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor.
            unfold container_get_usage_spec in *; simpl in *; subdestruct.
            rewrite e.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in Hper_data_rel; rewrite H2 in Hper_data_rel.
            inv Hper_data_rel; simpl in *; rewrite zeq_true in ×.

            rewrite <- ikern_re, <- ihost_re.
            rewrite Hdestruct1, Hdestruct2, Hcused.
            rewrite <- H6; simpl.
            rewrite AC_eq; reflexivity.
          × simpl.
            instantiate (1:= kctxt_pool).
            apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          inv H8; inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in ×.
          Transparent proc_id update uRData.
          simpl in ×.
          rewrite H2 in Hper_data_rel; inv Hper_data_rel; simpl in *; rewrite zeq_true in ×.
          assert (Hproc_id: Constant.TOTAL_CPU <
                            ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)) < Constant.num_proc
                            ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)) = main_thread).
          { eapply valid_thread_list.
            exploit all_cid_in_full_thread_list.
            simpl.
            instantiate (1 := l); intro; auto. }
          functional inversion H5; subst; simpl.
          × assert (AC_eq: AC = ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                            (AbstractDataType.AC a)).
            { generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
              simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel.
              unfold relate_AC_per_pd in AC_re.
              unfold B_GetContainerUsed´ in AC_re; simpl in ×.
              unfold curid, cpu in H9; rewrite H7, H8, H9 in AC_re.
              assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
              { rewrite main_thread_val.
                generalize current_CPU_ID_range.
                clear.
                intros.
                omega. }
              destruct Hproc_id as [Hproc_id | Hproc_id]; auto.
              + assert (Hno_main: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) main_thread).
                { revert Hproc_id.
                  revert Hmain_thread.
                  clear; intros.
                  omega. }
                rewrite zeq_false in AC_re; try assumption.
                rewrite zlt_lt_true in AC_re; try assumption.
              + rewrite Hproc_id in AC_re.
                rewrite zeq_true in AC_re; try assumption.
                rewrite Hproc_id; auto. }
            assert (Hcused: cused (ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                            (AbstractDataType.AC a)) = true).
            { inv Hrel.
              inv sh_shared_inv_re.
              rewrite <- H9.
              unfold B_GetContainerUsed´ in container_used_inv; simpl in ×.
              rewrite H8 in container_used_inv.
              exploit all_cid_in_full_thread_list.
              instantiate (1:= l); intros; auto.
              symmetry.
              unfold curid, cpu; auto. }
            refine_split´.
            { econstructor; [simpl | simpl; eauto].
               (gensem container_can_consume_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor.
              rewrite _x.
              unfold container_can_consume_spec; simpl in ×.
              rewrite Hcused.
              rewrite <- ikern_re, <- ihost_re.
              rewrite H4, H6.
              rewrite <- AC_eq.
              unfold c in H10.
              rewrite zle_le_true; auto.
              rewrite <- H1; reflexivity. }
            { reflexivity. }
            { instantiate (1:= kctxt_pool).
              apply relate_RData_gss; auto. }
            { auto. }
          × assert (AC_eq: AC = ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                            (AbstractDataType.AC a)).
            { generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
              simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel.
              unfold relate_AC_per_pd in AC_re.
              unfold B_GetContainerUsed´ in AC_re; simpl in ×.
              unfold curid, cpu in H9; rewrite H7, H8, H9 in AC_re.
              assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
              { rewrite main_thread_val.
                generalize current_CPU_ID_range.
                clear.
                intros.
                omega. }
              destruct Hproc_id as [Hproc_id | Hproc_id]; auto.
              + assert (Hno_main: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) main_thread).
                { revert Hproc_id.
                  revert Hmain_thread.
                  clear; intros.
                  omega. }
                rewrite zeq_false in AC_re; try assumption.
                rewrite zlt_lt_true in AC_re; try assumption.
              + rewrite Hproc_id in AC_re.
                rewrite zeq_true in AC_re; try assumption.
                rewrite Hproc_id; auto. }
            assert (Hcused: cused (ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                            (AbstractDataType.AC a)) = true).
            { inv Hrel.
              inv sh_shared_inv_re.
              rewrite <- H9.
              unfold B_GetContainerUsed´ in container_used_inv; simpl in ×.
              rewrite H8 in container_used_inv.
              exploit all_cid_in_full_thread_list.
              instantiate (1:= l); intros; auto.
              symmetry.
              unfold curid, cpu; auto. }
            refine_split´.
            { econstructor; [simpl | simpl; eauto].
               (gensem container_can_consume_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor.
              rewrite _x.
              unfold container_can_consume_spec; simpl in ×.
              rewrite Hcused.
              rewrite <- ikern_re, <- ihost_re.
              rewrite H4, H6.
              rewrite <- AC_eq.
              unfold c in H10.
              rewrite H10.
              rewrite <- H1; reflexivity. }
            { reflexivity. }
            { instantiate (1:= kctxt_pool).
              apply relate_RData_gss; auto. }
            { auto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv_monad H1; inv H1.
          unfold ObjPHBThreadVMM.single_get_CPU_ID_spec in H5; subdestruct; inv H5.
          Transparent proc_id update uRData.
          refine_split´; auto.
          × econstructor; [simpl | simpl; eauto].
             (gensem get_CPU_ID_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor.
            unfold get_CPU_ID_spec in *; simpl in *; subdestruct.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; simpl in *; rewrite zeq_true in ×.
            rewrite <- ikern_re, <- ihost_re, <- CPU_ID_re.
            rewrite Hdestruct0, Hdestruct1, H3.
            reflexivity.
          × simpl.
            instantiate (1:= kctxt_pool).
            apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv_monad H1; inv H1.
          unfold ObjPHBThreadVMM.single_get_curid_spec in H5; subdestruct; inv H5.
          Transparent proc_id update uRData.
          refine_split´; auto.
          × econstructor; [simpl | simpl; eauto].
             (gensem get_curid_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor.
            unfold get_curid_spec in *; simpl in *; subdestruct.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; simpl in *; rewrite zeq_true in ×.
            rewrite Hdestruct in cid_re.
            rewrite <- ikern_re, <- ihost_re, <- cid_re.
            rewrite Hdestruct1, Hdestruct0, H3.
            reflexivity.
          × simpl.
            instantiate (1:= kctxt_pool).
            apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          inv_monad H1; inv H1.
          Transparent proc_id update uRData.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem get_sync_chan_busy_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold get_sync_chan_busy_spec in ×.
            rewrite <- ikern_re, <- pg_re, <- ihost_re, <- ipt_re.
            unfold relate_SyncChanPool_per_pd in syncchpool_re.
            rewrite H9 in syncchpool_re.
            rewrite H3 in syncchpool_re.
            simpl in syncchpool_re; rewrite zeq_true in syncchpool_re.
            rewrite <- syncchpool_re.
            rewrite H4, H6, H7, H8, H10, H11, H1; reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in *; apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          destruct args; inv H1.
          inv H4.
          Transparent proc_id update uRData.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H1; inv H4.
          subst.
           (a {syncchpool : ZMap.set (Int.unsigned i)
                                            (SyncChanValid t v c (Int.repr (Int.unsigned i0)))
                                            (AbstractDataType.syncchpool a)}).
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem set_sync_chan_busy_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold set_sync_chan_busy_spec in ×.
            rewrite <- ikern_re, <- pg_re, <- ihost_re, <- ipt_re.
            unfold relate_SyncChanPool_per_pd in syncchpool_re.
            rewrite H5 in syncchpool_re.
            rewrite H10 in syncchpool_re.
            simpl in syncchpool_re; rewrite zeq_true in syncchpool_re.
            rewrite <- syncchpool_re.
            rewrite H6, H7, H8, H9, H11, H12; reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in ×.
            generalize Hrel; intros Hrelate.
            rewrite H10 in cid_re.
            inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
            constructor; simpl; auto.
            { rewrite init_re.
              rewrite H10.
              auto. }
            intros.
            generalize (per_data_re i1); intros.
            case_eq (zeq i1 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
            { rewrite ZMap.gss.
              rewrite <- cid_re in H3, H2.
              rewrite <- cid_re.
              rewrite H2 in H3.
              inv H3; simpl in *; rewrite zeq_true in ×.
              per_data_auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_AC_per_pd in *; simpl in ×.
                auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite init_re.
                rewrite cid_re; auto.
                rewrite H10; rewrite zeq_true; auto.
                rewrite H10 in PT_re.
                auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_SyncChanPool_per_pd.
                rewrite init_re; rewrite H10, H5.
                simpl; rewrite zeq_true.
                unfold relate_SyncChanPool_per_pd in syncchpool_re.
                rewrite init_re in syncchpool_re.
                rewrite H5 in syncchpool_re.
                rewrite H10 in syncchpool_re.
                simpl in syncchpool_re.
                rewrite cid_re in syncchpool_re.
                rewrite zeq_true in syncchpool_re.
                rewrite <- syncchpool_re; reflexivity.
              - auto.
              - inv inv_re0; constructor; auto. }
            { rewrite ZMap.gso; auto.
              clear BUILTIN_ENABLED.
              repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              subdestruct; auto.
              inv H3; simpl in ×.
              per_data_auto.
              - rewrite <- init_re in H10.
                unfold relate_AC_per_pd in *; simpl in ×.
                rewrite H10.
                rewrite H10 in AC_re0.
                auto.
              - auto.
              - rewrite <- init_re in H10.
                rewrite H10; rewrite H10 in PT_re.
                rewrite cid_re; rewrite cid_re in PT_re.
                rewrite zeq_false; auto; try (rewrite zeq_false in PT_re; auto).
              - unfold relate_SyncChanPool_per_pd in syncchpool_re0.
                unfold relate_SyncChanPool_per_pd.
                rewrite init_re, H5 in ×.
                rewrite zeq_false in syncchpool_re0;
                  try rewrite zeq_false; simpl; try (rewrite cid_re); auto.
              - auto.
              - inv inv_re0; constructor; auto.
              - constructor. }
            { inv sh_shared_inv_re.
              constructor; auto.
              rewrite H5.
              intros.
              inv H3. }
            { inv sh_mem_inv_re.
              rewrite <- cid_re.
              rewrite <- cid_re in H2.
              constructor; simpl; intros.
              case_eq (zeq i1 (ZMap.get (CPU_ID (update init_shared_adt l))
                                        (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H4.
                rewrite ZMap.gso in H13; eauto.
                inv H4; simpl in H14.
                eapply pperm_disjoint; eauto.
              - rewrite ZMap.gso in H4; auto.
                case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                         (cid (update init_shared_adt l)))); intros; subst.
                + rewrite ZMap.gss in H13; eauto.
                  inv H13.
                  simpl.
                  eapply pperm_disjoint; eauto.
                + rewrite ZMap.gso in H13; eauto. }
            Opaque proc_id update uRData.


        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          destruct args; inv H8.
          inv H9.
          Transparent proc_id update uRData.
          exploit ipc_receive_body_related; eauto.
          intros Hrelated.
          destruct Hrelated as ( & res´ & Hrelated_a & Hrelated_b & Hrelated_c & Hrelated_d & Hrelated_e & Hrelated_f); subst.
          esplit; refine_split; eauto.
          econstructor; [simpl | simpl; eauto].
           (gensem big_ipc_receive_body_spec); split; [reflexivity |].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor; simpl in ×.
          eauto.
          Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          inv_monad H1; inv H1.
          Transparent proc_id update uRData.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
            { (gensem rdmsr_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold rdmsr_spec in ×.
              rewrite <- ikern_re, <- ihost_re.
              rewrite H3, H4, H1.
              reflexivity. }
            { simpl; auto. }
          × instantiate (1:= kctxt_pool).
            simpl in *; apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          destruct args; inv H1.
          × inv H4.
            inv_monad H1; inv H1.
            Transparent proc_id update uRData.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
            functional inversion H5.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem wrmsr_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold wrmsr_spec in ×.
              rewrite <- ikern_re, <- ihost_re.
              rewrite H3, H4, H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
          × inv H4.
            inv_monad H1; inv H1.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
            functional inversion H6.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
              rewrite <- H5.
               (gensem wrmsr_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold wrmsr_spec in ×.
              rewrite <- ikern_re, <- ihost_re.
              rewrite H3, H4, H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          Transparent proc_id update uRData.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H1; simpl in *; subst.
          × (a {vmcs : ZMap.set (CPU_ID (update init_shared_adt l))
                                       (VMCSValid revid abrtid
                                                  (ZMap.set 16386
                                                            (Vint (Int.repr (Z.lor (Int.unsigned ctrls) 4))) d0))
                                       (AbstractDataType.vmcs a)}).
            refine_split´; eauto.
            rewrite _x in *; rewrite zeq_true in *; auto.
            unfold pv_adt in *; subst.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_set_intercept_intwin_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_set_intercept_intwin_spec in ×.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H6, H7, H8, H10, H11, H12.
              rewrite _x1; rewrite zeq_true.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in ×.
              generalize Hrel; intros Hrelate.
              {
                rewrite H9 in cid_re.
                inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
                constructor; simpl; auto.
                { rewrite init_re.
                  rewrite H9; auto. }
                intros.
                generalize (per_data_re i0); intros.
                case_eq (zeq i0 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
                { rewrite ZMap.gss.
                  rewrite <- cid_re in H3, H2.
                  rewrite <- cid_re.
                  rewrite H2 in H3.
                  inv H3; simpl in *; rewrite zeq_true in ×.
                  per_data_auto_simpl.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - unfold relate_AC_per_pd in *; simpl in ×.
                  auto.
                - rewrite zeq_true; auto.
                - auto.
                - rewrite init_re.
                  rewrite cid_re; auto.
                  rewrite zeq_true; auto.
                - unfold pv_adt.
                  rewrite init_re.
                  rewrite cid_re; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - unfold relate_SyncChanPool_per_pd.
                  unfold relate_SyncChanPool_per_pd in syncchpool_re.
                  simpl; auto.
                - auto.
                - unfold relate_uctxt_per_pd in *; simpl; auto.
                - rewrite CPU_ID_re.
                  rewrite ZMap.gss.
                  unfold d´0.
                  unfold res.
                  rewrite <- CPU_ID_re.
                  rewrite cid_re; rewrite <- _x.
                  rewrite cid_re; rewrite zeq_true.
                  reflexivity.
                - auto.
                - inv inv_re0; constructor; simpl; auto. }
            { rewrite ZMap.gso; auto.
              clear BUILTIN_ENABLED.
              repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              subdestruct; auto.
              inv H3; simpl in ×.
              per_data_auto_simpl.
              - unfold relate_AC_per_pd in *; simpl in *; auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
              - unfold relate_uctxt_per_pd in *; simpl; auto.
              - auto.
              - rewrite zeq_false; auto.
                rewrite <- _x; auto.
                rewrite cid_re; auto.
              - auto.
              - inv inv_re0; constructor; simpl; auto.
              - constructor. }
            { inv sh_shared_inv_re; constructor; simpl; auto. }
            { inv sh_mem_inv_re.
              rewrite <- cid_re in H2.
              rewrite <- cid_re.
              constructor; simpl; intros.
              case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l))
                                        (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H4.
                rewrite ZMap.gso in H14; eauto.
                inv H4; simpl in H15.
                eauto.
              - rewrite ZMap.gso in H4; auto.
                case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                         (cid (update init_shared_adt l)))); intros; subst.
                + rewrite ZMap.gss in H14; eauto.
                  inv H14.
                  simpl.
                  eapply pperm_disjoint; eauto.
                + rewrite ZMap.gso in H14; eauto. } } }
            
          × (a {vmcs: ZMap.set (CPU_ID (update init_shared_adt l))
                                      (VMCSValid revid abrtid
                                                 (ZMap.set 16386 (Vint (Int.repr (Z.land (Int.unsigned ctrls) 4294967291)))
                                                           d0)) (AbstractDataType.vmcs a)}).
            refine_split´; eauto.
            rewrite _x in *; rewrite zeq_true in *; auto.
            unfold pv_adt in *; subst.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_set_intercept_intwin_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_set_intercept_intwin_spec in ×.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H6, H7, H8, H10, H11, H12.
              rewrite zeq_false; auto. }
            { instantiate (1:= kctxt_pool).
              simpl in ×.
              generalize Hrel; intros Hrelate.
              rewrite H9 in cid_re.
              inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
              constructor; simpl; auto.
              { rewrite init_re.
                rewrite H9; auto. }
              intros.
              generalize (per_data_re i0); intros.
              case_eq (zeq i0 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
              { rewrite ZMap.gss.
                rewrite <- cid_re in H3, H2.
                rewrite <- cid_re.
                rewrite H2 in H3.
                inv H3; simpl in *; rewrite zeq_true in ×.
                per_data_auto_simpl.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - unfold relate_AC_per_pd in *; simpl in ×.
                  auto.
                - rewrite zeq_true; auto.
                - auto.
                - rewrite init_re.
                  rewrite cid_re; auto.
                  rewrite zeq_true; auto.
                - unfold pv_adt.
                  rewrite init_re.
                  rewrite cid_re; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - unfold relate_SyncChanPool_per_pd.
                  unfold relate_SyncChanPool_per_pd in syncchpool_re.
                  simpl; auto.
                - unfold relate_uctxt_per_pd in *; simpl; auto.
                - auto.
                - rewrite CPU_ID_re.
                  rewrite ZMap.gss.
                  unfold d´0.
                  unfold res.
                  rewrite <- CPU_ID_re.
                  rewrite cid_re; rewrite <- _x.
                  rewrite cid_re; rewrite zeq_true.
                  reflexivity.
                - auto.
                - inv inv_re0; constructor; simpl; auto. }
              { rewrite ZMap.gso; auto.
                clear BUILTIN_ENABLED.
                repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                       end.
                subdestruct; auto.
                inv H3; simpl in ×.
                per_data_auto_simpl.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
                - auto.
                - unfold relate_uctxt_per_pd in *; simpl; auto.
                - rewrite zeq_false; auto.
                  rewrite <- _x; auto.
                  rewrite cid_re; auto.
                - auto.
                - inv inv_re0; constructor; simpl; auto.
                - constructor. }
              { inv sh_shared_inv_re; constructor; simpl; auto. }
              { inv sh_mem_inv_re.
                rewrite <- cid_re in H2.
                rewrite <- cid_re.
                constructor; simpl; intros.
                case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l))
                                          (cid (update init_shared_adt l)))); intros; subst.
                - rewrite ZMap.gss in H4.
                  rewrite ZMap.gso in H14; eauto.
                  inv H4; simpl in H14.
                  eauto.
                - rewrite ZMap.gso in H4; auto.
                  case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                           (cid (update init_shared_adt l)))); intros; subst.
                  + rewrite ZMap.gss in H14; eauto.
                    inv H14.
                    simpl.
                    eapply pperm_disjoint; eauto.
                  + rewrite ZMap.gso in H14; eauto. } }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          destruct args; inv H8.
          inv H9.
          Transparent proc_id update uRData.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H1; simpl in *; subst.
           (a {vmcs : ZMap.set (CPU_ID (update init_shared_adt l))
                                     (VMCSValid revid abrtid d´0) (AbstractDataType.vmcs a)}).
          refine_split´; eauto.
          rewrite _x in *; rewrite zeq_true in *; auto.
          unfold pv_adt in *; subst.
          × econstructor; [simpl | simpl; eauto].
             (gensem vmx_set_desc1_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold vmx_set_desc1_spec in ×.
            rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
            rewrite H6, H7, H8, H10, H11, H12.
            reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in ×.
            generalize Hrel; intros Hrelate.
            rewrite H9 in cid_re.
            inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
            constructor; simpl; auto.
            { rewrite init_re; rewrite H9; auto. }
            intros.
            generalize (per_data_re i2); intros.
            case_eq (zeq i2 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
            { rewrite ZMap.gss.
              rewrite <- cid_re in H3, H2.
              rewrite <- cid_re.
              rewrite H2 in H3.
              inv H3; simpl in *; rewrite zeq_true in ×.
              per_data_auto_simpl.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_AC_per_pd in *; simpl in *; auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite init_re.
                rewrite cid_re; auto.
                rewrite zeq_true; auto.
              - unfold pv_adt.
                rewrite init_re.
                rewrite cid_re; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_SyncChanPool_per_pd.
                unfold relate_SyncChanPool_per_pd in syncchpool_re.
                simpl; auto.
              - unfold relate_uctxt_per_pd in *; simpl; auto.
              - auto.
              - rewrite CPU_ID_re.
                rewrite ZMap.gss.
                rewrite <- CPU_ID_re.
                rewrite cid_re; rewrite <- _x.
                rewrite cid_re; rewrite zeq_true.
                reflexivity.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
            { rewrite ZMap.gso; auto.
              clear BUILTIN_ENABLED.
              repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              subdestruct; auto.
              inv H3; simpl in ×.
              per_data_auto_simpl.
              - unfold relate_AC_per_pd in *; simpl in *; auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
              - unfold relate_uctxt_per_pd in *; simpl; auto.
              - auto.
              - rewrite zeq_false; auto.
                rewrite <- _x; auto.
                rewrite cid_re; auto.
              - auto.
              - inv inv_re0; constructor; simpl; auto.
              - constructor. }
            { inv sh_shared_inv_re; constructor; simpl; auto. }
            { inv sh_mem_inv_re.
              rewrite <- cid_re in H2.
              rewrite <- cid_re.
              constructor; simpl; intros.
              case_eq (zeq i2 (ZMap.get (CPU_ID (update init_shared_adt l))
                                        (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H4.
                rewrite ZMap.gso in H13; eauto.
                inv H4; simpl in H14.
                eauto.
              - rewrite ZMap.gso in H4; auto.
                case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                         (cid (update init_shared_adt l)))); intros; subst.
                + rewrite ZMap.gss in H13; eauto.
                  inv H13.
                  simpl.
                  eapply pperm_disjoint; eauto.
                + rewrite ZMap.gso in H13; eauto. }
              Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          destruct args; inv H8.
          inv H9.
          Transparent proc_id update uRData.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H1; simpl in *; subst.
           (a {vmcs : ZMap.set (CPU_ID (update init_shared_adt l))
                                     (VMCSValid revid abrtid d´0) (AbstractDataType.vmcs a)}).
          refine_split´; eauto.
          rewrite _x in *; rewrite zeq_true in *; auto.
          unfold pv_adt in *; subst.
          × econstructor; [simpl | simpl; eauto].
             (gensem vmx_set_desc2_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold vmx_set_desc2_spec in ×.
            rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
            rewrite H6, H7, H8, H10, H11, H12.
            reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in ×.
            generalize Hrel; intros Hrelate.
            rewrite H9 in cid_re.
            inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
            constructor; simpl; auto.
            { rewrite init_re; rewrite H9; auto. }
            intros.
            generalize (per_data_re i2); intros.
            case_eq (zeq i2 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
            { rewrite ZMap.gss.
              rewrite <- cid_re in H3, H2.
              rewrite <- cid_re.
              rewrite H2 in H3.
              inv H3; simpl in *; rewrite zeq_true in ×.
              per_data_auto_simpl.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_AC_per_pd in *; simpl in *; auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite init_re.
                rewrite cid_re; auto.
                rewrite zeq_true; auto.
              - unfold pv_adt.
                rewrite init_re.
                rewrite cid_re; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_SyncChanPool_per_pd.
                unfold relate_SyncChanPool_per_pd in syncchpool_re.
                simpl; auto.
              - unfold relate_uctxt_per_pd in *; simpl; auto.
              - auto.
              - rewrite CPU_ID_re.
                rewrite ZMap.gss.
                rewrite <- CPU_ID_re.
                rewrite cid_re; rewrite <- _x.
                rewrite cid_re; rewrite zeq_true.
                reflexivity.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
            { rewrite ZMap.gso; auto.
              clear BUILTIN_ENABLED.
              repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              subdestruct; auto.
              inv H3; simpl in ×.
              per_data_auto_simpl.
              - unfold relate_AC_per_pd in *; simpl in *; auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
              - unfold relate_uctxt_per_pd in *; simpl; auto.
              - auto.
              - rewrite zeq_false; auto.
                rewrite <- _x; auto.
                rewrite cid_re; auto.
              - auto.
              - inv inv_re0; constructor; simpl; auto.
              - constructor. }
            { inv sh_shared_inv_re; constructor; simpl; auto. }
            { inv sh_mem_inv_re.
              rewrite <- cid_re in H2.
              rewrite <- cid_re.
              constructor; simpl; intros.
              case_eq (zeq i2 (ZMap.get (CPU_ID (update init_shared_adt l))
                                        (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H4.
                rewrite ZMap.gso in H13; eauto.
                inv H4; simpl in H14.
                eauto.
              - rewrite ZMap.gso in H4; auto.
                case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                         (cid (update init_shared_adt l)))); intros; subst.
                + rewrite ZMap.gss in H13; eauto.
                  inv H13.
                  simpl.
                  eapply pperm_disjoint; eauto.
                + rewrite ZMap.gso in H13; eauto. }
              Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          destruct args; inv H8.
          destruct args; inv H9.
          inv H8.
          Transparent proc_id update uRData.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H1; simpl in *; subst.
          × (a {vmcs : ZMap.set (CPU_ID (update init_shared_adt l))
                                       (VMCSValid revid abrtid
                                                  (ZMap.set 16408 (Vint (Int.repr (Int.unsigned i1)))
                                                            (ZMap.set 16406
                                                                      (Vint (Int.repr
                                                                               (Z.lor (Z.lor (Z.lor 2147483648 (Int.unsigned i))
                                                                                             (Int.unsigned i0)) 2048))) d0)))
                                       (AbstractDataType.vmcs a)}).
            refine_split´; eauto.
            rewrite _x in *; rewrite zeq_true in *; auto.
            unfold pv_adt in *; subst.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_inject_event_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_inject_event_spec in ×.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H6, H7, H8, H10, H11, H12.
              unfold i3 in H13.
              rewrite H13, H14.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in ×.
              generalize Hrel; intros Hrelate.
              rewrite H9 in cid_re.
              inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
              constructor; simpl; auto.
              { rewrite init_re; rewrite H9; auto. }
              intros.
              generalize (per_data_re i7); intros.
              case_eq (zeq i7 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
              { rewrite ZMap.gss.
                rewrite <- cid_re in H3, H2.
                rewrite <- cid_re.
                rewrite H2 in H3.
                inv H3; simpl in *; rewrite zeq_true in ×.
                per_data_auto_simpl.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - rewrite zeq_true; auto.
                - auto.
                - rewrite init_re.
                  rewrite cid_re; auto.
                  rewrite zeq_true; auto.
                - unfold pv_adt.
                  rewrite init_re.
                  rewrite cid_re; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - unfold relate_SyncChanPool_per_pd.
                  unfold relate_SyncChanPool_per_pd in syncchpool_re.
                  simpl; auto.
                - unfold relate_uctxt_per_pd in *; simpl; auto.
                - auto.
                - rewrite CPU_ID_re.
                  rewrite ZMap.gss.
                  rewrite <- CPU_ID_re.
                  rewrite cid_re; rewrite <- _x.
                  rewrite cid_re; rewrite zeq_true.
                  unfold d2; unfold d1.
                  unfold i6; unfold i5; unfold i4; unfold i3.
                  reflexivity.
                - auto.
                - inv inv_re; constructor; simpl; auto. }
              { rewrite ZMap.gso; auto.
                clear BUILTIN_ENABLED.
                repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                       end.
                subdestruct; auto.
                inv H3; simpl in ×.
                per_data_auto_simpl.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
                - unfold relate_uctxt_per_pd in *; simpl; auto.
                - auto.
                - rewrite zeq_false; auto.
                  rewrite <- _x; auto.
                  rewrite cid_re; auto.
                - auto.
                - inv inv_re0; constructor; simpl; auto.
                - constructor. }
              { inv sh_shared_inv_re; constructor; simpl; auto. }
              { inv sh_mem_inv_re.
                rewrite <- cid_re in H2.
                rewrite <- cid_re.
                constructor; simpl; intros.
                case_eq (zeq i7 (ZMap.get (CPU_ID (update init_shared_adt l))
                                          (cid (update init_shared_adt l)))); intros; subst.
                - rewrite ZMap.gss in H4.
                  rewrite ZMap.gso in H15; eauto.
                  inv H4; simpl in H16.
                  eauto.
                - rewrite ZMap.gso in H4; auto.
                  case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                           (cid (update init_shared_adt l)))); intros; subst.
                  + rewrite ZMap.gss in H15; eauto.
                    inv H15.
                    simpl.
                    eapply pperm_disjoint; eauto.
                  + rewrite ZMap.gso in H15; eauto. } }

          × (a {vmcs: ZMap.set (CPU_ID (update init_shared_adt l))
                                      (VMCSValid revid abrtid
                                                 (ZMap.set 16406 (Vint (Int.repr (Z.lor (Z.lor 2147483648 (Int.unsigned i))
                                                                                        (Int.unsigned i0)))) d0))
                                      (AbstractDataType.vmcs a)}).
            refine_split´; eauto.
            rewrite _x in *; rewrite zeq_true in *; auto.
            unfold pv_adt in *; subst.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_inject_event_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_inject_event_spec in ×.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H6, H7, H8, H10, H11, H12.
              unfold i3 in H13.
              rewrite H13, H14.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in ×.
              generalize Hrel; intros Hrelate.
              rewrite H9 in cid_re.
              inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
              constructor; simpl; auto.
              { rewrite init_re; rewrite H9; auto. }
              intros.
              generalize (per_data_re i6); intros.
              case_eq (zeq i6 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
              { rewrite ZMap.gss.
                rewrite <- cid_re in H3, H2.
                rewrite <- cid_re.
                rewrite H2 in H3.
                inv H3; simpl in *; rewrite zeq_true in ×.
                per_data_auto_simpl.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - rewrite zeq_true; auto.
                - auto.
                - rewrite init_re.
                  rewrite cid_re; auto.
                  rewrite zeq_true; auto.
                - unfold pv_adt.
                  rewrite init_re.
                  rewrite cid_re; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - unfold relate_SyncChanPool_per_pd.
                  unfold relate_SyncChanPool_per_pd in syncchpool_re.
                  simpl; auto.
                - unfold relate_uctxt_per_pd in *; simpl; auto.
                - auto.
                - rewrite CPU_ID_re.
                  rewrite ZMap.gss.
                  rewrite <- CPU_ID_re.
                  rewrite cid_re; rewrite <- _x.
                  rewrite cid_re; rewrite zeq_true.
                  unfold d1.
                  unfold i5; unfold i4; unfold i3.
                  reflexivity.
                - auto.
                - inv inv_re; constructor; simpl; auto. }
              { rewrite ZMap.gso; auto.
                clear BUILTIN_ENABLED.
                repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                       end.
                subdestruct; auto.
                inv H3; simpl in ×.
                per_data_auto_simpl.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
                - unfold relate_uctxt_per_pd in *; simpl; auto.
                - auto.
                - rewrite zeq_false; auto.
                  rewrite <- _x; auto.
                  rewrite cid_re; auto.
                - auto.
                - inv inv_re0; constructor; simpl; auto.
                - constructor. }
              { inv sh_shared_inv_re; constructor; simpl; auto. }
              { inv sh_mem_inv_re.
                rewrite <- cid_re in H2.
                rewrite <- cid_re.
                constructor; simpl; intros.
                case_eq (zeq i6 (ZMap.get (CPU_ID (update init_shared_adt l))
                                          (cid (update init_shared_adt l)))); intros; subst.
                - rewrite ZMap.gss in H4.
                  rewrite ZMap.gso in H15; eauto.
                  inv H4; simpl in H16.
                  eauto.
                - rewrite ZMap.gso in H4; auto.
                  case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                           (cid (update init_shared_adt l)))); intros; subst.
                  + rewrite ZMap.gss in H15; eauto.
                    inv H15.
                    simpl.
                    eapply pperm_disjoint; eauto.
                  + rewrite ZMap.gso in H15; eauto. } }
          × unfold pv_adt in *; subst.
            unfold sh_adt0 in *; subst.
            unfold sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in *; auto.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_inject_event_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_inject_event_spec in ×.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H6, H7, H8, H10, H11, H12.
              unfold is_invalid in _x1; unfold i3 in _x1.
              rewrite zeq_false; auto. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H1.
          × inv H4.
            Transparent proc_id update uRData.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
            functional inversion H1.
            unfold sh_adt in ×.
            unfold pv_adt in ×.
            simpl in *; rewrite _x in *; rewrite zeq_true in *; auto.
             (a {vmcs: ZMap.set (CPU_ID (update init_shared_adt l))
                                       (VMCSValid revid abrtid
                                                  (CalRealIntelModule.write64_aux 8208 (VZ64 (Int64.unsigned i))
                                                                                  d0)) (AbstractDataType.vmcs a)}).
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_set_tsc_offset_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_set_tsc_offset_spec in ×.
              rewrite <- ikern_re, <- pg_re, <- ihost_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H6, H7, H8, H10, H11.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in ×.
              generalize Hrel; intros Hrelate.
              rewrite H9 in cid_re.
              inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
              constructor; simpl; auto.
              { rewrite init_re; rewrite H9; auto. }
              intros.
              generalize (per_data_re i0); intros.
              case_eq (zeq i0 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
              { rewrite ZMap.gss.
                rewrite <- cid_re in H3, H2.
                rewrite <- cid_re.
                rewrite H2 in H3.
                inv H3; simpl in *; rewrite zeq_true in ×.
                per_data_auto.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - rewrite init_re.
                  rewrite H9 in sh_cid_re.
                  rewrite sh_cid_re.
                  rewrite cid_re; auto.
                  rewrite zeq_true.
                  auto.
                - unfold relate_SyncChanPool_per_pd.
                  unfold relate_SyncChanPool_per_pd in syncchpool_re.
                  simpl; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                  rewrite CPU_ID_re.
                  rewrite ZMap.gss; auto.
                - rewrite zeq_true; auto.
                - inv inv_re; constructor; auto. }
              { rewrite ZMap.gso; auto.
                clear BUILTIN_ENABLED.
                rewrite H9 in sh_cid_re.
                repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                       end.
                subdestruct; auto.
                inv H3; simpl in ×.
                per_data_auto.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - auto.
                - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
                - rewrite zeq_false; auto.
                  rewrite sh_cid_re in _x.
                  rewrite cid_re; auto.
                - auto.
                - inv inv_re0; constructor; simpl; auto.
                - constructor. }
              { inv sh_shared_inv_re; constructor; simpl; auto. }
              { inv sh_mem_inv_re.
                rewrite <- cid_re in H2.
                rewrite <- cid_re.
                constructor; simpl; intros.
                case_eq (zeq i0 vm_handling_cid); intros; subst.
                - rewrite ZMap.gss in H4.
                  rewrite ZMap.gso in H12; eauto.
                  inv H4; simpl in H12.
                  eauto.
                - rewrite ZMap.gso in H4; auto.
                  case_eq (zeq j vm_handling_cid); intros; subst.
                  + rewrite ZMap.gss in H12; eauto.
                    inv H12.
                    simpl.
                    eapply pperm_disjoint; eauto.
                  + rewrite ZMap.gso in H12; eauto. } }

          × inv H4.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
            functional inversion H1.
            unfold sh_adt in ×.
            unfold pv_adt in ×.
            simpl in *; rewrite _x in *; rewrite zeq_true in *; auto.
             (a {vmcs: ZMap.set (CPU_ID (update init_shared_adt l))
                                       (VMCSValid revid abrtid
                                                  (CalRealIntelModule.write64_aux 8208 (VZ64 (Int64.unsigned i))
                                                                                  d0)) (AbstractDataType.vmcs a)}).
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
              rewrite <- H5.
               (gensem vmx_set_tsc_offset_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_set_tsc_offset_spec in ×.
              rewrite <- ikern_re, <- pg_re, <- ihost_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H7, H8, H9, H11, H12.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in ×.
              generalize Hrel; intros Hrelate.
              rewrite H10 in cid_re.
              inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
              constructor; simpl; auto.
              { rewrite init_re; rewrite H10; auto. }
              intros.
              generalize (per_data_re i0); intros.
              case_eq (zeq i0 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
              { rewrite ZMap.gss.
                rewrite <- cid_re in H3, H2.
                rewrite <- cid_re.
                rewrite H2 in H3.
                inv H3; simpl in *; rewrite zeq_true in ×.
                per_data_auto.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - rewrite init_re.
                  rewrite H10 in sh_cid_re.
                  rewrite sh_cid_re.
                  rewrite cid_re; auto.
                  rewrite zeq_true.
                  auto.
                - unfold relate_SyncChanPool_per_pd.
                  unfold relate_SyncChanPool_per_pd in syncchpool_re.
                  simpl; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                  rewrite CPU_ID_re.
                  rewrite ZMap.gss; auto.
                - rewrite zeq_true; auto.
                - inv inv_re; constructor; auto. }
              { rewrite ZMap.gso; auto.
                clear BUILTIN_ENABLED.
                rewrite H10 in sh_cid_re.
                repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                       end.
                subdestruct; auto.
                inv H3; simpl in ×.
                per_data_auto.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - auto.
                - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
                - rewrite zeq_false; auto.
                  rewrite sh_cid_re in _x.
                  rewrite cid_re; auto.
                - auto.
                - inv inv_re0; constructor; simpl; auto.
                - constructor. }
              { inv sh_shared_inv_re; constructor; simpl; auto. }
              { inv sh_mem_inv_re.
                rewrite <- cid_re in H2.
                rewrite <- cid_re.
                constructor; simpl; intros.
                case_eq (zeq i0 vm_handling_cid); intros; subst.
                - rewrite ZMap.gss in H4.
                  rewrite ZMap.gso in H13; eauto.
                  inv H4; simpl in H14.
                  eauto.
                - rewrite ZMap.gso in H4; auto.
                  case_eq (zeq j vm_handling_cid); intros; subst.
                  + rewrite ZMap.gss in H13; eauto.
                    inv H13.
                    simpl.
                    eapply pperm_disjoint; eauto.
                  + rewrite ZMap.gso in H13; eauto. } }
              Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          unfold pv_adt, sh_adt in *; subst.
          rewrite _x in *; rewrite zeq_true in ×.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
            { (gensem vmx_get_tsc_offset_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_get_tsc_offset_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10, H11, H12.
              rewrite H1.
              reflexivity. }
            { simpl; reflexivity. }
          × instantiate (1:= kctxt_pool).
            simpl in *; apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          Transparent proc_id update uRData.
          inv_monad H1.
          inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5.
          simpl in *; substx.
          simpl in *; rewrite _x in *; rewrite zeq_true in *; auto.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem vmx_get_exit_reason_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold vmx_get_exit_reason_spec in ×.
            rewrite <- ikern_re, <- pg_re, <- ihost_re, <- vmcs_re, <- CPU_ID_re.
            rewrite H4, H6, H7, H9, H10, H11.
            rewrite H1.
            reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in *; apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          unfold pv_adt, sh_adt in *; subst.
          rewrite _x in *; rewrite zeq_true in ×.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem vmx_get_exit_fault_addr_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold vmx_get_exit_fault_addr_spec.
            rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
            rewrite H4, H6, H7, H9, H10, H11.
            rewrite H1.
            reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in *; apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          unfold pv_adt, sh_adt in *; subst.
          rewrite _x in *; rewrite zeq_true in ×.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem vmx_get_exit_qualification_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold vmx_get_exit_qualification_spec.
            rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmx_re, <- CPU_ID_re.
            rewrite H4, H6, H7, H9, H10.
            rewrite H1.
            reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in *; apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          × unfold pv_adt, sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in ×.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_check_pending_event_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_check_pending_event_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10, H11, H12.
              rewrite H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
          × unfold pv_adt, sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in ×.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_check_pending_event_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_check_pending_event_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10, H11, H12.
              rewrite H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          × unfold pv_adt, sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in ×.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_check_int_shadow_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_check_int_shadow_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10, H11, H12.
              rewrite H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
          × unfold pv_adt, sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in ×.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_check_int_shadow_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_check_int_shadow_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10, H11, H12.
              rewrite H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          × unfold pv_adt, sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in ×.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_get_reg_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_get_reg_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- vmx_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10, H11.
              rewrite H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
          × unfold pv_adt, sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in ×.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_get_reg_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_get_reg_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- vmx_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10, H11, H12.
              rewrite H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          inv H8.
          Transparent proc_id update uRData.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H1.
          × simpl in *; substx.
            simpl in *; rewrite _x in *; rewrite zeq_true in *; auto.
             (a {vmx : ZMap.set (CPU_ID (update init_shared_adt l))
                                      (ZMap.set z (Vint (Int.repr (Int.unsigned i0)))
                                                (ZMap.get (CPU_ID (update init_shared_adt l))
                                                          (AbstractDataType.vmx a))) (AbstractDataType.vmx a)}).
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_set_reg_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_set_reg_spec in ×.
              rewrite <- ikern_re, <- pg_re, <- ihost_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H6, H7, H8, H10, H11.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in ×.
              generalize Hrel; intros Hrelate.
              rewrite H9 in cid_re.
              inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
              constructor; simpl; auto.
              { rewrite init_re; rewrite H9; auto. }
              intros.
              generalize (per_data_re i1); intros.
              case_eq (zeq i1 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
              { rewrite ZMap.gss.
                rewrite <- cid_re in H3, H2.
                rewrite <- cid_re.
                rewrite H2 in H3.
                inv H3; simpl in *; rewrite zeq_true in ×.
                per_data_auto.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - rewrite init_re.
                  rewrite H9 in sh_cid_re.
                  rewrite sh_cid_re.
                  rewrite cid_re; auto.
                  rewrite zeq_true.
                  auto.
                - unfold relate_SyncChanPool_per_pd.
                  unfold relate_SyncChanPool_per_pd in syncchpool_re.
                  simpl; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                  rewrite CPU_ID_re.
                  rewrite ZMap.gss; auto.
                  rewrite <- vmx_re.
                  reflexivity.
                - inv inv_re; constructor; simpl; auto. }
              { rewrite ZMap.gso; auto.
                clear BUILTIN_ENABLED.
                rewrite H9 in sh_cid_re.
                repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                       end.
                subdestruct; auto.
                inv H3; simpl in ×.
                per_data_auto.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - auto.
                - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
                - rewrite zeq_false; auto.
                  rewrite sh_cid_re in _x.
                  rewrite cid_re; auto.
                - inv inv_re0; constructor; simpl; auto.
                - constructor. }
              { inv sh_shared_inv_re; constructor; simpl; auto. }
              { inv sh_mem_inv_re.
                rewrite <- cid_re in H2.
                rewrite <- cid_re.
                constructor; simpl; intros.
                case_eq (zeq i1 vm_handling_cid); intros; subst.
                - rewrite ZMap.gss in H4.
                 rewrite ZMap.gso in H12; eauto.
                  inv H4; simpl in H13.
                  eauto.
                - rewrite ZMap.gso in H4; auto.
                  case_eq (zeq j vm_handling_cid); intros; subst.
                  + rewrite ZMap.gss in H12; eauto.
                    inv H12.
                    simpl.
                    eapply pperm_disjoint; eauto.
                  + rewrite ZMap.gso in H12; eauto. } }

          × simpl in *; substx.
            simpl in *; rewrite _x in *; rewrite zeq_true in *; auto.
             (a {vmcs : ZMap.set (CPU_ID (update init_shared_adt l))
                                       (VMCSValid revid abrtid (ZMap.set z (Vint (Int.repr (Int.unsigned i0))) d0))
                                       (AbstractDataType.vmcs a)}).
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_set_reg_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_set_reg_spec in ×.
              rewrite <- ikern_re, <- pg_re, <- ihost_re, <- vmcs_re, <- CPU_ID_re.
              rewrite H6, H7, H8, H10, H11, H12.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in ×.
              generalize Hrel; intros Hrelate.
              rewrite H9 in cid_re.
              inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
              constructor; simpl; auto.
              { rewrite init_re; rewrite H9; auto. }
              intros.
              generalize (per_data_re i1); intros.
              case_eq (zeq i1 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
              { rewrite ZMap.gss.
                rewrite <- cid_re in H3, H2.
                rewrite <- cid_re.
                rewrite H2 in H3.
                inv H3; simpl in *; rewrite zeq_true in ×.
                per_data_auto.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - rewrite init_re.
                  rewrite H9 in sh_cid_re.
                  rewrite sh_cid_re.
                  rewrite cid_re; auto.
                  rewrite zeq_true.
                  auto.
                - unfold relate_SyncChanPool_per_pd.
                  unfold relate_SyncChanPool_per_pd in syncchpool_re.
                  simpl; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                  rewrite CPU_ID_re.
                  rewrite ZMap.gss; auto.
                - rewrite zeq_true; auto.
                - inv inv_re; constructor; simpl; auto. }
              { rewrite ZMap.gso; auto.
                clear BUILTIN_ENABLED.
                rewrite H9 in sh_cid_re.
                repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                       end.
                subdestruct; auto.
                inv H3; simpl in ×.
                per_data_auto.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - auto.
                - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
                - rewrite zeq_false; auto.
                  rewrite sh_cid_re in _x.
                  rewrite cid_re; auto.
                - auto.
                - inv inv_re0; constructor; simpl; auto.
                - constructor. }
              { inv sh_shared_inv_re; constructor; simpl; auto. }
              { inv sh_mem_inv_re.
                rewrite <- cid_re in H2.
                rewrite <- cid_re.
                constructor; simpl; intros.
                case_eq (zeq i1 vm_handling_cid); intros; subst.
                - rewrite ZMap.gss in H4.
                  rewrite ZMap.gso in H13; eauto.
                  inv H4; simpl in H14.
                  eauto.
                - rewrite ZMap.gso in H4; auto.
                  case_eq (zeq j vm_handling_cid); intros; subst.
                  + rewrite ZMap.gss in H13; eauto.
                    inv H13.
                    simpl.
                    eapply pperm_disjoint; eauto.
                  + rewrite ZMap.gso in H13; eauto. } }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          unfold pv_adt, sh_adt in *; subst.
          rewrite _x in *; rewrite zeq_true in ×.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem vmx_get_next_eip_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold vmx_get_next_eip_spec.
            rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmcs_re, <- vmx_re, <- CPU_ID_re.
            rewrite H4, H6, H7, H9, H10, H11, H12.
            rewrite H1.
            reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in *; apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          × unfold pv_adt, sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in ×.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_get_io_width_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_get_io_width_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmx_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10, H11.
              rewrite H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
          × unfold pv_adt, sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in ×.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_get_io_width_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_get_io_width_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmx_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10, H11, H12.
              rewrite H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
          × unfold pv_adt, sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in ×.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_get_io_width_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_get_io_width_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmx_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10.
              unfold qsize in H11, H12; rewrite H11, H12.
              rewrite H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          × unfold pv_adt, sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in ×.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_get_io_write_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_get_io_write_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmx_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10, H11.
              rewrite H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
          × unfold pv_adt, sh_adt in *; subst.
            rewrite _x in *; rewrite zeq_true in ×.
            refine_split´; eauto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_get_io_write_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_get_io_write_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmx_re, <- CPU_ID_re.
              rewrite H4, H6, H7, H9, H10.
              unfold qdir in H11; rewrite H11.
              rewrite H1.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in *; apply relate_RData_gss; auto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          unfold pv_adt, sh_adt in *; subst.
          rewrite _x in *; rewrite zeq_true in ×.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem vmx_get_exit_io_rep_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold vmx_get_exit_io_rep_spec.
            rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmx_re, <- CPU_ID_re.
            rewrite H4, H6, H7, H9, H10.
            rewrite H1.
            reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in *; apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          unfold pv_adt, sh_adt in *; subst.
          rewrite _x in *; rewrite zeq_true in ×.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem vmx_get_exit_io_str_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold vmx_get_exit_io_str_spec.
            rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmx_re, <- CPU_ID_re.
            rewrite H4, H6, H7, H9, H10.
            rewrite H1.
            reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in *; apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          Transparent proc_id update uRData.
          inv_monad H1; inv H1.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H5; simpl in *; subst.
          unfold pv_adt, sh_adt in *; subst.
          rewrite _x in *; rewrite zeq_true in ×.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem vmx_get_exit_io_port_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            unfold vmx_get_exit_io_port_spec.
            rewrite <- ikern_re, <- ihost_re, <- pg_re, <- vmx_re, <- CPU_ID_re.
            rewrite H4, H6, H7, H9, H10.
            rewrite H1.
            reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in *; apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          destruct args; inv H8.
          inv H9.
          Transparent proc_id update uRData.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H1; simpl in *; subst.
          × unfold pv_adt, sh_adt in ×.
            functional inversion H6.
            functional inversion H11; substx; simpl in ×.
             (a {ept
                       : ZMap.set (CPU_ID (update init_shared_adt l))
                                  (ZMap.set (EPT_PML4_INDEX (Int.unsigned i))
                                            (EPML4EValid
                                               (ZMap.set (EPT_PDPT_INDEX (Int.unsigned i))
                                                         (EPDPTEValid
                                                            (ZMap.set (EPT_PDIR_INDEX (Int.unsigned i))
                                                                      (EPDTEValid
                                                                         (ZMap.set (EPT_PTAB_INDEX (Int.unsigned i))
                                                                                   (EPTEValid
                                                                                      (Z.lor
                                                                                         (Z.lor
                                                                                            (Z.land (Int.unsigned i0)
                                                                                                    4294963200) 71)
                                                                                         (2 × (2 × (2 × Int.unsigned i1)))))
                                                                                   eptab)) epdt)) epdpt))
                                            (ept d)) (AbstractDataType.ept a)}).
            refine_split´; eauto.
            rewrite _x in *; rewrite zeq_true in *; auto.
            { econstructor; [simpl | simpl; eauto].
               (gensem vmx_set_mmap_spec); split; [reflexivity |].
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; simpl in ×.
              unfold vmx_set_mmap_spec in ×.
              unfold ept_add_mapping_spec.
              unfold ept_invalidate_mappings_spec.
              rewrite <- ikern_re, <- ihost_re, <- pg_re, <- ept_re, <- CPU_ID_re.
              rewrite H13, H14, H15, H17, H18, H19, H20, H21, H22, <- H23.
              rewrite zle_lt_true; auto.
              rewrite zle_le_true; auto.
              rewrite zeq_true; simpl.
              rewrite <- ikern_re, <- ihost_re.
              rewrite H25, H26.
              reflexivity. }
            { instantiate (1:= kctxt_pool).
              simpl in ×.
              generalize Hrel; intros Hrelate.
              rewrite H16 in cid_re.
              inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
              constructor; simpl; auto.
              { rewrite init_re; rewrite H16; auto. }
              intros.
              generalize (per_data_re i2); intros.
              case_eq (zeq i2 (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
              { rewrite ZMap.gss.
                rewrite <- cid_re in H3, H2.
                rewrite <- cid_re.
                rewrite H2 in H3.
                inv H3; simpl in *; rewrite zeq_true in ×.
                per_data_auto_simpl.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - rewrite zeq_true; auto.
                - auto.
                - rewrite init_re.
                  rewrite cid_re; auto.
                  rewrite zeq_true; auto.
                - rewrite init_re.
                  rewrite cid_re; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - rewrite zeq_true; auto.
                - unfold relate_SyncChanPool_per_pd.
                  unfold relate_SyncChanPool_per_pd in syncchpool_re.
                  simpl; auto.
                - unfold relate_uctxt_per_pd in *; simpl; auto.
                - rewrite CPU_ID_re.
                  rewrite ZMap.gss.
                  rewrite <- CPU_ID_re.
                  rewrite cid_re; rewrite <- _x.
                  rewrite cid_re; rewrite zeq_true.
                  reflexivity.
                - auto.
                - auto.
                - inv inv_re; constructor; simpl.
                 auto. }
              { rewrite ZMap.gso; auto.
                clear BUILTIN_ENABLED.
                repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                       end.
                subdestruct; auto.
                inv H3; simpl in ×.
                per_data_auto_simpl.
                - unfold relate_AC_per_pd in *; simpl in *; auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - auto.
                - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
                - unfold relate_uctxt_per_pd in *; simpl; auto.
                - rewrite zeq_false; auto.
                  rewrite <- _x; auto.
                  rewrite cid_re; auto.
                - auto.
                - auto.
                - inv inv_re0; constructor; simpl.
                  auto.
                - constructor. }
              { inv sh_shared_inv_re; constructor; simpl; auto. }
              { inv sh_mem_inv_re.
                rewrite <- cid_re in H2.
                rewrite <- cid_re.
                constructor; simpl; intros.
                case_eq (zeq i2 (ZMap.get (CPU_ID (update init_shared_adt l))
                                          (cid (update init_shared_adt l)))); intros; subst.
                - rewrite ZMap.gss in H4.
                  rewrite ZMap.gso in H28; eauto.
                  inv H4; simpl in H29.
                  eauto.
                - rewrite ZMap.gso in H4; auto.
                  case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                           (cid (update init_shared_adt l)))); intros; subst.
                  + rewrite ZMap.gss in H28; eauto.
                    inv H28.
                    simpl.
                    eapply pperm_disjoint; eauto.
                  + rewrite ZMap.gso in H28; eauto. } }
          × unfold pv_adt, sh_adt in ×.
            functional inversion H6.
            inv H4; substx; simpl in ×.
            omega.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst; inv Hsem.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst; inv Hsem.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          Transparent proc_id update uRData.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in *; rewrite H2 in Hper_data_rel; inv Hper_data_rel; rewrite zeq_true in ×.
          functional inversion H6; simpl in *; subst.
          unfold pv_adt in *; subst.
          rewrite _x in *; rewrite zeq_true in ×.
           (((a {ept : ZMap.set (CPU_ID (update init_shared_adt l))
                                      (CalRealIntelModule.Calculate_EPDPTE
                                         3 (ZMap.set 0 (EPML4EValid (ZMap.init EPDPTEUndef)) (ept d)))
                                      (AbstractDataType.ept a)})
                     {vmcs : ZMap.set (CPU_ID (update init_shared_adt l))
                                      (CalRealIntelModule.Calculate_VMCS_at_i
                                         (CPU_ID (update init_shared_adt l))
                                         (ZMap.get (CPU_ID (update init_shared_adt l))
                                                   (cid (update init_shared_adt l))) (vmcs d) real_vmxinfo
                                         pml4ept_b stack_b idt_b msr_bitmap_b io_bitmap_b host_rip_b)
                                      (AbstractDataType.vmcs a)})
                    {vmx : ZMap.set (CPU_ID (update init_shared_adt l))
                                    (CalRealIntelModule.Calculate_VMX_at_i (vmx d))
                                    (AbstractDataType.vmx a)}).
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (IntelPrimSemantics.vmcs_set_defaults_compatsem vmx_init_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            econstructor; eauto.
            unfold vmx_init_spec.
            rewrite H17 in cid_re.
            rewrite <- init_re, <- pg_re, <- in_intr_re, <- ikern_re, <- ihost_re, <- ipt_re,
            <- ept_re, <- vmcs_re, <- vmx_re, <- cid_re, <- CPU_ID_re.
            rewrite H5, H10, H12, H15, H16, H17, H18.
            rewrite <- _x; rewrite H19.
            reflexivity.
          × instantiate (1:= kctxt_pool).
            simpl in ×.
            generalize Hrel; intros Hrelate.
            rewrite H17 in cid_re.
            inv Hrel; intros; simpl in *; rewrite cid_re in *; rewrite init_re in ×.
            constructor; simpl; auto.
            { rewrite init_re; rewrite H17; auto. }
            intros.
            generalize (per_data_re i); intros.
            case_eq (zeq i (ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a))); intros; subst.
            { rewrite ZMap.gss.
              rewrite <- cid_re in H3, H2.
              rewrite <- cid_re in H1.
              rewrite H2 in H1.
              inv H1; simpl in *; rewrite zeq_true in ×.
              per_data_auto_simpl.
              - rewrite _x.
                rewrite zeq_true.
                auto.
              - rewrite _x.
                rewrite zeq_true.
                auto.
              - unfold relate_AC_per_pd in *; simpl in *; auto.
              - rewrite _x.
                rewrite zeq_true.
                auto.
              - auto.
              - rewrite init_re.
                rewrite _x.
                rewrite zeq_true.
                auto.
              - rewrite <- cid_re.
                auto.
              - rewrite _x.
                rewrite zeq_true.
                auto.
              - rewrite _x.
                rewrite zeq_true.
                auto.
              - rewrite _x.
                rewrite zeq_true.
                auto.
              - unfold relate_SyncChanPool_per_pd.
                unfold relate_SyncChanPool_per_pd in syncchpool_re.
                simpl; auto.
              - unfold relate_uctxt_per_pd in *; simpl; auto.
              - rewrite cid_re.
                rewrite zeq_true.
                rewrite CPU_ID_re.
                rewrite ZMap.gss.
                reflexivity.
              - rewrite cid_re.
                rewrite zeq_true.
                rewrite CPU_ID_re.
                rewrite ZMap.gss.
                rewrite <- _x.
                rewrite <- CPU_ID_re.
                reflexivity.
              - rewrite cid_re.
                rewrite zeq_true.
                rewrite CPU_ID_re.
                rewrite ZMap.gss.
                reflexivity.
              - inv inv_re; constructor; simpl; auto. }
            { rewrite ZMap.gso; auto.
              clear BUILTIN_ENABLED.
              rewrite H17 in sh_cid_re.
              repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              subdestruct; auto.
              inv H1; simpl in ×.
              per_data_auto.
              - unfold relate_AC_per_pd in *; simpl in *; auto.
              - auto.
              - auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
              - rewrite zeq_false; auto.
                rewrite sh_cid_re in _x.
                rewrite cid_re; auto.
              - rewrite zeq_false; auto.
                rewrite sh_cid_re in _x.
                rewrite cid_re; auto.
              - rewrite zeq_false; auto.
                rewrite sh_cid_re in _x.
                rewrite cid_re; auto.
              - inv inv_re0; constructor; simpl; auto.
              - constructor. }
            { inv sh_shared_inv_re; constructor; simpl; auto. }
            { inv sh_mem_inv_re.
              rewrite <- cid_re in H2.
              rewrite <- cid_re.
              constructor; simpl; intros.
              case_eq (zeq i vm_handling_cid); intros; subst.
              - rewrite ZMap.gss in H3.
                rewrite ZMap.gso in H20; eauto.
                inv H3; simpl in H21.
                eauto.
              - rewrite ZMap.gso in H3; auto.
                case_eq (zeq j vm_handling_cid); intros; subst.
                + rewrite ZMap.gss in H20; eauto.
                  inv H20.
                  simpl.
                  eapply pperm_disjoint; eauto.
                + rewrite ZMap.gso in H20; eauto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          functional inversion H1; inv H4.
           (a {intr_flag: false}).
          Transparent proc_id update uRData.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem cli_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
          × instantiate (1:= kctxt_pool).
            simpl in ×.
            inv Hrel.
            constructor; simpl; auto.
            intros.
            generalize (per_data_re i); intros; simpl in ×.
            case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros; subst.
            { rewrite ZMap.gss.
              rewrite H2 in H3; inv H3; rewrite zeq_true in *; simpl.
              per_data_auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
            { rewrite ZMap.gso; auto.
              clear BUILTIN_ENABLED.
              repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              subdestruct; auto.
              inv H3; simpl in ×.
              per_data_auto.
              - auto.
              - auto.
              - auto.
              - rewrite zeq_false; auto.
                rewrite zeq_false in intr_flag_re; auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
              - auto.
              - inv inv_re; constructor; simpl; auto.
              - constructor. }
            { inv sh_shared_inv_re; constructor; simpl; auto. }
            { inv sh_mem_inv_re.
              constructor; simpl; intros.
              case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H4.
                rewrite ZMap.gso in H7; eauto.
                inv H4; simpl in H8.
                eauto.
              - rewrite ZMap.gso in H4; auto.
                case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                         (cid (update init_shared_adt l)))); intros; subst.
                + rewrite ZMap.gss in H7; eauto.
                  inv H7.
                  simpl.
                  eapply pperm_disjoint; eauto.
                + rewrite ZMap.gso in H7; eauto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          functional inversion H1; inv H4.
           (a {intr_flag: true}).
          Transparent proc_id update uRData.
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem sti_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
          × instantiate (1:= kctxt_pool).
            simpl in ×.
            inv Hrel.
            constructor; simpl; auto.
            intros.
            generalize (per_data_re i); intros; simpl in ×.
            case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros; subst.
            { rewrite ZMap.gss.
              rewrite H2 in H3; inv H3; rewrite zeq_true in *; simpl in ×.
              per_data_auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
            { rewrite ZMap.gso; auto.
              clear BUILTIN_ENABLED.
              repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              subdestruct; auto.
              inv H3; simpl in ×.
              per_data_auto.
              - auto.
              - auto.
              - auto.
              - rewrite zeq_false; auto.
                rewrite zeq_false in intr_flag_re; auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl in *; auto.
              - auto.
              - inv inv_re; constructor; simpl; auto.
              - constructor. }
            { inv sh_shared_inv_re; constructor; simpl; auto. }
            { inv sh_mem_inv_re.
              constructor; simpl; intros.
              case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H4.
                rewrite ZMap.gso in H7; eauto.
                inv H4; simpl in H8.
                eauto.
              - rewrite ZMap.gso in H4; auto.
                case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                         (cid (update init_shared_adt l)))); intros; subst.
                + rewrite ZMap.gss in H7; eauto.
                  inv H7.
                  simpl.
                  eapply pperm_disjoint; eauto.
                + rewrite ZMap.gso in H7; eauto. }
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          Transparent proc_id update uRData.
          exploit serial_intr_disable_related; simpl in *; eauto; intros Hrelated.
          destruct Hrelated as ( & Hrelated_a & Hrelated_b & Hrelated_c & Hrelated_d & Hrelated_e).
           .
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem serial_intr_disable_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            eauto.
          × instantiate (1:= kctxt_pool).
            rewrite <- Hrelated_c in Hrelated_b.
            auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          Transparent proc_id update uRData.
          exploit serial_intr_enable_related; simpl in *; eauto. intros Hrelated.
          destruct Hrelated as ( & Hrelated_a & Hrelated_b & Hrelated_c & Hrelated_d & Hrelated_e).
           .
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem serial_intr_enable_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            eauto.
          × instantiate (1:= kctxt_pool).
            rewrite <- Hrelated_c in Hrelated_b.
            auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          destruct args; inv H1.
          inv H4.
          Transparent proc_id update uRData.
          exploit serial_putc_related; simpl in *; eauto. intros Hrelated.
          destruct Hrelated as ( & Hrelated_a & Hrelated_b & Hrelated_c & Hrelated_d & Hrelated_e).
           .
          refine_split´; eauto.
          × econstructor; [simpl | simpl; eauto].
             (gensem serial_putc_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; simpl in ×.
            eauto.
          × instantiate (1:= kctxt_pool).
            rewrite <- Hrelated_c in Hrelated_b.
            auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          inv H6.
          Transparent proc_id update uRData.
          exploit cons_buf_read_related; simpl in *; eauto. intros Hrelated.
          destruct Hrelated as ( & res´ & Hrelated_a & Hrelated_b & Hrelated_c & Hrelated_d & Hrelated_e & Hrelated_f); subst.
           .
          refine_split´; eauto.
          econstructor; [simpl | simpl; eauto].
           (gensem cons_buf_read_spec); split; [reflexivity |].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor; simpl in ×.
          eauto.
          Opaque proc_id update uRData.
        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst; inv Hsem; simpl in *; inv Hsig.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst; inv Hsem; simpl in *; inv Hsig.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          functional inversion H1; simpl in *; subst; auto.
          Transparent proc_id update uRData.
          refine_split´; auto.
          × econstructor; [simpl | simpl; eauto].
             (gensem proc_create_postinit_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor.
            unfold proc_create_postinit_spec in *; simpl in *; subdestruct.
            generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
            simpl in ×.
            rewrite H2 in Hper_data_rel.
            inv Hper_data_rel; simpl in *; rewrite zeq_true in ×.
            rewrite <- ikern_re, <- ihost_re.
            rewrite H5, H6.
            repeat (rewrite zeq_true; auto).
            reflexivity.
          × simpl.
            instantiate (1:= kctxt_pool).
            apply relate_RData_gss; auto.
            Opaque proc_id update uRData.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst; inv Hsem; simpl in *; inv Hsig.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst; inv Hsem; simpl in *; inv Hsig.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst; inv Hsem; simpl in *; inv Hsig.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst; inv Hsem; simpl in *; inv Hsig.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst; inv Hsem; simpl in *; inv Hsig.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst; inv Hsem; simpl in *; inv Hsig.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          subdestruct; inv H3.
          refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
          econstructor; [simpl | simpl; eauto].
           sextcall_longoffloat_compatsem; split; [reflexivity | ].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor; auto.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          subdestruct; inv H3.
          refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
          econstructor; [simpl | simpl; eauto].
           sextcall_longuoffloat_compatsem; split; [reflexivity | ].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor; auto.

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          subdestruct; inv H3.
          { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
            econstructor; [simpl | simpl; eauto].
             sextcall_floatoflong_compatsem; split; [reflexivity | ];
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; auto. }
          { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
            econstructor; [simpl | simpl; eauto].
             sextcall_floatoflong_compatsem; split; [reflexivity | ];
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; auto. }

        +
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          subdestruct; inv H3.
          { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
            econstructor; [simpl | simpl; eauto].
             sextcall_floatoflongu_compatsem; split; [reflexivity | ];
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; auto. }
          { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
            econstructor; [simpl | simpl; eauto].
             sextcall_floatoflongu_compatsem; split; [reflexivity | ];
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; auto. }

        +
            destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
            inv Hsem; simpl in *; inv Hsig.
            subdestruct; inv H3.
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_singleoflong_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_singleoflong_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }
            
        +
            destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
            inv Hsem; simpl in *; inv Hsig.
            subdestruct; inv H3.
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_singleoflongu_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_singleoflongu_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }

          +
            destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
            inv Hsem; simpl in *; inv Hsig.
            subdestruct; inv H3.
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_divls_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_divls_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }

          +
            destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
            inv Hsem; simpl in *; inv Hsig.
            subdestruct; inv H3.
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_divlu_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_divlu_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }

          +
            destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
            inv Hsem; simpl in *; inv Hsig.
            subdestruct; inv H3.
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_modls_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_modls_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }

          +
            destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
            inv Hsem; simpl in *; inv Hsig.
            subdestruct; inv H3.
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_modlu_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }
            { refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
              econstructor; [simpl | simpl; eauto].
               sextcall_modlu_compatsem; split; [reflexivity | ];
               s; refine_split´; auto; try reflexivity; simpl.
              repeat constructor; auto. }

          +
            destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
            inv Hsem; simpl in *; inv Hsig.
            subdestruct; inv H3.
            refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
            econstructor; [simpl | simpl; eauto].
             sextcall_shll_compatsem; split; [reflexivity | ];
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; auto.

          +
            destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
            inv Hsem; simpl in *; inv Hsig.
            subdestruct; inv H3.
            refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
            econstructor; [simpl | simpl; eauto].
             sextcall_shrlu_compatsem; split; [reflexivity | ];
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; auto.

          +
            destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
            inv Hsem; simpl in *; inv Hsig.
            subdestruct; inv H3.
            refine_split´; auto; [ | eapply relate_RData_gss; auto; instantiate (1:= kctxt_pool); eauto].
            econstructor; [simpl | simpl; eauto].
             sextcall_shrl_compatsem; split; [reflexivity | ];
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor; auto.

      - assert (Hext:= H); apply exec_external_prop in H; auto; inv H.
        refine_split´; auto.
        + inv Hext.
          clear BUILTIN_ENABLED.
          destruct ef; inv Hcase; inv H.
          × econstructor; simpl; eauto; simpl in *; rewritesb.
            apply CompCertBuiltins.builtin_sem_i64_neg; auto.
          × econstructor; simpl; eauto; simpl in *; rewritesb.
            apply CompCertBuiltins.builtin_sem_i64_add; auto.
          × econstructor; simpl; eauto; simpl in *; rewritesb.
            apply CompCertBuiltins.builtin_sem_i64_sub; auto.
          × econstructor; simpl; eauto; simpl in *; rewritesb.
            apply CompCertBuiltins.builtin_sem_i64_mul; auto.
          × econstructor; simpl; eauto; rewritesb.
            destruct args; [inv H1 | ].
            inv H1; inv H7.
            { repeat constructor; auto. }
            { repeat constructor; auto. }
          × econstructor; simpl; eauto; rewritesb.
            inv H1; repeat constructor; auto.
            destruct args; [ inv H4 |].
            simpl in H4; subdestruct.
            { inv H4; inv H3; subst; [ econstructor; auto | ].
              econstructor; auto; unfold Mem.store in ×.
              simpl in *; lift_unfold; destruct H1; split; auto. }
            { inv H4; inv H3; subst; [ econstructor; auto | ].
              econstructor; auto; unfold Mem.store in ×.
              simpl in *; lift_unfold; destruct H1; split; auto. }
            { inv H4; inv H3; subst; [ econstructor; auto | ].
              econstructor; auto; unfold Mem.store in ×.
              simpl in *; lift_unfold; destruct H1; split; auto. }
            { inv H4; inv H3; subst; [ econstructor; auto | ].
              econstructor; auto; unfold Mem.store in ×.
              simpl in *; lift_unfold; destruct H1; split; auto. }
            { inv H4; inv H3; subst; [ econstructor; auto | ].
              econstructor; auto; unfold Mem.store in ×.
              simpl in *; lift_unfold; destruct H1; split; auto. }

          × econstructor; simpl; eauto.
            simpl in H4; subst.
            econstructor; eauto.
            inv H8; repeat constructor; auto.
          × econstructor; simpl; eauto; rewritesb.
            econstructor; eauto.
            inv H1; repeat constructor; auto.
            subdestruct; inv H5; inv H4; constructor; auto; lift_unfold; destruct H1; tauto.
          × destruct m´0 as [m0 d0].
            destruct args; inv H1.
            simpl; econstructor; try eauto.
            { unfold Mem.alloc in H3; simpl in H3.
              unfold Mem.store in H4; simpl in H4.
              econstructor.
              - instantiate (1 := b).
                instantiate (1:= (m0, a)); lift_unfold.
                destruct H3 as (H3a, H3b); destruct H4 as (H4a, H4b); subst; tauto.
              - lift_unfold; tauto. }
            { eauto. }
          × econstructor; simpl; eauto; rewritesb.
            econstructor; eauto.
            lift_unfold; intuition.
          × econstructor; simpl; eauto; rewritesb.
            econstructor; eauto.
            lift_unfold; intuition.
          × econstructor; simpl; eauto.
            constructor; auto.
          × econstructor; simpl; eauto; rewritesb.
            constructor; auto.
        + apply relate_RData_gss.
          auto.
          eauto.
    Fail idtac.
   Qed.
    Next Obligation.
      rename H1 into kctxt_pool.
      rename H4 into Hrel.
      rename H3 into Hlval.
      destruct (ef_cases ef) as [Hcase|Hcase].
      - destruct ef; inv Hcase.
        inv H; inv H1.
        destruct H as [Hprims Hsem].
        inv_layer Hprims; no_event_remove BUILTIN_ENABLED.

        +
          assert (in_event: has_event palloc = true) by (unfold has_event; auto); rewrite in_event; clear in_event.
          assert (prim_num: prim_id_num palloc = 1) by (unfold prim_id_num; simpl; auto).
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          unfold initial_thread_kctxt.
          unfold prim_thread_init_pc.
          rewrite peq_false; [ | intro contra; inv contra].
          inversion Hlval; subst.
          inv H5; inv H7.
          subdestruct; eauto.
          destruct BUILTIN_ENABLED as (Hchoice_check & Hl).
          Transparent proc_id uRData update.
          exploit big_palloc_related; eauto; simpl in *; eauto.
          intros Hrelated.
          destruct Hrelated as ( & res´ &Hrelated_a & Hrelated_b & Hrelated_c &
                                Hrelated_d & Hrelated_e & Hrelated_f).
           ; subst.
          refine_split´; eauto.
          simpl; econstructor; [simpl | simpl; eauto].
           (gensem big_palloc_spec); split; [reflexivity |].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor.
          eauto.
          Opaque proc_id uRData update.

        +
          assert (in_event: has_event pt_resv = true) by (unfold has_event; auto); rewrite in_event; clear in_event.
          assert (prim_num: prim_id_num pt_resv = 2) by (unfold prim_id_num; simpl; auto).
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          destruct args; inv H8.
          unfold initial_thread_kctxt.
          unfold prim_thread_init_pc.
          rewrite peq_false; [ | intro contra; inv contra].
          inversion Hlval; subst.
          inversion H6; subst.
          inversion H8; subst.
          inv H4; inv H5; inv H7.
          clear H8 H6.
          inv H9.
          inv H11.
          subdestruct; eauto.
          destruct BUILTIN_ENABLED as (Hchoice_check & Hl).
          Transparent proc_id uRData update.
          exploit big_ptResv_related; eauto; simpl in *; eauto.
          intros Hrelated.
          destruct Hrelated as ( & res´ &Hrelated_a & Hrelated_b & Hrelated_c & Hrelated_d & Hrelated_e & Hrelated_f).
           ; subst.
          refine_split´; eauto.
          simpl; econstructor; [simpl | simpl; eauto].
           (gensem big_ptResv_spec); split; [reflexivity |].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor.
          eauto.
          Opaque proc_id uRData update.

        +
          assert (in_event: has_event thread_wakeup = true) by (unfold has_event; auto); rewrite in_event; clear in_event.
          assert (prim_num: prim_id_num thread_wakeup = 4) by (unfold prim_id_num; simpl; auto).
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          unfold initial_thread_kctxt.
          unfold prim_thread_init_pc.
          rewrite peq_false; [ | intro contra; inv contra].
          inversion Hlval; subst.
          inv H5; inv H7.
          subdestruct; eauto.
          destruct BUILTIN_ENABLED as (Hchoice_check & Hl).
          Transparent proc_id uRData update.
          exploit thread_wakeup_related; eauto; simpl in *; eauto.
          intros Hrelated.
          destruct Hrelated as ( & res´ & Hrelated_a & Hrelated_b & Hrelated_c & Hrelated_d & Hrelated_e & Hrelated_f).
           ; subst.
          refine_split´; eauto.
          simpl; econstructor; [simpl | simpl; eauto].
           (gensem big_thread_wakeup_spec); split; [reflexivity |].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor.
          eauto.
          Opaque proc_id uRData update.

        +
          assert (in_event: has_event sched_init = true) by (unfold has_event; auto); rewrite in_event; clear in_event.
          assert (prim_num: prim_id_num sched_init = 5) by (unfold prim_id_num; simpl; auto).
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          unfold initial_thread_kctxt.
          unfold prim_thread_init_pc.
          rewrite peq_false; [ | intro contra; inv contra].
          inversion Hlval; subst.
          inv H5; inv H7.
          subdestruct; eauto.
          destruct BUILTIN_ENABLED as (Hchoice_check & Hl).
          Transparent proc_id uRData update.
          exploit sched_init_related; eauto; simpl in *; eauto.
          intros Hrelated.
          destruct Hrelated as ( & Hrelated_a & Hrelated_b & Hrelated_c & Hrelated_d).
           ; subst.
          refine_split´; eauto.
          simpl; econstructor; [simpl | simpl; eauto].
           (gensem big_sched_init_spec); split; [reflexivity |].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor.
          eauto.
          Opaque proc_id uRData update.

        +
          assert (in_event: has_event proc_create = true) by (unfold has_event; auto); rewrite in_event.
          assert (prim_num: prim_id_num proc_create = 3) by (unfold prim_id_num; simpl; auto).
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H1.
          destruct args; inv H5.
          destruct args; inv H4; subst.
          inversion Hlval.
          inv H6.
          inv H15.
          inv H8.
          inv H4; inv H13; inv H5.
          destruct H11 as (fun_id & H11).
          rewrite peq_false in BUILTIN_ENABLED; [|intro contra; inv contra].
          rewrite peq_false in BUILTIN_ENABLED; [|intro contra; inv contra].
          rewrite in_event in BUILTIN_ENABLED.
          destruct BUILTIN_ENABLED as (Hchoice_check & Hl).
          unfold prim_thread_init_pc.
          rewrite peq_true; auto.
          Transparent proc_id uRData update.
          generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
          simpl in ×.
          rewrite H2 in Hper_data_rel.
          inv Hper_data_rel.
          rewrite zeq_true in ×.
          functional inversion H7.
          functional inversion H18.
          unfold single_big_proc_create_spec_share; simpl in ×.

          rewrite prim_num.
          unfold choice_check.
          rewrite prim_num.
          unfold ObjPHBThreadSCHED.single_big_proc_create_spec_test.
          unfold ObjPHBThreadSCHED.single_big_proc_create_spec_share.
          simpl in ×.
          unfold big_proc_create_spec; subst.
          rewritesb.
          rewrite H22, H21, H23.
          unfold curid, cpu0, c in H24; simpl in *; rewrite H24.
          case_eq (in_dec zeq (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                               Z.of_nat (length (cchildren (AC d)))) full_thread_list); intros; subst; [ | elim n0; auto].
          unfold in H26.
          unfold cpu0, e, l2 in H26.
          unfold curid0, to, cpu0 in H26.
          unfold ofs in H26.
          unfold curid, c in H26.
          rewrite H26.
          rewrite H23 in H16; inversion H16; substx.
          rewrite H22 in ×.
          assert (Hcused : cused (ZMap.get (proc_id (uRData l)) (AbstractDataType.AC a)) = true).
          { inv Hrel.
            inv sh_shared_inv_re.
            rewrite <- H17.
            symmetry.
            unfold B_GetContainerUsed´ in container_used_inv.
            rewrite H23 in container_used_inv.
            apply container_used_inv.
            generalize (all_cid_in_full_thread_list l).
            simpl.
            intros; auto. }
          assert (Hcused_ch : cused (ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))
                                               × 8 + 1 + Z.of_nat (length (cchildren (AC d))))
                                              (AbstractDataType.AC a)) = false).
          { inv Hrel.
            rewrite <- H24.
            inv sh_shared_inv_re.
            unfold B_GetContainerUsed´ in container_used_inv.
            rewrite H23 in container_used_inv.
            symmetry.
            apply container_used_inv; auto. }

          unfold relate_AC_per_pd in AC_re.
          rewrite H22 in ×.
          unfold B_GetContainerUsed´ in AC_re.
          rewrite H23 in AC_re.
          rewrite H17 in AC_re.
          assert (valid_procid: proc_id (uRData l) = main_thread TOTAL_CPU < proc_id (uRData l) < num_proc).
          { generalize (all_cid_in_full_thread_list l); intros.
            generalize (valid_thread_list (proc_id (uRData l))); intros.
            eapply H19 in H3.
            tauto. }
           assert (HAC_eq : AC d = (ZMap.get (proc_id (uRData l)) (AbstractDataType.AC a))).
          { destruct valid_procid as [ valid_proc | valid_proc].
            - rewrite <- valid_proc in AC_re.
              rewrite zeq_true in AC_re.
              rewrite AC_re; auto.
            - destruct (zeq (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))) main_thread); auto.
              rewrite zlt_lt_true in AC_re; auto. }
          generalize (CPU_ID_valid l); intros CPU_ID_valid; simpl in ×.
          generalize main_thread_val; intros main_thread_val.
          generalize current_CPU_ID_range; intros current_CPU_ID_range.
          clear H16.
          clear AC_re.
          rename HAC_eq into AC_re.
          exploit (abs_rel_proc_create_prop ge kctxt_pool l dp a d b buc ofs_uc q).
          { auto. }
          { simpl; auto. }
          { unfold single_big_proc_create_spec; simpl.
            unfold single_big_proc_create_spec_share; simpl.
            rewrite H13, H14, H15, H21.
            rewrite zlt_lt_true; auto.
            rewrite zlt_true; auto.
            rewrite zle_le_true; auto.
            rewrite H5.
            rewrite H23.
            rewrite H17, H24, H25, H22.
            rewrite H26.
            auto. }
          { unfold big_proc_create_spec; simpl.
            rewrite <- pg_re, <- ikern_re, <- ihost_re, <- ipt_re.
            rewrite H13, H14, H15, H21.
            rewrite <- cid_re.
            rewrite Hcused.
            rewrite <- AC_re.
            rewrite Hcused_ch.
            rewrite zlt_lt_true; auto.
            rewrite zlt_true; auto.
            rewrite zle_le_true; auto.
            rewrite <- big_log_re; rewrite H23.
            rewrite <- CPU_ID_re, <- big_oracle_re.
            rewrite H26.
            reflexivity. }
          { reflexivity. }
          { instantiate (1:= elf_id).
            reflexivity. }
          intros Hproc_prop.
          simpl in Hproc_prop; destruct Hproc_prop as (Hproc_init_dproc & Hproc_kctxt).
          assert (ch_id_range:8 < (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                                   Z.of_nat (length (cchildren (AC d)))) < 1024).
          { assert (0 Z.of_nat (length (cchildren (AC d)))).
            { eauto using Nat2Z.is_nonneg. }
            revert H3.
            clear H12.
            revert _x.
            revert valid_procid.
            simpl.
            rewrite main_thread_val.
            generalize current_CPU_ID_range.
            remember (Z.of_nat (length (cchildren (AC d)))) as .
            remember (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))) as parent; clear.
            intros; omega. }
          { assert (ch_id_neq: (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                ((ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                   × 8 + 1 + Z.of_nat (length (cchildren (AC d))))).
            { assert (0 Z.of_nat (length (cchildren (AC d)))).
              { eauto using Nat2Z.is_nonneg. }
              revert H3.
              clear H12.
              revert _x.
              revert valid_procid.
              simpl.
              remember (Z.of_nat (length (cchildren (AC d)))) as .
              remember (ZMap.get (CPU_ID (update init_shared_adt l))
                                 (cid (update init_shared_adt l))) as parent; clear.
              intros; omega. }
          refine_split; auto.
          × simpl; econstructor; [simpl | simpl; eauto].
             (proc_create_compatsem big_proc_create_spec); split; [reflexivity |].
             s; refine_split´; auto; try reflexivity; simpl.
            repeat constructor.
            eapply extcall_proccreate_sem_intro; eauto.
            unfold big_proc_create_spec; subst.
            rewritesb.
            simpl in ×.
            rewrite H21, H14, H15, H13.
            rewrite AC_re.
            rewrite Hcused.
            rewrite AC_re in Hcused_ch.
            rewrite Hcused_ch.
            rewrite zlt_lt_true; auto.
            rewrite zlt_true; auto.
            rewrite zle_le_true; auto.
            rewrite H23.
            rewrite AC_re in H26.
            rewrite H26.
            reflexivity.
            { rewrite <- AC_re; auto. }
            { rewrite <- AC_re; assumption. }
            { rewrite <- AC_re; assumption. }
          × instantiate (1:= ZMap.set (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                                       Z.of_nat (length (cchildren (AC d))))
                                      ((ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l))
                                                           (cid (update init_shared_adt l)) × 8 + 1 +
                                                  Z.of_nat (length (cchildren (AC d)))) (kctxt a))
                                         # ESP <- (Vptr b (Int.repr ((ZMap.get (CPU_ID (update init_shared_adt l))
                                                                               (cid (update init_shared_adt l))
                                                                      × 8 + 1 +
                                                                      Z.of_nat (length (cchildren (AC d))) + 1)
                                                                     × 4096 - 4)))) # RA <- (Vptr Int.zero) (kctxt a)).
            constructor; simpl; auto.
            { rewrite H22; intros; auto. }
            { rewrite H22; intros; auto. }
            { rewrite AC_re; auto. }
            { rewrite AC_re; auto. }
            intros .
            case_eq (zeq (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros.
            { subst.
              rewrite ZMap.gss.
              per_data_auto_simpl.
              - rewrite H22; auto.
              - rewrite H22; auto.
              - rewrite AC_re; auto.
              - rewrite AC_re; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_AC_per_pd.
                simpl.
                rewrite ZMap.gso; auto.
                rewrite ZMap.gss.
                rewrite <- Hcused.
                rewrite <- AC_re.
                rewrite H22.
                exploit single_big_proc_create_preserves_cused; eauto.
                { simpl; inv Hrel; inv sh_shared_inv_re; simpl; auto. }
                { simpl; reflexivity. }
                { intros.
                  simpl in H16.
                  rewrite <- H16.
                  rewrite H17.
                  destruct valid_procid as [ valid_proc | valid_proc].
                  - rewrite <- valid_proc.
                    rewrite zeq_true; auto.
                  - destruct (zeq (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))
                                  main_thread); auto.
                    rewrite zlt_lt_true; auto. }
                rewrite <- AC_re.
                assumption.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite H22; auto.
                rewrite zeq_true.
                auto.
              - rewrite H22.
                auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_SyncChanPool_per_pd.
                simpl.
                rewrite H22.
                rewrite zeq_true; auto.
              - unfold relate_uctxt_per_pd.
                simpl.
                rewrite ZMap.gso; auto.
                rewrite H22.
                exploit single_big_proc_create_preserves_cused; eauto.
                { simpl; inv Hrel; inv sh_shared_inv_re; simpl; auto. }
                { simpl; reflexivity. }
                { intros.
                  simpl in H16.
                  rewrite <- H16.
                  rewrite H17.
                  unfold relate_uctxt_per_pd in uctxt_re.
                  unfold B_GetContainerUsed´ in uctxt_re.
                  rewrite H23, H22, H17 in uctxt_re.
                  auto. }
                rewrite <- AC_re.
                assumption.
              - auto.
              - auto.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
            { subst.
              rewrite ZMap.gso; auto.
              generalize (per_data_re _ _ _ _ Hrel ); intros Hper_data_rel.
              repeat match goal with
                       | [H: context[if _ then _ else _] |- _] ⇒ clear H
                     end.
              subdestruct; simpl; auto.
              inv Hper_data_rel.
              per_data_auto.
              - rewrite AC_re; auto.
              - rewrite AC_re; auto.
              - unfold relate_AC_per_pd in *; simpl in ×.
                rewrite H22 in ×.
                exploit single_big_proc_create_preserves_cused; eauto.
                { simpl; inv Hrel; inv sh_shared_inv_re; simpl; auto. }
                { simpl; reflexivity. }
                { intros.
                  simpl in H16.
                  case_eq (zeq (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))
                                   × 8 + 1 + Z.of_nat (length (cchildren (AC d))))).
                  + intros; subst.
                    unfold B_GetContainerUsed´ in AC_re0.
                    rewrite H23 in AC_re0; rewrite H24 in AC_re0.
                    rewrite <- AC_re.
                    rewrite zlt_lt_true in AC_re0; auto.
                    rewrite AC_re0.
                    rewrite AC_re.
                    rewrite ZMap.gss.
                    unfold B_GetContainerUsed; simpl.
                    rewrite <- CPU_ID_valid in ×.
                    rewrite zeq_false.
                    rewrite <- AC_re.
                    case (zle_le 0 (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                                    Z.of_nat (length (cchildren (AC d)))) 8); intros.
                    { revert a0.
                      revert ch_id_range.
                      remember (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                                Z.of_nat (length (cchildren (AC d))) ) as child_id.
                      clear.
                      intros.
                      omega. }
                    { repeat rewrite zeq_true.
                      rewrite zeq_false.
                      rewrite zlt_lt_true; auto.
                      simpl.
                      rewrite Hproc_init_dproc.
                      simpl.
                      reflexivity.
                      rewrite main_thread_val.
                      generalize current_CPU_ID_range.
                      generalize ch_id_range.
                      remember (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                                Z.of_nat (length (cchildren (AC d))) ) as child_id.
                      clear.
                      intros.
                      intro contra.
                      subst.
                      omega. }
                    generalize current_CPU_ID_range.
                    generalize ch_id_range.
                    rewrite <- AC_re.
                    remember (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                              Z.of_nat (length (cchildren (AC d))) ) as child_id.
                    clear.
                    intros.
                    intro contra; subst.
                    omega.

                  + intros.
                    unfold relate_AC_per_pd in AC_re0.
                    simpl.
                    rewrite ZMap.gso; auto.
                    rewrite ZMap.gso; auto.
                    exploit (single_big_proc_create_preserves_cused_snd ); eauto.
                    { simpl; inv Hrel; inv sh_shared_inv_re; simpl; auto. }
                    { simpl; reflexivity. }
                    { intros; rewrite <- H4.
                      assumption. }
                    { intros; simpl in H20.
                      rewrite <- H20.
                      unfold B_GetContainerUsed´ in AC_re0.
                      rewrite H23 in AC_re0; auto. }
                    rewrite <- AC_re; auto. }
              - auto.
              - rewrite H22.
                rewrite H22 in PT_re0.
                auto.
              - unfold relate_SyncChanPool_per_pd.
                simpl; rewrite H22.
                rewrite zeq_true; auto.
              - unfold relate_uctxt_per_pd in *; simpl in ×.
                rewrite H22 in ×.
                exploit single_big_proc_create_preserves_cused; eauto.
                { simpl; inv Hrel; inv sh_shared_inv_re; simpl; auto. }
                { simpl; reflexivity. }
                { intros.
                  simpl in H16.
                  rewrite <- AC_re.
                  case_eq (zeq (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l))
                                   × 8 + 1 + Z.of_nat (length (cchildren (AC d))))).
                  + intros; subst.
                    unfold B_GetContainerUsed´ in uctxt_re0.
                    rewrite H23 in uctxt_re0; rewrite H24 in uctxt_re0.
                    rewrite zlt_lt_true in uctxt_re0.
                    unfold B_GetContainerUsed; simpl.
                    rewrite zeq_false; auto.
                    case (zle_le 0 (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                                    Z.of_nat (length (cchildren (AC d)))) 8); intros.
                    { revert a0.
                      revert ch_id_range.
                      remember (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                                Z.of_nat (length (cchildren (AC d))) ) as child_id.
                      clear.
                      intros.
                      omega. }
                    { repeat rewrite zeq_true.
                      generalize current_CPU_ID_range.
                      generalize ch_id_range; intros.
                      rewrite zeq_false.
                      rewrite zlt_lt_true; auto.
                      rewrite uctxt_re0.
                      rewrite Hproc_init_dproc.
                      simpl.
                      unfold uctxt_bieq.
                      rewrite ZMap.gss.
                      simpl.
                      unfold Int.zero.
                      rewrite Int.eq_true.
                      split.
                      intros.
                      rewrite ZMap.gss.
                      generalize (sh_shared_inv_re kctxt_pool (update init_shared_adt l) dp a Hrel).
                      intros Htemp.
                      inv Htemp.
                      unfold B_GetContainerUsed´ in uctxt_used_inv.
                      rewrite H23 in uctxt_used_inv.
                      rewrite (uctxt_used_inv _ H24).
                      rewrite ZMap.gso.
                      symmetry.
                      rewrite ZMap.gso.
                      reflexivity.
                      assumption.
                      assumption.

                       (Vptr buc ofs_uc).
                      rewrite ZMap.gss.
                      rewrite ZMap.gss.
                      auto.
                      rewrite <- CPU_ID_valid in ×.

                      rewrite main_thread_val.
                      generalize current_CPU_ID_range.
                      generalize ch_id_range.
                      remember (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                                Z.of_nat (length (cchildren (AC d))) ) as child_id.
                      clear.
                      intros.
                      intro contra.
                      subst.
                      omega. }

                    generalize current_CPU_ID_range.
                    rewrite <- CPU_ID_valid.
                    generalize ch_id_range.
                    remember (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                              Z.of_nat (length (cchildren (AC d))) ) as child_id.
                    clear.
                    intros.
                    intro contra; subst.
                    omega.

                    assumption.
                  + intros.
                    unfold relate_uctxt_per_pd in uctxt_re0.
                    simpl.
                    rewrite ZMap.gso; auto.
                    exploit (single_big_proc_create_preserves_cused_snd ); eauto.
                    { simpl; inv Hrel; inv sh_shared_inv_re; simpl; auto. }
                    { simpl; reflexivity. }
                    { intros; rewrite <- H4.
                      assumption. }
                    { intros.
                      simpl in H20.
                      rewrite <- H20.
                      unfold B_GetContainerUsed´ in uctxt_re0.
                      rewrite H23 in uctxt_re0.
                      auto. } }
              - auto.
              - inv inv_re0; constructor; simpl; auto.
              - constructor. }
            { generalize Hrel; intro Hrelated; inv Hrel; inv sh_shared_inv_re.
              constructor; simpl; auto.
              - rewrite H22; intro contra; inv contra.
              - rewrite H22; intro contra; inv contra.
              - rewrite zeq_true; intros.
                rewrite H5 in syncchpool_inv.
                apply syncchpool_inv in H3.
                rewrite <- H3.
                symmetry.
                unfold B_GetlastPush.
                rewrite H23.
                eapply single_big_oracle_query_preserves_B_GetlastPush_aux; auto.
              - intros.
                case_eq (zeq tid (ZMap.get (CPU_ID (update init_shared_adt l))
                                           (cid (update init_shared_adt l)) × 8 + 1 +
                                  Z.of_nat (length (cchildren (AC d))))).
                + intros.
                  rewrite <- AC_re.
                  unfold B_GetContainerUsed.
                  subst.
                  rewrite ZMap.gss.
                  rewrite zeq_false; auto.
                  case_eq (zle_le 0 (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                                     Z.of_nat (length (cchildren (AC d)))) 8); intros; subst.
                  × clear H19; revert a0.
                    revert ch_id_range.
                    remember (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                              Z.of_nat (length (cchildren (AC d))) ) as child_id.
                    clear.
                    intros; omega.
                  × intros.
                    simpl.
                    rewrite zeq_true; auto.
                    rewrite zeq_true; auto.
                  × rewrite <- CPU_ID_valid in ×.
                    generalize current_CPU_ID_range.
                    generalize ch_id_range.
                    remember (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                              Z.of_nat (length (cchildren (AC d))) ) as child_id.
                    clear.
                    intros.
                    intro contra; subst.
                    omega.
                + intros.
                  rewrite H4.
                  rewrite <- AC_re.
                  rewrite H4.
                  exploit (single_big_proc_create_preserves_cused_snd tid (uRData l, d)); eauto.
                  { rewrite H4.
                    simpl; auto. }
                  { simpl.
                    rewrite <- H4.
                    assumption. }
                  rewrite ZMap.gso; auto.
                  intros.
                  simpl in H19.
                  rewrite <- H19.
                  case_eq (zeq tid (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))).
                  × intros.
                    subst.
                    rewrite ZMap.gss.
                    simpl.
                    auto.
                  × intros; rewrite ZMap.gso; auto.
                    apply container_used_inv in H3; auto.
                    unfold B_GetContainerUsed´ in H3; rewrite H23 in H3; auto.
                  × rewrite <- H4.
                    auto.

              - intros.
                case_eq (zeq tid (ZMap.get (CPU_ID (update init_shared_adt l))
                                           (cid (update init_shared_adt l)) × 8 + 1 +
                                  Z.of_nat (length (cchildren (AC d))))).
                + intros.
                  subst.
                  simpl in H3.
                  unfold B_GetContainerUsed in H3.
                  remember (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                            Z.of_nat (length (cchildren (AC d))) ) as child_id.
                  case_eq (zeq child_id (CPU_ID (update init_shared_adt l) + 1)).
                  intros; subst.
                  rewrite e in H3.
                  rewrite zeq_true in H3.
                  inv H3.

                  intros.
                  rewrite zeq_false in H3; [ | try assumption].
                  case_eq (zle_le 0 child_id 8); intros; subst.
                  rewrite <- CPU_ID_valid in ×.
                  generalize current_CPU_ID_range.
                  generalize ch_id_range.
                  intros.
                  remember (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) × 8 + 1 +
                            Z.of_nat (length (cchildren (AC d))) ) as child_id.
                  elim n0.
                  generalize ch_id_range, a0.
                  clear.
                  intros.
                  omega.

                  rewrite H20 in H3.
                  simpl in H3.
                  rewrite zeq_true in H3.
                  rewrite zeq_true in H3.
                  inv H3.

                + intros.
                  rewrite <- AC_re.
                  rewrite ZMap.gso; auto.
                  apply uctxt_used_inv.
                  unfold B_GetContainerUsed´.
                  rewrite H23.
                  exploit (single_big_proc_create_preserves_cused_snd tid (uRData l, d)); eauto.
                  { rewrite H4.
                    simpl; auto. }
                  { simpl.
                    rewrite <- H4.
                    assumption. }
              - unfold valid_AT_log_type_B.
                intros.
                inv Hdef.
                destruct big_oracle_inv as (Htemp1 & Htemp2 & Htemp3 & Htemp4).
                unfold valid_AT_oracle_B in Htemp3.
                generalize (valid_AT_log_inv _ H23); intros Hvalid_log.
                generalize (Htemp3 l0 (CPU_ID (update init_shared_adt l)) Hvalid_log); intros Hvalid_oracle_log.
                unfold valid_AT_log_B in ×.
                destruct Hvalid_oracle_log as (? & Hvalid_oracle_log).
                unfold B_CalAT_log_real in *; simpl.
                rewrite Hvalid_oracle_log; esplit; auto.
              - unfold valid_TCB_log_type_B.
                intros.
                inv Hdef.
                revert H26.
                rewrite H22 in sh_cid_re.
                rewrite sh_cid_re.
                rewrite sh_CPU_ID_re.
                unfold valid_TCB_log_B.
                unfold B_CalTCB_log_real.
                clear.
                intros; subst.
                subdestruct.
                esplit; auto.
              - intros.
                destruct big_oracle_inv as (Htemp1 & Htemp2 & Htemp3 & Htemp4).
                unfold valid_AT_oracle_B in Htemp3.
                unfold valid_AT_log_type_B in valid_AT_log_inv.
                generalize (valid_AT_log_inv l0 H23); intros Hvalid_AT_log.
                generalize (Htemp3 _ (CPU_ID (update init_shared_adt l)) Hvalid_AT_log);
                  intros Hvalid_AT_oracle_res.
                unfold valid_AT_log_type_B in valid_AT_log_inv.
                unfold valid_AT_log_B in valid_AT_log_inv.
                destruct Hvalid_AT_log as (at0 & Hvalid_AT_log).
                eapply (valid_B_Cal_AT_log_inv i0 l0 at0) in H3; eauto.
                inv H16.
                unfold B_CalAT_log_real in Hvalid_AT_log, H19.
                simpl in H19.
                case_eq (B_CalAT_log (big_oracle (update init_shared_adt l)
                                                 (CPU_ID (update init_shared_adt l)) l0 ++ l0)
                                     (real_AT (ZMap.init ATUndef))); intros.
                rewrite H16 in H19.
                inv H19.
                eapply (B_CalAT_log_keeps_allocated_pages _ _ _ _ _ Hvalid_AT_log H16 i0 H3); eauto.
                rewrite H16 in H19; inv H19.
                rewrite <- sh_big_log_re.
                auto. }
            { inv Hrel.
              inv sh_mem_inv_re.
              constructor; simpl; intros.
              case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l))
                                        (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H16.
                rewrite ZMap.gso in H19; eauto.
                inv H16; simpl in H20.
                eauto.
              - rewrite ZMap.gso in H16; auto.
                case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                         (cid (update init_shared_adt l)))); intros; subst.
                + rewrite ZMap.gss in H19; eauto.
                  inv H19.
                  simpl.
                  eapply pperm_disjoint; eauto.
                + rewrite ZMap.gso in H19; eauto. }
            × simpl.
              rewrite <- AC_re.
              rewrite H4.
              rewrite H4 in Hproc_kctxt.
              simpl in Hproc_kctxt.
              unfold prim_thread_init_pc in Hproc_kctxt.
              rewrite peq_true in Hproc_kctxt; auto. }
          Opaque proc_id uRData update.

        +
          assert (in_event: has_event acquire_lock_CHAN = true) by (unfold has_event; auto); rewrite in_event; clear in_event.
          assert (prim_num: prim_id_num acquire_lock_CHAN = 6) by (unfold prim_id_num; simpl; auto).
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          unfold initial_thread_kctxt.
          unfold prim_thread_init_pc.
          rewrite peq_false; [ | intro contra; inv contra].
          inversion Hlval; subst.
          inv H5; inv H7.
          subdestruct; eauto.
          destruct BUILTIN_ENABLED as (Hchoice_check & Hl).
          Transparent proc_id uRData update.
          exploit acquire_lock_CHAN_related; eauto; simpl in *; auto.
          subst.
          intros Hrelated.
          destruct Hrelated as ( & Hrelated_a & Hrelated_b & Hrelated_c & Hrelated_d & Hrelated_e).
           ; subst.
          refine_split´; eauto.
          simpl; econstructor; [simpl | simpl; eauto].
           (gensem big_acquire_lock_SC_spec); split; [reflexivity |].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor.
          eauto.
          Opaque proc_id uRData update.

        +
          assert (in_event: has_event release_lock_CHAN = true) by (unfold has_event; auto); rewrite in_event; clear in_event.
          assert (prim_num: prim_id_num release_lock_CHAN = 7) by (unfold prim_id_num; simpl; auto).
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          inv H9.
          unfold initial_thread_kctxt.
          unfold prim_thread_init_pc.
          rewrite peq_false; [ | intro contra; inv contra].
          inversion Hlval; subst.
          inv H5; inv H7.
          subdestruct; eauto.
          destruct BUILTIN_ENABLED as (Hchoice_check & Hl).
          Transparent proc_id uRData update.
          exploit release_lock_CHAN_related; eauto; simpl in *; auto.
          subst.
          intros Hrelated.
          destruct Hrelated as ( & Hrelated_a & Hrelated_b & Hrelated_c & Hrelated_d & Hrelated_e & Hrelated_f).
           ; subst.
          refine_split´; eauto.
          simpl; econstructor; [simpl | simpl; eauto].
           (gensem big_release_lock_SC_spec); split; [reflexivity |].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor.
          eauto.
          Opaque proc_id uRData update.

        +
          assert (in_event: has_event ipc_send_body = true) by (unfold has_event; auto); rewrite in_event; clear in_event.
          assert (prim_num: prim_id_num ipc_send_body = 8) by (unfold prim_id_num; simpl; auto).
          destruct Hsem as [s [sem [? [Hsig [Hsem [? ?]]]]]]; subst.
          inv Hsem; simpl in *; inv Hsig.
          destruct args; inv H6.
          destruct args; inv H9.
          destruct args; inv H8.
          inv H9.
          unfold initial_thread_kctxt.
          unfold prim_thread_init_pc.
          rewrite peq_false; [ | intro contra; inv contra].
          inversion Hlval; subst.
          inv H5; inv H7.
          inv H5; inv H8.
          inv H5; inv H7.
          subdestruct; eauto.
          destruct BUILTIN_ENABLED as (Hchoice_check & Hl).
          Transparent proc_id uRData update.
          exploit big_ipc_send_body_related; eauto; simpl in *; auto.
          subst.
          intros Hrelated.
          destruct Hrelated as ( & res´ & Hrelated_a & Hrelated_b & Hrelated_c & Hrelated_d & Hrelated_e & Hrelated_f).
           ; subst.
          refine_split´; eauto.
          simpl; econstructor; [simpl | simpl; eauto].
           (gensem big_ipc_send_body_spec); split; [reflexivity |].
           s; refine_split´; auto; try reflexivity; simpl.
          repeat constructor.
          eauto.
          Opaque proc_id uRData update.

      - destruct ef; simpl; try inv BUILTIN_ENABLED.
        inv Hcase.
    Fail idtac.
    Qed.
    Next Obligation.
      rename H into name.
      rename H1 into kctxt_pool.
      rename H3 into Hrel.
      case (peq name thread_yield); intros; subst;
      [rewrite peq_true in BUILTIN_ENABLED; contradiction | ].
      rewrite peq_false in BUILTIN_ENABLED; try contradiction; auto.
      case (peq name thread_sleep); intros; subst;
      [rewrite peq_true in BUILTIN_ENABLED; contradiction | ].
      rewrite peq_false in BUILTIN_ENABLED; try contradiction; auto.
      inv_layer H7; try (inv H8; fail).


      -
        assert (in_event: has_event vmx_run_vm = false) by (unfold has_event; auto).
        rewrite in_event in *; subst.
        inv H8; subst.
        inv H3; subst.
        unfold ObjPHBThreadINTELVIRT.single_vm_run_spec in H8; subdestruct; simpl in *; subst.
        Transparent update uRData proc_id.
        generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
        simpl in *; rewrite H2 in Hper_data_rel.
        inv Hper_data_rel; rewrite zeq_true in ×.
         (((a {vmx: ZMap.set (CPU_ID (update init_shared_adt l))
                                   (ZMap.set 34 (((Pregmap.init Vundef) # EDI <- (rs EDI)) # EBP <- (rs EBP) EDI)
                                             (ZMap.set 33 (((Pregmap.init Vundef) # EDI <- (rs EDI)) # EBP <- (rs EBP) EBP)
                                                       (vmx d))) (AbstractDataType.vmx a)})
                   {vmcs : ZMap.set (CPU_ID (update init_shared_adt l))
                                    (VMCSValid revid abrtid (ZMap.set 26654 (Vint i) data))
                                    (AbstractDataType.vmcs a)}) {ihost : false}).
        refine_split´; auto; simpl.
        + instantiate (1 := IntelPrimSemantics.primcall_vmx_enter_compatsem vm_run_spec vmx_run_vm).
          eapply get_layer_primitive_left; auto; decision.
        + econstructor; eauto; simpl.
          econstructor; eauto; simpl.
          unfold vm_run_spec.
          inv H8.
          rewrite <- ikern_re, <- ihost_re, <- pg_re, <- init_re.
          rewrite e in vmcs_re, vmx_re, ept_re.
          rewrite zeq_true in ept_re, vmcs_re, vmx_re.
          rewrite <- vmcs_re, <- vmx_re, <- CPU_ID_re.
          rewrite Hdestruct0, Hdestruct1, Hdestruct2, Hdestruct3, Hdestruct4, Hdestruct5, Hdestruct6.
          simpl; reflexivity.
        + inv H8; simpl; auto.
        + instantiate (1:= kctxt_pool).
          inv H8; simpl; auto.
          inv Hrel; simpl.
          constructor; simpl; auto.
          intros.
          case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros.
          × subst; rewrite ZMap.gss.
            { per_data_auto_simpl.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - rewrite Hdestruct3 in ×.
                auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - auto.
              - auto.
              - auto.
              - auto.
                rewrite e in ×.
                rewrite zeq_true in ×.
                simpl in ×.
                rewrite <- CPU_ID_re.
                rewrite ZMap.gss.
                reflexivity.
              - auto.
                rewrite e in ×.
                rewrite zeq_true in ×.
                simpl in ×.
                rewrite <- CPU_ID_re.
                rewrite ZMap.gss.
                reflexivity.
              - inv inv_re; constructor; simpl; auto. }
          × generalize (per_data_re i0); intros.
            rewrite ZMap.gso; auto.
            repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                   end.
            subdestruct; eauto.
            inv H3; simpl in ×.
            { per_data_auto.
              - rewrite zeq_false; auto.
                rewrite zeq_false in ihost_re0; auto.
              - auto.
              - auto.
              - rewrite Hdestruct3 in ×.
                auto.
              - auto.
              - rewrite <- e.
                rewrite zeq_false; auto.
              - rewrite <- e.
                rewrite zeq_false; auto.
              - inv inv_re0; constructor; simpl; auto. }
            constructor.
          × inv sh_shared_inv_re; constructor; simpl; auto.
          × inv sh_mem_inv_re.
            constructor; simpl; intros.
            case_eq (zeq i0 (ZMap.get (CPU_ID (update init_shared_adt l))
                                      (cid (update init_shared_adt l)))); intros; subst.
            { rewrite ZMap.gss in H3.
              rewrite ZMap.gso in H5; eauto.
              inv H3; simpl in H6.
              eauto. }
            { rewrite ZMap.gso in H3; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H5; eauto.
                inv H5.
                simpl.
                eapply pperm_disjoint; eauto.
              - rewrite ZMap.gso in H5; eauto. }
            Opaque proc_id uRData update.

      -
        assert (in_event: has_event vmx_return_from_guest = false) by (unfold has_event; auto).
        rewrite in_event in *; subst.
        inv H8; subst.
        inv H3; subst.
        unfold ObjPHBThreadINTELVIRT.single_vmx_return_from_guest_spec in H8; subdestruct; simpl in *; subst.
        Transparent update uRData proc_id.
        generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
        simpl in *; rewrite H2 in Hper_data_rel.
        inv Hper_data_rel; rewrite zeq_true in ×.
         (a {ihost : true}
                  {vmx : ZMap.set (CPU_ID (update init_shared_adt l))
                                  (ZMap.set 31 Vone (ZMap.set 28 (Vint i1) (ZMap.set 27 (Vint i0)
                                                                                     (ZMap.set 14 (Vint i) (vmx d)))))
                                  (AbstractDataType.vmx a)}).
        refine_split´; auto; simpl.
        + instantiate (1 := IntelPrimSemantics.primcall_vmx_return_compatsem vmx_return_from_guest_spec).
          eapply get_layer_primitive_left; auto; decision.
        + econstructor; eauto; simpl.
          econstructor; eauto; simpl.
          unfold vmx_return_from_guest_spec.
          inv H8.
          rewrite <- ikern_re, <- ihost_re, <- pg_re, <- init_re.
          rewrite e in vmcs_re, vmx_re, ept_re.
          rewrite zeq_true in ept_re, vmcs_re, vmx_re.
          rewrite <- vmcs_re, <- vmx_re, <- CPU_ID_re.
          rewrite Hdestruct0, Hdestruct1, Hdestruct2, Hdestruct3, Hdestruct4, Hdestruct5, Hdestruct6, Hdestruct7, Hdestruct8.
          simpl; reflexivity.
        + inv H8; simpl; auto.
        + instantiate (1:= kctxt_pool).
          inv H8; simpl; auto.
          inv Hrel; simpl.
          constructor; simpl; auto.
          intros.
          case_eq (zeq i2 (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros.
          × subst; rewrite ZMap.gss.
            { per_data_auto_simpl.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - rewrite Hdestruct3 in ×.
                auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
                rewrite e in ×.
                rewrite zeq_true in ×.
                simpl in ×.
                rewrite <- CPU_ID_re.
                rewrite ZMap.gss.
                reflexivity.
              - inv inv_re; constructor; simpl; auto. }
          × generalize (per_data_re i2); intros.
            rewrite ZMap.gso; auto.
            repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                   end.
            subdestruct; eauto.
            inv H3; simpl in ×.
            { per_data_auto.
              - rewrite zeq_false; auto.
                rewrite zeq_false in ihost_re0; auto.
              - auto.
              - auto.
              - rewrite Hdestruct3 in ×.
                auto.
              - auto.
              - rewrite <- e.
                rewrite zeq_false; auto.
              - inv inv_re0; constructor; simpl; auto. }
            constructor.
          × inv sh_shared_inv_re; constructor; simpl; auto.
          × inv sh_mem_inv_re.
            constructor; simpl; intros.
            case_eq (zeq i2 (ZMap.get (CPU_ID (update init_shared_adt l))
                                      (cid (update init_shared_adt l)))); intros; subst.
            { rewrite ZMap.gss in H3.
              rewrite ZMap.gso in H5; eauto.
              inv H3; simpl in H6.
              eauto. }
            { rewrite ZMap.gso in H3; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H5; eauto.
                inv H5.
                simpl.
                eapply pperm_disjoint; eauto.
              - rewrite ZMap.gso in H5; eauto. }
            Opaque proc_id uRData update.



      -
        assert (in_event: has_event host_in = false) by (unfold has_event; auto).
        rewrite in_event in *; subst.
        inv H8; subst.
        inv H3; subst.
        functional inversion H8; subst.
        Transparent update uRData proc_id.
        generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
        simpl in *; rewrite H2 in Hper_data_rel.
        inv Hper_data_rel; rewrite zeq_true in ×.
         (a {ihost: true}).
        refine_split´; auto; simpl.
        + instantiate (1 := primcall_general_compatsem hostin_spec).
          eapply get_layer_primitive_left; auto; decision.
        + econstructor; eauto; simpl.
          econstructor; eauto; simpl.
          unfold hostin_spec; simpl; rewritesb; simpl in ×.
          rewrite H5, H6; reflexivity.
        + instantiate (1:= kctxt_pool).
          inv Hrel; simpl.
          constructor; simpl; auto.
          intros.
          case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros.
          × subst; rewrite ZMap.gss.
            { per_data_auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - auto.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
          × generalize (per_data_re i); intros.
            rewrite ZMap.gso; auto.
            repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                   end.
            subdestruct; eauto; inv H3; simpl in ×.
            { per_data_auto.
              - rewrite zeq_false; auto.
                rewrite zeq_false in ihost_re0; auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - inv inv_re0; constructor; simpl; auto. }
            constructor.
          × inv sh_shared_inv_re; constructor; simpl; auto.
          × inv sh_mem_inv_re.
            constructor; simpl; intros.
            case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)))); intros; subst.
            { rewrite ZMap.gss in H3.
              rewrite ZMap.gso in H9; eauto.
              inv H3; simpl in H10.
              eauto. }
            { rewrite ZMap.gso in H3; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H9; eauto.
                inv H9.
                simpl.
                eapply pperm_disjoint; eauto.
              - rewrite ZMap.gso in H9; eauto. }
            Opaque proc_id uRData update.

      -
        assert (in_event: has_event host_out = false) by (unfold has_event; auto).
        rewrite in_event in *; subst.
        inv H8; subst.
        inv H3; subst.
        functional inversion H8; subst.
        Transparent update uRData proc_id.
        generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
        simpl in *; rewrite H2 in Hper_data_rel.
        inv Hper_data_rel; rewrite zeq_true in ×.
         (a {ihost: false}).
        refine_split´; auto; simpl.
        + instantiate (1 := primcall_general_compatsem hostout_spec).
          eapply get_layer_primitive_left; auto; decision.
        + econstructor; eauto; simpl.
          econstructor; eauto; simpl.
          unfold hostout_spec; simpl; rewritesb; simpl in ×.
          rewrite H5, H6, H7, H9; reflexivity.
        + instantiate (1:= kctxt_pool).
          inv Hrel; simpl.
          constructor; simpl; auto.
          intros.
          case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros.
          × subst; rewrite ZMap.gss.
            { per_data_auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - auto.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
          × generalize (per_data_re i); intros.
            rewrite ZMap.gso; auto.
            repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                   end.
            subdestruct; eauto; inv H3; simpl in ×.
            { per_data_auto.
              - rewrite zeq_false; auto.
                rewrite zeq_false in ihost_re0; auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - auto.
              - inv inv_re0; constructor; simpl; auto. }
            constructor.
          × inv sh_shared_inv_re; constructor; simpl; auto.
          × inv sh_mem_inv_re.
            constructor; simpl; intros.
            case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)))); intros; subst.
            { rewrite ZMap.gss in H3.
              rewrite ZMap.gso in H10; eauto.
              inv H3; simpl in H11.
              eauto. }
            { rewrite ZMap.gso in H3; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H10; eauto.
                inv H10.
                simpl.
                eapply pperm_disjoint; eauto.
              - rewrite ZMap.gso in H10; eauto. }
            Opaque proc_id uRData update.

      -
        assert (in_event: has_event trap_get = false) by (unfold has_event; auto).
        rewrite in_event in *; subst.
        inv H8; subst.
        inv H3; subst.
        Transparent update uRData proc_id.
        generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
        simpl in *; rewrite H2 in Hper_data_rel.
        inv Hper_data_rel; rewrite zeq_true in ×.
        refine_split´; auto.
        + instantiate (1:= primcall_trap_info_get_compatsem trap_info_get_spec).
          eapply get_layer_primitive_left; auto; decision.
        + econstructor; eauto; simpl.
          econstructor; eauto; simpl.
          unfold ObjPHBThreadDEV.single_trap_info_get_spec.
          unfold trap_info_get_spec; simpl.
          subst.
          rewrite ti_re; auto.
        + instantiate (1:= kctxt_pool).
          eapply relate_RData_gss in Hrel; eauto.
          Opaque update uRData proc_id.

      -
        assert (in_event: has_event trap_set = false) by (unfold has_event; auto).
        rewrite in_event in *; subst.
        inv H8; subst.
        inv H3; subst.
        Transparent update uRData proc_id.
        generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
        simpl in *; rewrite H2 in Hper_data_rel.
        inv Hper_data_rel; rewrite zeq_true in ×.
        refine_split´; auto.
        + instantiate (1:= primcall_trap_info_ret_compatsem trap_info_ret_spec).
          eapply get_layer_primitive_left; auto; decision.
        + econstructor; eauto; simpl.
          econstructor; eauto; simpl.
          unfold ObjPHBThreadDEV.single_trap_info_ret_spec.
          unfold trap_info_ret_spec; simpl.
          subst.
          rewrite ti_re; auto.
        + instantiate (1:= kctxt_pool).
          eapply relate_RData_gss in Hrel; eauto.
          Opaque update uRData proc_id.

      -
        assert (in_event: has_event proc_start_user = false) by (unfold has_event; auto).
        rewrite in_event in *; subst.
        inv H8; subst.
        inv H3; subst.
        Transparent update uRData proc_id.
        generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
        simpl in ×.
        rewrite H2 in Hper_data_rel.
        functional inversion H8.
        assert (Hproc_id: Constant.TOTAL_CPU <
                          ZMap.get (CPU_ID (update init_shared_adt l))
                                   (cid (update init_shared_adt l)) < Constant.num_proc
                          ZMap.get (CPU_ID (update init_shared_adt l))
                                   (cid (update init_shared_adt l)) = main_thread).
        { eapply valid_thread_list.
          exploit all_cid_in_full_thread_list.
          simpl.
          instantiate (1 := l); intro; auto. }
         (((a {ikern : false}) {ipt : false}) {PT : ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)}).
        refine_split´; auto.
        + instantiate (1:= primcall_start_user_compatsem proc_start_user_spec).
          eapply get_layer_primitive_left; auto; decision.
        + inv Hper_data_rel; rewrite zeq_true in ×.
          econstructor; eauto; simpl.
          econstructor; eauto; simpl.
          unfold proc_start_user_spec.
          rewritesb.
          unfold pv_adt in *; simpl in ×.
          rewrite H6, H7, H9, H10, H11.
          f_equal.
          f_equal.
          assert (uctxt d = ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l))
                                                (cid (update init_shared_adt l)))
                                      (AbstractDataType.uctxt a)).
          { unfold relate_uctxt_per_pd in uctxt_re.
            unfold B_GetContainerUsed´ in uctxt_re.
            unfold curid, cpu in H13; rewrite H6, H12, H13 in uctxt_re.
            { assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
              { rewrite main_thread_val.
                generalize current_CPU_ID_range.
                clear.
                intros.
                omega. }
              destruct Hproc_id as [Hproc_id | Hproc_id]; auto.
              + assert (Hno_main: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) main_thread).
                { revert Hproc_id.
                  revert Hmain_thread.
                  clear; intros.
                  omega. }
                rewrite zeq_false in uctxt_re; try assumption.
                rewrite zlt_lt_true in uctxt_re; try assumption.
                unfold uctxt_bieq in ×.
                destruct (ZMap.get 12 (uctxt d)) eqn:?; try contradiction.
                elim HUEIP_val_cond2; auto.
                case_eq (Int.eq i (Int.repr 0)); intros; subst; auto.
                elim HUEIP_val_cond1.
                f_equal.
                unfold Int.eq in H3.
                destruct ( zeq (Int.unsigned i) (Int.unsigned (Int.repr 0))).
                apply Int_unsigned_eq in e.
                auto.
                inv H3.
                rewrite H3 in uctxt_re.
                auto.
              + rewrite Hproc_id in uctxt_re.
                rewrite zeq_true in uctxt_re; try assumption.
                rewrite Hproc_id; auto.
                unfold uctxt_bieq in ×.
                destruct (ZMap.get 12 (uctxt d)) eqn:?; try contradiction.
                elim HUEIP_val_cond2; auto.
                case_eq (Int.eq i (Int.repr 0)); intros; subst; auto.
                elim HUEIP_val_cond1.
                f_equal.
                unfold Int.eq in H3.
                destruct ( zeq (Int.unsigned i) (Int.unsigned (Int.repr 0))).
                apply Int_unsigned_eq in e.
                auto.
                inv H3.
                rewrite H3 in uctxt_re.
                auto. } }
          rewrite H6 in ×.
          rewrite CPU_ID_re.
          rewrite <- cid_re.
          congruence.
        + instantiate (1:= kctxt_pool).
          eapply relate_RData_gss in Hrel; eauto.
          inv Hrel; simpl.
          constructor; simpl; auto.
          intros.
          case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros.
          × subst; rewrite ZMap.gss.
            inv Hper_data_rel.
            unfold pv_adt; simpl in ×.
            { rewrite H6 in *; rewrite zeq_true in ×.
              per_data_auto_simpl.
              - rewrite H6; auto.
              - rewrite H6; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_AC_per_pd in *; simpl; auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite H6; unfold curid; rewrite zeq_true; auto.
              - rewrite H6.
                auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl; auto.
              - unfold relate_uctxt_per_pd in *; simpl; auto.
              - auto.
              - auto.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
          × generalize (per_data_re i).
            rewrite ZMap.gso; auto.
            rewrite ZMap.gso; auto; intros.
            repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                   end.
            subdestruct; eauto; inv H5; simpl in ×.
            { rewrite H6 in ×.
              per_data_auto_simpl.
              - rewrite H6; auto.
              - rewrite H6; auto.
              - rewrite zeq_false; auto.
                rewrite zeq_false in ikern_re; auto.
              - unfold relate_AC_per_pd in *; simpl; auto.
              - auto.
              - auto.
              - rewrite H6; rewrite zeq_false; auto.
                rewrite zeq_false in PT_re; auto.
              - rewrite H6; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in ipt_re; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in intr_flag_re; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in in_intr_re; auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl; auto.
              - unfold relate_uctxt_per_pd in *; simpl; auto.
              - auto.
              - auto.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
            constructor.
          × inv sh_shared_inv_re; constructor; simpl; auto.
          × inv sh_mem_inv_re.
            constructor; simpl; intros.
            case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)))); intros; subst.
            { rewrite ZMap.gss in H3.
              rewrite ZMap.gso in H5; eauto.
              inv H3.
              simpl in H14.
              unfold pv_adt in H14; simpl in H14.
              eapply pperm_disjoint; eauto.
              rewrite ZMap.gss; auto.
              rewrite ZMap.gso; auto. }
            { rewrite ZMap.gso in H3; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H5; eauto.
                inv H5.
                simpl.
                eapply pperm_disjoint; eauto.
                rewrite ZMap.gso; auto.
                rewrite ZMap.gss; auto.
              - rewrite ZMap.gso in H5; eauto.
                eapply pperm_disjoint in H; eauto.
                rewrite ZMap.gso; auto.
                rewrite ZMap.gso; auto. }
            Opaque proc_id uRData update.

      -
        assert (in_event: has_event proc_exit_user = false) by (unfold has_event; auto).
        rewrite in_event in *; subst.
        inv H8; subst.
        inv H3; subst.
        Transparent update uRData proc_id.
        generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
        simpl in ×.
        rewrite H2 in Hper_data_rel.
        functional inversion H8.
        assert (Hproc_id: Constant.TOTAL_CPU <
                          ZMap.get (CPU_ID (update init_shared_adt l))
                                   (cid (update init_shared_adt l)) < Constant.num_proc
                          ZMap.get (CPU_ID (update init_shared_adt l))
                                   (cid (update init_shared_adt l)) = main_thread).
        { eapply valid_thread_list.
          exploit all_cid_in_full_thread_list.
          simpl.
          instantiate (1 := l); intro; auto. }
        assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
        { rewrite main_thread_val.
          generalize current_CPU_ID_range.
          clear.
          intros.
          omega. }
         ((((a {ikern : true}) {PT : 0}) {ipt : true})
                  {uctxt: ZMap.set (ZMap.get (CPU_ID (update init_shared_adt l))
                                             (cid (update init_shared_adt l))) uctx4 (AbstractDataType.uctxt a)}).
        esplit; refine_split´; auto.
        + instantiate (1:= primcall_exit_user_compatsem proc_exit_user_spec).
          eapply get_layer_primitive_left; auto; decision.
        + inv Hper_data_rel; rewrite zeq_true in ×.
          econstructor; eauto; simpl.
          econstructor; eauto; simpl.
          unfold proc_exit_user_spec.
          rewritesb.
          unfold pv_adt in *; simpl in ×.
          rewrite H5, H6, H7, H9.
          fold uctx1; fold uctx2.
          fold uctx3; fold uctx4.
          rewrite H9 in cid_re.
          rewrite CPU_ID_re.
          rewrite <- cid_re.
          rewrite <- CPU_ID_re.
          reflexivity.
        + instantiate (1:= kctxt_pool).
          inv Hrel; simpl.
          constructor; simpl; auto.
          intros.
          case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros.
          { subst; rewrite ZMap.gss; simpl in ×.
            inv Hper_data_rel.
            rewrite H9 in *; rewrite zeq_true in ×.
            per_data_auto_simpl.
            - rewrite H9; auto.
            - rewrite H9; auto.
            - rewrite zeq_true; auto.
            - rewrite zeq_true; auto.
            - unfold relate_AC_per_pd in *; simpl; auto.
            - rewrite zeq_true; auto.
            - auto.
            - rewrite H9; rewrite zeq_true; auto.
            - rewrite H9.
              auto.
            - rewrite zeq_true; auto.
            - rewrite zeq_true; auto.
            - rewrite zeq_true; auto.
            - unfold relate_SyncChanPool_per_pd in *; simpl; auto.
            - unfold relate_uctxt_per_pd in *; simpl; auto.
              rewrite H9.
              unfold B_GetContainerUsed´; unfold cpu, curid in H12; rewrite H11, H12.
              rewrite ZMap.gss; auto.
              destruct (zeq (ZMap.get (CPU_ID (update init_shared_adt l))
                                      (cid (update init_shared_adt l))) main_thread); auto.
              unfold uctxt_bieq.
              unfold uctx4.
              unfold uctx3.
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gss.
              case_eq (Int.eq v12 (Int.repr 0)); intros.

              elim H14.
              unfold Int.eq in H.
              destruct (zeq (Int.unsigned v12) (Int.unsigned (Int.repr 0))).
              apply Int_unsigned_eq in e0.
              auto.
              inv H.

              reflexivity.

              destruct (zlt_lt 8 (ZMap.get (CPU_ID (update init_shared_adt l))
                                           (cid (update init_shared_adt l))) 1024); auto.
              unfold uctxt_bieq.
              unfold uctx4.
              unfold uctx3.
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gss.
              case_eq (Int.eq v12 (Int.repr 0)); intros.

              elim H14.
              unfold Int.eq in H.
              destruct (zeq (Int.unsigned v12) (Int.unsigned (Int.repr 0))).
              apply Int_unsigned_eq in e.
              auto.
              inv H.

              reflexivity.

            - auto.
            - auto.
            - auto.
            - inv inv_re; constructor; simpl; auto. }
          × generalize (per_data_re i).
            rewrite ZMap.gso; auto; intros.
            repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                   end.
            subdestruct; eauto; inv H13; simpl in ×.
            { rewrite H9 in ×.
              per_data_auto_simpl.
              - rewrite H9; auto.
              - rewrite H9; auto.
              - rewrite zeq_false; auto.
                rewrite zeq_false in ikern_re; auto.
              - unfold relate_AC_per_pd in *; simpl; auto.
              - auto.
              - auto.
              - rewrite H9; rewrite zeq_false; auto.
                rewrite zeq_false in PT_re; auto.
              - rewrite H9; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in ipt_re; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in intr_flag_re; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in in_intr_re; auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl; auto.
              - unfold relate_uctxt_per_pd in ×.
                rewrite H9 in ×.
                simpl.
                rewrite ZMap.gso; auto.
              - auto.
              - auto.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
            constructor.
          × inv sh_shared_inv_re; constructor; simpl; auto.
            { simpl in H9.
              rewrite H9.
              intros contra; inv contra. }
            
            { intros.
              case_eq (zeq tid (ZMap.get (CPU_ID (update init_shared_adt l))
                                         (cid (update init_shared_adt l)))); intros; subst.
              - unfold B_GetContainerUsed´ in H.
                simpl in H11; rewrite H11 in H.
                unfold curid, cpu in H12.
                simpl in H12.
                rewrite H12 in H.
                inv H.
              - rewrite ZMap.gso; auto. }
          × inv sh_mem_inv_re.
            constructor; simpl; intros.
            case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)))); intros; subst.
            { rewrite ZMap.gss in H3.
              rewrite ZMap.gso in H13; eauto.
              inv H3; simpl in H15.
              unfold pv_adt in H15; simpl in H15.
              eapply pperm_disjoint; eauto. }
            { rewrite ZMap.gso in H3; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H13; eauto.
                inv H13.
                simpl.
                eapply pperm_disjoint; eauto.
              - rewrite ZMap.gso in H13; eauto. }
            Opaque proc_id uRData update.

      -
        assert (in_event: has_event proc_start_user2 = false) by (unfold has_event; auto).
        rewrite in_event in *; subst.
        inv H8; subst.
        inv H3; subst.
        Transparent update uRData proc_id.
        generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
        simpl in ×.
        rewrite H2 in Hper_data_rel.
        functional inversion H8.
        assert (Hproc_id: Constant.TOTAL_CPU <
                          ZMap.get (CPU_ID (update init_shared_adt l))
                                   (cid (update init_shared_adt l)) < Constant.num_proc
                          ZMap.get (CPU_ID (update init_shared_adt l))
                                   (cid (update init_shared_adt l)) = main_thread).
        { eapply valid_thread_list.
          exploit all_cid_in_full_thread_list.
          simpl.
          instantiate (1 := l); intro; auto. }
         (((a {ikern : false}) {ipt : false}) {PT : ZMap.get (AbstractDataType.CPU_ID a) (AbstractDataType.cid a)}).
        refine_split´; auto.
        + instantiate (1:= primcall_start_user_compatsem2 proc_start_user_spec).
          eapply get_layer_primitive_left; auto; decision.
        + inv Hper_data_rel; rewrite zeq_true in ×.
          econstructor; eauto; simpl.
          econstructor; eauto; simpl.
          unfold proc_start_user_spec.
          rewritesb.
          unfold pv_adt in *; simpl in ×.
          rewrite H6, H7, H9, H10, H11.
          f_equal.
          f_equal.
          assert (uctxt d = ZMap.get (ZMap.get (CPU_ID (update init_shared_adt l))
                                                (cid (update init_shared_adt l)))
                                      (AbstractDataType.uctxt a)).
          { unfold relate_uctxt_per_pd in uctxt_re.
            unfold B_GetContainerUsed´ in uctxt_re.
            unfold curid, cpu in H13; rewrite H6, H12, H13 in uctxt_re.
            { assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
              { rewrite main_thread_val.
                generalize current_CPU_ID_range.
                clear.
                intros.
                omega. }
              destruct Hproc_id as [Hproc_id | Hproc_id]; auto.
              + assert (Hno_main: ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)) main_thread).
                { revert Hproc_id.
                  revert Hmain_thread.
                  clear; intros.
                  omega. }
                rewrite zeq_false in uctxt_re; try assumption.
                rewrite zlt_lt_true in uctxt_re; try assumption.
                unfold uctxt_bieq in ×.
                destruct (ZMap.get 12 (uctxt d)) eqn:?; try contradiction.
                elim HUEIP_val_cond2; auto.
                case_eq (Int.eq i (Int.repr 0)); intros; subst; auto.
                elim HUEIP_val_cond1.
                f_equal.
                unfold Int.eq in H3.
                destruct ( zeq (Int.unsigned i) (Int.unsigned (Int.repr 0))).
                apply Int_unsigned_eq in e.
                auto.
                inv H3.
                rewrite H3 in uctxt_re.
                auto.
              + rewrite Hproc_id in uctxt_re.
                rewrite zeq_true in uctxt_re; try assumption.
                rewrite Hproc_id; auto.
                unfold uctxt_bieq in ×.
                destruct (ZMap.get 12 (uctxt d)) eqn:?; try contradiction.
                elim HUEIP_val_cond2; auto.
                case_eq (Int.eq i (Int.repr 0)); intros; subst; auto.
                elim HUEIP_val_cond1.
                f_equal.
                unfold Int.eq in H3.
                destruct ( zeq (Int.unsigned i) (Int.unsigned (Int.repr 0))).
                apply Int_unsigned_eq in e.
                auto.
                inv H3.
                rewrite H3 in uctxt_re.
                auto. } }
          rewrite H6 in ×.
          rewrite CPU_ID_re.
          rewrite <- cid_re.
          congruence.
        + instantiate (1:= kctxt_pool).
          eapply relate_RData_gss in Hrel; eauto.
          inv Hrel; simpl.
          constructor; simpl; auto.
          intros.
          case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros.
          × subst; rewrite ZMap.gss.
            inv Hper_data_rel.
            unfold pv_adt; simpl in ×.
            { rewrite H6 in *; rewrite zeq_true in ×.
              per_data_auto_simpl.
              - rewrite H6; auto.
              - rewrite H6; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_AC_per_pd in *; simpl; auto.
              - rewrite zeq_true; auto.
              - auto.
              - rewrite H6; unfold curid; rewrite zeq_true; auto.
              - rewrite H6.
                auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - rewrite zeq_true; auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl; auto.
              - unfold relate_uctxt_per_pd in *; simpl; auto.
              - auto.
              - auto.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
          × generalize (per_data_re i).
            rewrite ZMap.gso; auto.
            rewrite ZMap.gso; auto; intros.
            repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                   end.
            subdestruct; eauto; inv H5; simpl in ×.
            { rewrite H6 in ×.
              per_data_auto_simpl.
              - rewrite H6; auto.
              - rewrite H6; auto.
              - rewrite zeq_false; auto.
                rewrite zeq_false in ikern_re; auto.
              - unfold relate_AC_per_pd in *; simpl; auto.
              - auto.
              - auto.
              - rewrite H6; rewrite zeq_false; auto.
                rewrite zeq_false in PT_re; auto.
              - rewrite H6; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in ipt_re; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in intr_flag_re; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in in_intr_re; auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl; auto.
              - unfold relate_uctxt_per_pd in *; simpl; auto.
              - auto.
              - auto.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
            constructor.
          × inv sh_shared_inv_re; constructor; simpl; auto.
          × inv sh_mem_inv_re.
            constructor; simpl; intros.
            case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)))); intros; subst.
            { rewrite ZMap.gss in H3.
              rewrite ZMap.gso in H5; eauto.
              inv H3.
              simpl in H14.
              unfold pv_adt in H14; simpl in H14.
              eapply pperm_disjoint; eauto.
              rewrite ZMap.gss; auto.
              rewrite ZMap.gso; auto. }
            { rewrite ZMap.gso in H3; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H5; eauto.
                inv H5.
                simpl.
                eapply pperm_disjoint; eauto.
                rewrite ZMap.gso; auto.
                rewrite ZMap.gss; auto.
              - rewrite ZMap.gso in H5; eauto.
                eapply pperm_disjoint in H; eauto.
                rewrite ZMap.gso; auto.
                rewrite ZMap.gso; auto. }
            Opaque proc_id uRData update.

      -
        assert (in_event: has_event proc_exit_user2 = false) by (unfold has_event; auto).
        rewrite in_event in *; subst.
        inv H8; subst.
        inv H3; subst.
        Transparent update uRData proc_id.
        generalize (per_data_re _ _ _ _ Hrel (proc_id (uRData l))); intros Hper_data_rel.
        simpl in ×.
        rewrite H2 in Hper_data_rel.
        functional inversion H8.
        assert (Hproc_id: Constant.TOTAL_CPU <
                          ZMap.get (CPU_ID (update init_shared_adt l))
                                   (cid (update init_shared_adt l)) < Constant.num_proc
                          ZMap.get (CPU_ID (update init_shared_adt l))
                                   (cid (update init_shared_adt l)) = main_thread).
        { eapply valid_thread_list.
          exploit all_cid_in_full_thread_list.
          simpl.
          instantiate (1 := l); intro; auto. }
        assert (Hmain_thread: 0 < main_thread Constant.TOTAL_CPU).
        { rewrite main_thread_val.
          generalize current_CPU_ID_range.
          clear.
          intros.
          omega. }
         ((((a {ikern : true}) {PT : 0}) {ipt : true})
                  {uctxt: ZMap.set (ZMap.get (CPU_ID (update init_shared_adt l))
                                             (cid (update init_shared_adt l))) uctx4 (AbstractDataType.uctxt a)}).
        esplit; refine_split´; auto.
        + instantiate (1:= primcall_exit_user_compatsem2 proc_exit_user_spec).
          eapply get_layer_primitive_left; auto; decision.
        + inv Hper_data_rel; rewrite zeq_true in ×.
          econstructor; eauto; simpl.
          econstructor; eauto; simpl.
          unfold proc_exit_user_spec.
          rewritesb.
          unfold pv_adt in *; simpl in ×.
          rewrite H5, H6, H7, H9.
          fold uctx1; fold uctx2.
          fold uctx3; fold uctx4.
          rewrite H9 in cid_re.
          rewrite CPU_ID_re.
          rewrite <- cid_re.
          rewrite <- CPU_ID_re.
          reflexivity.
        + instantiate (1:= kctxt_pool).
          inv Hrel; simpl.
          constructor; simpl; auto.
          intros.
          case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l)) (cid (update init_shared_adt l)))); intros.
          { subst; rewrite ZMap.gss; simpl in ×.
            inv Hper_data_rel.
            rewrite H9 in *; rewrite zeq_true in ×.
            per_data_auto_simpl.
            - rewrite H9; auto.
            - rewrite H9; auto.
            - rewrite zeq_true; auto.
            - rewrite zeq_true; auto.
            - unfold relate_AC_per_pd in *; simpl; auto.
            - rewrite zeq_true; auto.
            - auto.
            - rewrite H9; rewrite zeq_true; auto.
            - rewrite H9.
              auto.
            - rewrite zeq_true; auto.
            - rewrite zeq_true; auto.
            - rewrite zeq_true; auto.
            - unfold relate_SyncChanPool_per_pd in *; simpl; auto.
            - unfold relate_uctxt_per_pd in *; simpl; auto.
              rewrite H9.
              unfold B_GetContainerUsed´; unfold cpu, curid in H12; rewrite H11, H12.
              rewrite ZMap.gss; auto.
              destruct (zeq (ZMap.get (CPU_ID (update init_shared_adt l))
                                      (cid (update init_shared_adt l))) main_thread); auto.
              unfold uctxt_bieq.
              unfold uctx4.
              unfold uctx3.
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gss.
              case_eq (Int.eq v12 (Int.repr 0)); intros.

              elim H14.
              unfold Int.eq in H.
              destruct (zeq (Int.unsigned v12) (Int.unsigned (Int.repr 0))).
              apply Int_unsigned_eq in e0.
              auto.
              inv H.

              reflexivity.

              destruct (zlt_lt 8 (ZMap.get (CPU_ID (update init_shared_adt l))
                                           (cid (update init_shared_adt l))) 1024); auto.
              unfold uctxt_bieq.
              unfold uctx4.
              unfold uctx3.
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gso; [ | intro contra´; inv contra´].
              rewrite ZMap.gss.
              case_eq (Int.eq v12 (Int.repr 0)); intros.

              elim H14.
              unfold Int.eq in H.
              destruct (zeq (Int.unsigned v12) (Int.unsigned (Int.repr 0))).
              apply Int_unsigned_eq in e.
              auto.
              inv H.

              reflexivity.

            - auto.
            - auto.
            - auto.
            - inv inv_re; constructor; simpl; auto. }
          × generalize (per_data_re i).
            rewrite ZMap.gso; auto; intros.
            repeat match goal with
                     | [H: context[if _ then _ else _] |- _] ⇒ clear H
                   end.
            subdestruct; eauto; inv H13; simpl in ×.
            { rewrite H9 in ×.
              per_data_auto_simpl.
              - rewrite H9; auto.
              - rewrite H9; auto.
              - rewrite zeq_false; auto.
                rewrite zeq_false in ikern_re; auto.
              - unfold relate_AC_per_pd in *; simpl; auto.
              - auto.
              - auto.
              - rewrite H9; rewrite zeq_false; auto.
                rewrite zeq_false in PT_re; auto.
              - rewrite H9; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in ipt_re; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in intr_flag_re; auto.
              - rewrite zeq_false; auto; rewrite zeq_false in in_intr_re; auto.
              - unfold relate_SyncChanPool_per_pd in *; simpl; auto.
              - unfold relate_uctxt_per_pd in ×.
                rewrite H9 in ×.
                simpl.
                rewrite ZMap.gso; auto.
              - auto.
              - auto.
              - auto.
              - inv inv_re; constructor; simpl; auto. }
            constructor.
          × inv sh_shared_inv_re; constructor; simpl; auto.
            { simpl in H9.
              rewrite H9.
              intros contra; inv contra. }
            
            { intros.
              case_eq (zeq tid (ZMap.get (CPU_ID (update init_shared_adt l))
                                         (cid (update init_shared_adt l)))); intros; subst.
              - unfold B_GetContainerUsed´ in H.
                simpl in H11; rewrite H11 in H.
                unfold curid, cpu in H12.
                simpl in H12.
                rewrite H12 in H.
                inv H.
              - rewrite ZMap.gso; auto. }
          × inv sh_mem_inv_re.
            constructor; simpl; intros.
            case_eq (zeq i (ZMap.get (CPU_ID (update init_shared_adt l))
                                     (cid (update init_shared_adt l)))); intros; subst.
            { rewrite ZMap.gss in H3.
              rewrite ZMap.gso in H13; eauto.
              inv H3; simpl in H15.
              unfold pv_adt in H15; simpl in H15.
              eapply pperm_disjoint; eauto. }
            { rewrite ZMap.gso in H3; auto.
              case_eq (zeq j (ZMap.get (CPU_ID (update init_shared_adt l))
                                       (cid (update init_shared_adt l)))); intros; subst.
              - rewrite ZMap.gss in H13; eauto.
                inv H13.
                simpl.
                eapply pperm_disjoint; eauto.
              - rewrite ZMap.gso in H13; eauto. }
            Opaque proc_id uRData update.
    Fail idtac.
    Qed.
  End ABS_REL_INSTANCE.

End PHBTHREADABSREL.