Library mcertikos.ticketlog.MQTicketLockOp


This file defines the abstract data and the primitives for the PAbQueue layer, which will introduce abstraction of kernel context
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 LAsm.
Require Import LoadStoreSem1.
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 AbstractDataType.

Require Export ObjCPU.
Require Export ObjMM.
Require Export ObjFlatMem.
Require Export ObjMultiprocessor.
Require Export ObjQLock.
Require Export ObjSerialDevice.
Require Export ObjKeyboard.
Require Export ObjInterruptController.
Require Export ObjX86.

Require Import INVLemmaQLock.
Require Import CalTicketLock.

Abstract Data and Primitives at this layer

Section WITHMEM.

  Local Open Scope Z_scope.

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

**Definition of the raw data at this layer

  Section correct_qlock_status_inv.

    Definition qlock_length (l: MultiLog):=
       self_c other_c b q tq
             (Hcal: Q_CalLock l = Some (self_c, other_c, b, q, tq)),
        ( i, In i q → 0 i < TOTAL_CPU).

    Definition qlock_pool_length (l: MultiLogPool):=
       i ofs r q
             (Hrange: index2Z i ofs = Some r)
             (Hdef: ZMap.get r l = MultiDef q),
        qlock_length q.

    Definition correct_qlock_status (l: MultiLog) (k: LockStatus) (cpu: Z):=
       self_c other_c b q tq
             (Hcal: Q_CalLock l = Some (self_c, other_c, b, q, tq)),
        (k = LockFalse¬ In cpu q)
         ( lb, k = LockOwn lb , q = cpu :: b = LHOLD ¬ In cpu ).

    Definition correct_qlock_pool_status (l: MultiLogPool) (k: Lock) (cpu: Z):=
       i ofs r q
             (Hrange: index2Z i ofs = Some r)
             (Hdef: ZMap.get r l = MultiDef q),
        correct_qlock_status q (ZMap.get r k) cpu.

    Local Transparent Q_CalLock.

    Lemma correct_qlock_pool_status_gss´:
       l k c z
             (HP: correct_qlock_pool_status l k c)
             (Hc: correct_qlock_status (ZMap.get z k) c),
        correct_qlock_pool_status
          (ZMap.set z (MultiDef ) l) k c.
    Proof.
      intros. unfold correct_qlock_pool_status in *; intros.
      destruct (zeq r z); subst.
      - rewrite ZMap.gss in ×. inv Hdef. trivial.
      - rewrite ZMap.gso in *; eauto.
    Qed.

    Inductive MatchLock: LockStatusLockStatusProp :=
    | MLFALSE: MatchLock LockFalse LockFalse
    | MLOWN: a b,
               MatchLock (LockOwn a) (LockOwn b).

    Lemma correct_qlock_status_Match:
       l k c
             (Hc: correct_qlock_status l k c)
             (Hl: MatchLock k),
        correct_qlock_status l c.
    Proof.
      intros. unfold correct_qlock_status in *; intros.
      exploit Hc; eauto.
      intros (HF & HO).
      split; intros; subst; eauto.
      - inv Hl. eapply HF; eauto.
      - inv Hl. eapply HO; eauto.
    Qed.

    Lemma correct_qlock_pool_status_gss:
       l k c z
             (HP: correct_qlock_pool_status l k c)
             (Hc: correct_qlock_status c),
        correct_qlock_pool_status
          (ZMap.set z (MultiDef ) l) (ZMap.set z k) c.
    Proof.
      intros.
      unfold correct_qlock_pool_status in *; intros.
      destruct (zeq r z); subst.
      - rewrite ZMap.gss in ×. inv Hdef.
        trivial.
      - rewrite ZMap.gso in *; eauto.
    Qed.

    Lemma correct_qlock_pool_status_gss_match:
       l k c z
             (HP: correct_qlock_pool_status l k c)
             (Hc: correct_qlock_status (ZMap.get z k) c)
             (Hl: MatchLock (ZMap.get z k)),
        correct_qlock_pool_status
          (ZMap.set z (MultiDef ) l) (ZMap.set z k) c.
    Proof.
      intros. eapply correct_qlock_pool_status_gss; eauto.
      eapply correct_qlock_status_Match; eauto.
    Qed.

    Lemma qlock_pool_length_gss:
       l z
             (Hq: qlock_pool_length l)
             (Hlen: qlock_length ),
        qlock_pool_length
          (ZMap.set z (MultiDef ) l).
    Proof.
      intros. unfold qlock_pool_length in *; intros.
      destruct (zeq r z); subst.
      - rewrite ZMap.gss in ×. inv Hdef. trivial.
      - rewrite ZMap.gso in *; eauto.
    Qed.

    Lemma qlock_length_cons:
       i e l
             (Hrange: 0 i < TOTAL_CPU)
             (Hl: qlock_length l),
        qlock_length (TEVENT i e :: l).
    Proof.
      unfold qlock_length; intros.
      simpl in Hcal.
      subdestruct; inv Hcal; eauto.
      - inv H; trivial. inv H0.
      - exploit Hl; eauto. right; trivial.
      - rewrite app_comm_cons in H.
        eapply in_app_or in H.
        inv H; eauto. inv H0; trivial. inv H.
    Qed.

    Lemma real_qlock_pool_length_pb:
       n l
             (Hc: qlock_pool_length l),
        qlock_pool_length (real_multi_log_pb n l).
    Proof.
      induction n; simpl; intros; trivial.
      eapply qlock_pool_length_gss; eauto.
      unfold qlock_length. simpl. intros. inv Hcal.
      inv H.
    Qed.

    Lemma real_qlock_pool_length:
       l
             (HP: qlock_pool_length l),
         qlock_pool_length (real_multi_log l).
    Proof.
      unfold real_multi_log.
      eapply real_qlock_pool_length_pb.
    Qed.

    Lemma real_correct_qlock_pool_status_pb:
       n l d c
             (Hc: correct_qlock_pool_status l d c),
        correct_qlock_pool_status (real_multi_log_pb n l)
                                  (real_lock_pb n d) c.
    Proof.
      induction n; simpl; intros; trivial.
      eapply correct_qlock_pool_status_gss; eauto.
      unfold correct_qlock_status.
      intros. simpl in Hcal. inv Hcal.
      split; intros.
      - intro HF; inv HF.
      - inv H.
    Qed.

    Lemma real_correct_qlock_pool_status:
       l d c
             (HP: correct_qlock_pool_status l d c),
        correct_qlock_pool_status
          (real_multi_log l) (real_lock d) c.
    Proof.
      unfold real_multi_log.
      unfold real_lock.
      eapply real_correct_qlock_pool_status_pb.
    Qed.

    Lemma count_occ_append:
       l x,
        (count_occ zeq (l ++ ) x = count_occ zeq l x + count_occ zeq x) % nat.
    Proof.
      induction l; simpl; intros; auto.
      destruct (zeq a x); subst.
      - erewrite IHl. omega.
      - eauto.
    Qed.

    Lemma Q_CalLock_unique:
       l self_c other_c b q tq
             (Hcal: Q_CalLock l = Some (self_c, other_c, b, q, tq)),
        ( i, In i qcount_occ zeq q i = 1%nat).
    Proof.
      induction l; simpl; intros.
      - inv Hcal. inv H.
      - subdestruct; inv Hcal; try (exploit IHl; eauto; fail).
        + simpl in H. destruct H as [HF|HF]; inv HF.
          simpl. rewrite zeq_true. trivial.
        + specialize (IHl _ _ _ _ _ refl_equal).
          exploit (IHl i).
          { right; trivial. }
          simpl. destruct (zeq i z); subst.
          × rewrite zeq_true. intros.
            eapply (count_occ_In zeq) in H.
            omega.
          × rewrite zeq_false; auto.
        + specialize (IHl _ _ _ _ _ refl_equal).
          destruct (zeq i i0); subst.
          × rewrite app_comm_cons.
            rewrite count_occ_append.
            replace (count_occ zeq (z :: l3) i0) with O.
            {
              simpl. rewrite zeq_true. trivial.
            }
            {
              destruct (count_occ zeq (z :: l3) i0) eqn: Hpos; trivial.
              assert (Hpos´: (count_occ zeq (z :: l3) i0 > O)%nat) by omega.
              eapply count_occ_In in Hpos´. destruct Hpos´ as [HF|HF]; subst; congruence.
            }
          × rewrite app_comm_cons.
            rewrite count_occ_append.
            rewrite IHl.
            {
              simpl. rewrite zeq_false; auto.
            }
            {
              rewrite app_comm_cons in H.
              apply in_app_or in H.
              destruct H as [HF| HF]; trivial.
              destruct HF as [HF´|HF´]; inv HF´. congruence.
            }
    Qed.

    Lemma correct_status_HOLD_LOCK:
       l c b
      ,
        correct_qlock_status (TEVENT c (TTICKET HOLD_LOCK) :: l) (LockOwn b) c.
    Proof.
      unfold correct_qlock_status.
      simpl; intros.
      subdestruct; inv Hcal.
      - split; intros. inv H. inv H.
        refine_split´; trivial.
        intro HF.
        assert (Hin: In z (z :: l3)).
        {
          left; trivial.
        }
        eapply Q_CalLock_unique in Hdestruct; eauto.
        simpl in Hdestruct. rewrite zeq_true in Hdestruct.
        inv Hdestruct. eapply (count_occ_In zeq) in HF. omega.
      - split; intros. inv H. inv H.
        refine_split´; trivial.
        intro HF.
        assert (Hin: In z (z :: l3)).
        {
          left; trivial.
        }
        eapply Q_CalLock_unique in Hdestruct; eauto.
        simpl in Hdestruct. rewrite zeq_true in Hdestruct.
        inv Hdestruct. eapply (count_occ_In zeq) in HF. omega.
    Qed.

    Lemma correct_qlock_status_shared:
       l d c e
             (Hc: correct_qlock_status l d c),
        correct_qlock_status (TEVENT (TSHARED e) :: l) d c.
    Proof.
      unfold correct_qlock_status; intros.
      simpl in Hcal. subdestruct; inv Hcal.
      - exploit Hc; eauto.
      - exploit Hc; eauto.
    Qed.

    Lemma correct_status_INC_NOW:
       l c lb
             (Hc: correct_qlock_status l (LockOwn lb) c),
        correct_qlock_status (TEVENT c (TTICKET INC_NOW) :: l) LockFalse c.
    Proof.
      unfold correct_qlock_status.
      simpl; intros.
      subdestruct; inv Hcal.
      - exploit Hc; eauto.
        intros (_ & HL). specialize (HL _ refl_equal).
        destruct HL as ( & Heq & _ & Hnot). inv Heq.
        split; intros; trivial. inv H.
      - exploit Hc; eauto.
        intros (_ & HL). specialize (HL _ refl_equal).
        destruct HL as ( & Heq & _ & Hnot). inv Heq.
        split; intros; trivial. inv H.
    Qed.

    Local Opaque Q_CalLock.

  End correct_qlock_status_inv.

  Section cal_qlock_status.

    Definition cal_qlock_status (l: MultiLog) (k: LockStatus) (cpu: Z):=
       self_c other_c b q tq
             (Hcal: Q_CalLock l = Some (self_c, other_c, b, q, tq)),
        (k = LockOwn falseCalValidBit l = Some FreeToPull)
         (k = LockOwn trueCalValidBit l = Some (PullBy cpu)).

    Definition cal_qlock_pool_status (l: MultiLogPool) (k: Lock) (cpu: Z):=
       i ofs r q
             (Hrange: index2Z i ofs = Some r)
             (Hdef: ZMap.get r l = MultiDef q),
        cal_qlock_status q (ZMap.get r k) cpu.

  End cal_qlock_status.

  Section last_pool.

    Definition valid_last (c: Z) (q: MultiLog) :=
      ( i1 e1,
         In (TEVENT i1 e1) q → 0 i1 < TOTAL_CPU)
       (match q with
            | nilTrue
            | TEVENT c0 _ :: _c0 = c
          end).

    Definition valid_last_pool (c: Z) (l: MultiLogPool) :=
       i ofs r q
             (Hrange: index2Z i ofs = Some r)
             (Hdef: ZMap.get r l = MultiDef q),
        valid_last c q.

    Definition valid_from (c: Z) (q: MultiLog) (o: MultiOracle):=
       i1 e1, In (TEVENT i1 e1) qi1 c l0, In (TEVENT i1 e1) (o c l0).

    Definition valid_from_pool (c: Z) (l: MultiLogPool) (op: MultiOraclePool):=
       i ofs r q o
             (Hrange: index2Z i ofs = Some r)
             (Hdef: ZMap.get r l = MultiDef q)
             (Horacle: ZMap.get r op = o),
        valid_from c q o.

  End last_pool.

**Definition of the invariants at MPTNew layer 0th page map is reserved for the kernel thread
  Record high_level_invariant (abd: RData) :=
    mkInvariant {
        CPU_ID_range: 0 (CPU_ID abd) < TOTAL_CPU;
        valid_curid: 0 ZMap.get (CPU_ID abd) (cid abd) < num_proc;
        valid_multi_oracle_pool_inv: valid_multi_oracle_pool (multi_oracle abd);
        get_free_qlock_oracle_pool_inv: get_free_qlock_oracle_pool (multi_oracle abd);
        valid_qlock_pool_inv: valid_qlock_pool (multi_log abd);
        valid_qlock_pool_status_inv:
          correct_qlock_pool_status (multi_log abd) (lock abd) (CPU_ID abd);
        cal_qlock_pool_status_inv:
          cal_qlock_pool_status (multi_log abd) (lock abd) (CPU_ID abd);
        get_free_qlock_pool_inv:
          get_free_qlock_pool (multi_log abd);
        qlock_pool_length_inv:
          qlock_pool_length (multi_log abd);
        valid_last_pool_inv:
          valid_last_pool (CPU_ID abd) (multi_log abd);
        valid_from_inv: valid_from_pool (CPU_ID abd) (multi_log abd) (multi_oracle abd)
      }.

Definition of the abstract state ops

  Global Instance mqticketlockop_data_ops : CompatDataOps RData :=
    {
      empty_data := init_adt multi_oracle_init0;
      high_level_invariant := high_level_invariant;
      low_level_invariant := low_level_invariant;
      kernel_mode adt := ikern adt = true ihost adt = true
    }.

Proofs that the initial abstract_data should satisfy the invariants

  Section Property_Abstract_Data.

    Lemma correct_qlock_pool_status_init:
       l z,
        correct_qlock_pool_status (ZMap.init MultiUndef) l z.
    Proof.
      unfold correct_qlock_pool_status.
      intros. rewrite ZMap.gi in Hdef. inv Hdef.
    Qed.

    Lemma qlock_pool_length_init:
      qlock_pool_length (ZMap.init MultiUndef).
    Proof.
      unfold qlock_pool_length.
      intros. rewrite ZMap.gi in Hdef. inv Hdef.
    Qed.

    Lemma cal_qlock_pool_status_init:
       l z,
        cal_qlock_pool_status (ZMap.init MultiUndef) l z.
    Proof.
      unfold cal_qlock_pool_status.
      intros. rewrite ZMap.gi in Hdef. inv Hdef.
    Qed.

    Lemma get_free_qlock_pool_init:
      get_free_qlock_pool (ZMap.init MultiUndef).
    Proof.
      unfold get_free_qlock_pool.
      intros. rewrite ZMap.gi in Hdef. inv Hdef.
    Qed.

    Lemma valid_last_pool_init:
      valid_last_pool current_CPU_ID (ZMap.init MultiUndef).
    Proof.
      unfold valid_last_pool.
      intros. rewrite ZMap.gi in Hdef. inv Hdef.
    Qed.

    Lemma valid_from_pool_init:
       op, valid_from_pool current_CPU_ID (ZMap.init MultiUndef) op.
    Proof.
      unfold valid_from_pool.
      intros. rewrite ZMap.gi in Hdef. inv Hdef.
    Qed.

    Lemma empty_data_high_level_invariant:
      high_level_invariant (init_adt multi_oracle_init0).
    Proof.
      constructor; trivial; simpl; try congruence; intros.
      - apply current_CPU_ID_range.
      - rewrite ZMap.gi; intuition.
      - apply valid_ticket_oracle0.
      - eapply get_free_qlock_oracle0.
      - apply valid_qlock_pool_init.
      - apply correct_qlock_pool_status_init.
      - apply cal_qlock_pool_status_init.
      - apply get_free_qlock_pool_init.
      - apply qlock_pool_length_init.
      - apply valid_last_pool_init.
      - apply valid_from_pool_init.
    Qed.

Definition of the abstract state

    Global Instance mqticketlockop_data_prf : CompatData RData.
    Proof.
      constructor.
      - apply low_level_invariant_incr.
      - apply empty_data_low_level_invariant.
      - apply empty_data_high_level_invariant.
    Qed.

  End Property_Abstract_Data.

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

Proofs that the primitives satisfies the invariants at this layer

  Section INV.

    Global Instance setPG_inv: PreservesInvariants setPG_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance fstore_inv: PreservesInvariants fstore´_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance setCR3_inv: SetCR3Invariants setCR3_spec.
    Proof.
      constructor; intros.
      - functional inversion H; inv H0; constructor; trivial.
      - functional inversion H; inv H0; constructor; auto.
      - functional inversion H; simpl in *; assumption.
    Qed.

    Global Instance trapin_inv: PrimInvariants trapin_spec.
    Proof.
      PrimInvariants_simpl_auto.
    Qed.

    Global Instance trapout_inv: PrimInvariants trapout´_spec.
    Proof.
      PrimInvariants_simpl_auto.
    Qed.

    Global Instance hostin_inv: PrimInvariants hostin_spec.
    Proof.
      PrimInvariants_simpl_auto.
    Qed.

    Global Instance hostout_inv: PrimInvariants hostout´_spec.
    Proof.
      PrimInvariants_simpl_auto.
    Qed.

    Global Instance set_curid_inv: PreservesInvariants set_curid_spec.
    Proof.
      preserves_invariants_simpl_auto; eauto 2.
      - rewrite ZMap.gss. auto.
    Qed.

    Global Instance set_curid_init_inv: PreservesInvariants set_curid_init_spec.
    Proof.
      preserves_invariants_simpl_auto; eauto 2.
      case_eq (zeq (CPU_ID d) (Int.unsigned i)); subst.
      - intros.
        rewrite e; rewrite ZMap.gss; omega.
      - intros; rewrite ZMap.gso; auto.
    Qed.

    Lemma cal_qlock_pool_status_gss´:
       l k c z
             (HP: cal_qlock_pool_status l k c)
             (Hc: cal_qlock_status (ZMap.get z k) c),
        cal_qlock_pool_status
          (ZMap.set z (MultiDef ) l) k c.
    Proof.
      intros. unfold cal_qlock_pool_status in *; intros.
      destruct (zeq r z); subst.
      - rewrite ZMap.gss in ×. inv Hdef. trivial.
      - rewrite ZMap.gso in *; eauto.
    Qed.


    Lemma cal_qlock_pool_status_gss:
       l k c z
             (HP: cal_qlock_pool_status l k c)
             (Hc: cal_qlock_status c),
        cal_qlock_pool_status
          (ZMap.set z (MultiDef ) l) (ZMap.set z k) c.
    Proof.
      intros.
      unfold cal_qlock_pool_status in *; intros.
      destruct (zeq r z); subst.
      - rewrite ZMap.gss in ×. inv Hdef.
        trivial.
      - rewrite ZMap.gso in *; eauto.
    Qed.

    Lemma real_cal_qlock_pool_status_pb:
       n l d c
             (Hc: cal_qlock_pool_status l d c),
        cal_qlock_pool_status (real_multi_log_pb n l)
                                  (real_lock_pb n d) c.
    Proof.
      induction n; simpl; intros; trivial.
      eapply cal_qlock_pool_status_gss; eauto.
      unfold cal_qlock_status.
      intros. simpl in Hcal. inv Hcal.
      split; intro HF; inv HF.
    Qed.

    Lemma real_cal_qlock_pool_status:
       l d c
             (HP: cal_qlock_pool_status l d c),
        cal_qlock_pool_status
          (real_multi_log l) (real_lock d) c.
    Proof.
      unfold real_multi_log.
      unfold real_lock.
      eapply real_cal_qlock_pool_status_pb.
    Qed.

    Lemma get_free_qlock_pool_gss:
       l z
             (HP: get_free_qlock_pool l)
             (Hc: get_free_qlock ),
        get_free_qlock_pool
          (ZMap.set z (MultiDef ) l).
    Proof.
      intros. unfold get_free_qlock_pool in *; intros.
      destruct (zeq r z); subst.
      - rewrite ZMap.gss in ×. inv Hdef. trivial.
      - rewrite ZMap.gso in *; eauto.
    Qed.

    Lemma real_get_free_qlock_pool_pb:
       n l
             (Hc: get_free_qlock_pool l),
        get_free_qlock_pool (real_multi_log_pb n l).
    Proof.
      induction n; simpl; intros; trivial.
      eapply get_free_qlock_pool_gss; eauto.
      unfold get_free_qlock.
      intros. destruct Hincr as (prefix & Hcons).
      apply app_eq_nil in Hcons. destruct Hcons; subst.
      trivial.
    Qed.

    Lemma real_get_free_qlock_pool:
       l
             (HP: get_free_qlock_pool l),
        get_free_qlock_pool
          (real_multi_log l).
    Proof.
      unfold real_multi_log.
      eapply real_get_free_qlock_pool_pb.
    Qed.

    Lemma valid_last_pool_gss:
       c l z
             (HP: valid_last_pool c l)
             (Hc: valid_last c ),
        valid_last_pool c (ZMap.set z (MultiDef ) l).
    Proof.
      intros. unfold valid_last_pool in *; intros.
      destruct (zeq r z); subst.
      - rewrite ZMap.gss in ×. inv Hdef. trivial.
      - rewrite ZMap.gso in *; eauto.
    Qed.

    Lemma real_valid_last_pool_pb:
       n c l
             (Hc: valid_last_pool c l),
        valid_last_pool c (real_multi_log_pb n l).
    Proof.
      induction n; simpl; intros; trivial.
      eapply valid_last_pool_gss; eauto.
      unfold valid_last.
      split; intros; trivial. inv H.
    Qed.

    Lemma real_valid_last_pool:
       c l
             (HP: valid_last_pool c l),
        valid_last_pool c
          (real_multi_log l).
    Proof.
      unfold real_multi_log.
      eapply real_valid_last_pool_pb.
    Qed.

    Lemma valid_from_pool_gss:
       c l z op
             (HP: valid_from_pool c l op)
             (Hc: valid_from c (ZMap.get z op)),
        valid_from_pool c (ZMap.set z (MultiDef ) l) op.
    Proof.
      intros. unfold valid_from_pool in *; intros.
      destruct (zeq r z); subst.
      - rewrite ZMap.gss in ×. inv Hdef. trivial.
      - rewrite ZMap.gso in *; eauto.
    Qed.

    Lemma real_valid_from_pool_pb:
       n c l o
             (Hc: valid_from_pool c l o),
        valid_from_pool c (real_multi_log_pb n l) o.
    Proof.
      induction n; simpl; intros; trivial.
      eapply valid_from_pool_gss; eauto.
      unfold valid_from. intros. inv H.
    Qed.

    Lemma real_valid_from_pool:
       c l o
             (HP: valid_from_pool c l o),
        valid_from_pool c (real_multi_log l) o.
    Proof.
      unfold real_multi_log.
      eapply real_valid_from_pool_pb.
    Qed.

    Global Instance ticket_lock_init_inv: PreservesInvariants ticket_lock_init0_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant.
      - eassumption.
      - eassumption.
      - eapply real_valid_qlock_pool; trivial.
      - eapply real_correct_qlock_pool_status; eauto.
      - eapply real_cal_qlock_pool_status; eauto.
      - eapply real_get_free_qlock_pool; eauto.
      - eapply real_qlock_pool_length; eauto.
      - eapply real_valid_last_pool; eauto.
      - eapply real_valid_from_pool; eauto.
    Qed.

    Lemma cal_qlock_status_shared_pull:
       l c
             (Hc: cal_qlock_status l (LockOwn false) c),
        cal_qlock_status (TEVENT c (TSHARED OPULL) :: l) (LockOwn true) c.
    Proof.
      unfold cal_qlock_status; intros.
      split; intros HF; inv HF.
      Local Transparent Q_CalLock.
      simpl in Hcal. subdestruct; subst.
      - inv Hcal. exploit Hc; eauto.
        intros (HF & _). rewrite HF; trivial.
      - inv Hcal. exploit Hc; eauto.
        intros (HF & _). rewrite HF; trivial.
      Local Opaque Q_CalLock.
    Qed.

    Lemma cal_qlock_status_shared:
       l c m
             (Hc: cal_qlock_status l (LockOwn true) c),
        cal_qlock_status (TEVENT c (TSHARED (OMEME m)) :: l) (LockOwn false) c.
    Proof.
      unfold cal_qlock_status; intros.
      Local Transparent Q_CalLock.
      simpl in Hcal. subdestruct; subst.
      - inv Hcal. exploit Hc; eauto.
        intros (_ & HT).
        split; intros HP; inv HP.
        rewrite HT; trivial.
        rewrite zeq_true. trivial.
      - inv Hcal. exploit Hc; eauto.
        intros (_ & HF). rewrite HF; trivial.
        rewrite zeq_true. split; intros HP; inv HP. trivial.
      Local Opaque Q_CalLock.
    Qed.

    Lemma Q_CalLock_append:
       e l p
             (Hcal: Q_CalLock (e :: l) = Some p),
       self_c other_c b q tq,
        Q_CalLock l = Some (self_c, other_c, b, q, tq).
    Proof.
      Local Transparent Q_CalLock.
      simpl; intros. subdestruct; inv Hcal; eauto 20.
      Local Opaque Q_CalLock.
    Qed.

    Lemma valid_last_event:
       c l e (Hc: 0 c < TOTAL_CPU)
             (Hvalid: valid_last c l),
        valid_last c (TEVENT c e :: l).
    Proof.
      unfold valid_last. intros. destruct Hvalid as (Hvalid1 & Hvalid2).
      split; trivial. intros. inv H.
      - inv H0. eauto.
      - eauto.
    Qed.

    Lemma valid_from_event:
       c l e o (Hvalid: valid_from c l o),
        valid_from c (TEVENT c e :: l) o.
    Proof.
      unfold valid_from. intros. inv H.
      - inv H1. elim H0. trivial.
      - eauto.
    Qed.

    Lemma get_free_qlock_shared:
       l c e self_c other_c q tq
             (Hfree: get_free_qlock l)
             (Hcal: Q_CalLock l = Some (self_c, other_c, LHOLD, q, tq)),
        get_free_qlock (TEVENT c (TSHARED e) :: l).
    Proof.
      Local Transparent Q_CalLock.
      unfold get_free_qlock; intros.
      simpl in ×. destruct Hincr as (prefix & Hprefix).
      destruct prefix.
      - simpl in ×. inv Hprefix. simpl in ×.
        rewrite Hcal in Hcal0.
        destruct Hk; subst; subdestruct.
      - inv Hprefix. eauto.
      Local Opaque Q_CalLock.
    Qed.

    Global Instance acquire_shared_inv: SAcquireInvariants acquire_shared1_spec0.
    Proof.
      constructor; unfold acquire_shared1_spec; intros; subdestruct;
      inv H; inv H0; constructor; auto.
      - eapply valid_qlock_pool_gss´; eauto.
      - simpl. eapply correct_qlock_pool_status_gss_match; eauto.
        + eapply correct_qlock_status_shared; eauto.
        + rewrite Hdestruct3. constructor.
      - simpl. eapply cal_qlock_pool_status_gss; eauto.
        eapply cal_qlock_status_shared_pull; eauto.
        unfold cal_qlock_pool_status in ×.
        rewrite <- Hdestruct3. eauto.
      - simpl. eapply get_free_qlock_pool_gss; eauto.
        eapply Q_CalLock_append in Hdestruct6.
        destruct Hdestruct6 as (self_c & other_c & b & q & tq & Hcal).
        unfold correct_qlock_pool_status in ×.
        exploit valid_qlock_pool_status_inv0; eauto.
        rewrite Hdestruct3. intros (_ & HT).
        exploit HT; eauto.
        intros (_ & _ & ? & _); subst.
        eapply get_free_qlock_shared; eauto.
      - eapply qlock_pool_length_gss; eauto.
        eapply qlock_length_cons; eauto.
      - eapply valid_last_pool_gss; eauto.
        simpl. eapply valid_last_event; eauto.
      - eapply valid_from_pool_gss; eauto.
        simpl. eapply valid_from_event; eauto.
    Qed.

    Global Instance release_shared_inv: ReleaseInvariants release_shared1_spec0.
    Proof.
      constructor; unfold release_shared1_spec; intros; subdestruct;
      inv H; inv H0; constructor; auto.
      - eapply valid_qlock_pool_gss´; eauto.
      - simpl. eapply correct_qlock_pool_status_gss_match; eauto.
        + eapply correct_qlock_status_shared; eauto.
        + rewrite Hdestruct3. constructor.
      - simpl. eapply cal_qlock_pool_status_gss; eauto.
        specialize (valid_qlock_pool_status_inv0 _ _ _ _ Hdestruct2 Hdestruct5).
        eapply cal_qlock_status_shared; eauto.
        unfold cal_qlock_pool_status in ×. rewrite <- Hdestruct3. eauto.
      - simpl. eapply get_free_qlock_pool_gss; eauto.
        eapply Q_CalLock_append in Hdestruct6.
        destruct Hdestruct6 as (self_c & other_c & b & q & tq & Hcal).
        unfold correct_qlock_pool_status in ×.
        exploit valid_qlock_pool_status_inv0; eauto.
        rewrite Hdestruct3. intros (_ & HT).
        exploit HT; eauto.
        intros (_ & _ & ? & _); subst.
        eapply get_free_qlock_shared; eauto.
      - eapply qlock_pool_length_gss; eauto.
        eapply qlock_length_cons; eauto.
      - eapply valid_last_pool_gss; eauto.
        simpl. eapply valid_last_event; eauto.
      - eapply valid_from_pool_gss; eauto.
        simpl. eapply valid_from_event; eauto.
    Qed.

    Lemma cal_status_INC_NOW:
       l c
             (Hc: cal_qlock_status l (LockOwn false) c),
        cal_qlock_status (TEVENT c (TTICKET INC_NOW) :: l) LockFalse c.
    Proof.
      Local Transparent Q_CalLock.
      unfold cal_qlock_status.
      simpl; intros.
      subdestruct; inv Hcal.
      - exploit Hc; eauto.
        intros (HL & _). specialize (HL refl_equal). rewrite HL.
        split; intros HP; inv HP.
      - exploit Hc; eauto.
        intros (HL & _). specialize (HL refl_equal).
        rewrite HL.
        split; intros HP; inv HP.
      Local Opaque Q_CalLock.
    Qed.

    Lemma get_free_INC_NOW:
       l c
             (Hfree: get_free_qlock l)
             (Hcal: CalValidBit l = Some FreeToPull),
        get_free_qlock (TEVENT c (TTICKET INC_NOW) :: l).
    Proof.
      Local Transparent Q_CalLock.
      unfold get_free_qlock.
      simpl; intros. destruct Hincr as (prefix & Hprefix).
      destruct prefix.
      - simpl in ×. inv Hprefix. simpl in ×.
        rewrite Hcal. trivial.
      - inv Hprefix. eauto.
      Local Opaque Q_CalLock.
    Qed.

    Global Instance pass_lock_inv: PreservesInvariants pass_qlock_spec.
    Proof.
      preserves_invariants_simpl_auto1.
      - eapply valid_qlock_pool_gss´; eauto.
      - eapply correct_qlock_pool_status_gss; eauto.
        eapply correct_status_INC_NOW; eauto.
        instantiate (1:= false).
        specialize (valid_qlock_pool_status_inv0 _ _ _ _ Hdestruct1 Hdestruct2).
        rewrite <- Hdestruct4. trivial.
      - eapply cal_qlock_pool_status_gss; eauto.
        eapply cal_status_INC_NOW; eauto.
        specialize (cal_qlock_pool_status_inv0 _ _ _ _ Hdestruct1 Hdestruct2).
        rewrite <- Hdestruct4. trivial.
      - simpl. eapply get_free_qlock_pool_gss; eauto.
        eapply get_free_INC_NOW; eauto.
        eapply Q_CalLock_append in Hdestruct3.
        destruct Hdestruct3 as (self_c & other_c & b & q & tq & Hcal).
        unfold cal_qlock_pool_status in ×.
        exploit cal_qlock_pool_status_inv0; eauto.
        rewrite Hdestruct4. intros (HT & _).
        exploit HT; eauto.
      - eapply qlock_pool_length_gss; eauto.
        eapply qlock_length_cons; eauto.
      - eapply valid_last_pool_gss; eauto.
        simpl. eapply valid_last_event; eauto.
      - eapply valid_from_pool_gss; eauto.
        simpl. eapply valid_from_event; eauto.
    Qed.

    Lemma qlock_length_append:
       l
             (Hlen1: qlock_length l)
             (Hto: i0 e, In (TEVENT i0 e) → 0 i0 < TOTAL_CPU),
        qlock_length ( ++ l).
    Proof.
      induction ; simpl; intros; trivial.
      destruct a.
      eapply qlock_length_cons; eauto.
    Qed.

    Lemma Q_CalWait_length:
       n l i to
             (Hrange: 0 i < TOTAL_CPU)
             (Hlen: qlock_length l)
             (Hto: q i0 e, In (TEVENT i0 e) (to i q) → 0 i0 < TOTAL_CPU)
             (Hwait: Q_CalWait n i l to = Some ),
        qlock_length .
    Proof.
      Local Transparent Q_CalWait.
      induction n; simpl; intros; try congruence.
      subdestruct; subst.
      - inv Hwait.
        eapply qlock_length_cons; eauto.
        eapply qlock_length_append; eauto.
      - eapply IHn; try eapply Hwait; eauto.
        eapply qlock_length_cons; eauto.
        eapply qlock_length_append; eauto.
    Qed.

    Lemma get_free_GET_NOW:
       l c
             (Hget: get_free_qlock l),
        get_free_qlock (TEVENT c (TTICKET GET_NOW) :: l).
    Proof.
      Local Transparent Q_CalLock.
      unfold get_free_qlock.
      simpl; intros.
      destruct Hincr as (prefix & Hprefix).
      destruct prefix.
      - simpl in ×. inv Hprefix. simpl in ×.
        subdestruct; inv Hcal.
        + exploit Hget; eauto.
           nil; trivial.
          intros →. trivial.
        + exploit Hget; eauto.
           nil; trivial.
          intros →. trivial.
      - inv Hprefix. eauto.
      Local Opaque Q_CalLock.
    Qed.

    Lemma valid_from_oracle:
       i l to,
        valid_from i l to
        valid_from i (to i l ++ l) to.
    Proof.
      unfold valid_from. intros.
      eapply in_app_or in H0. inv H0; eauto.
    Qed.

    Lemma Q_CalWait_get_free:
       n l i to
             (Hfree: get_free_qlock l)
             (Hto: get_free_qlock_oracle to)
             (Hfrom: valid_from i l to)
             (Hwait: Q_CalWait n i l to = Some ),
        get_free_qlock .
    Proof.
      induction n; simpl; intros; try congruence.
      subdestruct; subst.
      - inv Hwait.
        eapply get_free_GET_NOW; eauto.
      - eapply IHn; try eapply Hwait; eauto.
        eapply get_free_GET_NOW; eauto.
        eapply valid_from_event.
        eapply valid_from_oracle; eauto.
    Qed.

    Lemma Q_CalWait_valid_last:
       n l i to
             (Hvalid: i1 e1,
                        In (TEVENT i1 e1) l → 0 i1 < TOTAL_CPU)
             (Hto: valid_multi_oracle_domain to)
             (Hwait: Q_CalWait n i l to = Some )
             (Hi: 0 i < TOTAL_CPU),
      ( i1 e1,
         In (TEVENT i1 e1) → 0 i1 < TOTAL_CPU).
    Proof.
      induction n; simpl; intros; try congruence.
      subdestruct; subst.
      - inv Hwait.
        inv H.
        + inv H0. eauto.
        + eapply in_app_or in H0. inv H0.
          × eapply Hto. eauto.
          × eapply Hvalid. eauto.
      - eapply IHn; try eapply Hwait; eauto.
        intros. inv H0.
        + inv H1. eauto.
        + eapply in_app_or in H1. inv H1.
          × eapply Hto. eauto.
          × eapply Hvalid. eauto.
    Qed.

    Lemma Q_CalWait_valid_from:
       n l i to
             (Hfrom: valid_from i l to)
             (Hwait: Q_CalWait n i l to = Some ),
        valid_from i to.
    Proof.
      induction n; simpl; intros; try congruence.
      subdestruct; subst.
      - inv Hwait. eapply valid_from_event.
        eapply valid_from_oracle; eauto.
      - inv Hwait. eapply IHn in H0; eauto.
        eapply valid_from_event.
        eapply valid_from_oracle; eauto.
    Qed.

    Local Opaque Q_CalWait.

    Lemma cal_status_HOLD_LOCK:
       l c
             (Hget: get_free_qlock l),
        cal_qlock_status (TEVENT c (TTICKET HOLD_LOCK) :: l) (LockOwn false) c.
    Proof.
      Local Transparent Q_CalLock.
      unfold cal_qlock_status.
      simpl; intros.
      subdestruct; inv Hcal.
      - split; intros HP; inv HP.
        unfold get_free_qlock in ×. eapply Hget in Hdestruct; eauto.
        rewrite Hdestruct; trivial.
         nil; trivial.
      - split; intros HP; inv HP.
        unfold get_free_qlock in ×. eapply Hget in Hdestruct; eauto.
        rewrite Hdestruct; trivial.
         nil; trivial.
      Local Opaque Q_CalLock.
    Qed.

    Lemma get_free_HOLD_LOCK:
       l c
             (Hget: get_free_qlock l),
        get_free_qlock (TEVENT c (TTICKET HOLD_LOCK) :: l).
    Proof.
      Local Transparent Q_CalLock.
      unfold get_free_qlock.
      simpl; intros.
      destruct Hincr as (prefix & Hprefix).
      destruct prefix.
      - simpl in ×. inv Hprefix. simpl in ×.
        subdestruct; inv Hcal.
        + destruct Hk as [HT|HT]; inv HT.
        + destruct Hk as [HT|HT]; inv HT.
      - inv Hprefix. eauto.
      Local Opaque Q_CalLock.
    Qed.

    Lemma Q_CalLock_nil:
       l self_c other_c k
             (Hcal: Q_CalLock l = Some (self_c, other_c, k, nil, nil)),
        k = LEMPTY.
    Proof.
      Local Transparent Q_CalLock.
      induction l; simpl; intros; eauto.
      - inv Hcal. trivial.
      - subdestruct; inv Hcal; trivial.
    Qed.

    Lemma get_free_INC_TICKET:
       l c n
             (Hget: get_free_qlock l),
        get_free_qlock (TEVENT c (TTICKET (INC_TICKET n)) :: l).
    Proof.
      Local Transparent Q_CalLock.
      unfold get_free_qlock.
      simpl; intros.
      destruct Hincr as (prefix & Hprefix).
      destruct prefix.
      - simpl in ×. inv Hprefix. simpl in ×.
        subdestruct; inv Hcal.
        + exploit Hget; eauto.
           nil; trivial.
          eapply Q_CalLock_nil in Hdestruct; subst. eauto.
          intros ->; trivial.
        + exploit Hget; eauto.
           nil; trivial.
          intros ->; trivial.
      - inv Hprefix. eauto.
        Local Opaque Q_CalLock.
    Qed.

    Global Instance wait_lock_inv: PreservesInvariants wait_qlock_spec.
    Proof.
      preserves_invariants_simpl_auto1.
      - eapply valid_qlock_pool_gss´; eauto.
      - eapply correct_qlock_pool_status_gss; eauto.
        eapply correct_status_HOLD_LOCK; eauto.
      - eapply cal_qlock_pool_status_gss; eauto.
        eapply cal_status_HOLD_LOCK; eauto.
        eapply Q_CalWait_get_free; try eapply Hdestruct9; eauto.
        eapply get_free_INC_TICKET; eauto.
        eapply Q_CalLock_append in Hdestruct4.
        destruct Hdestruct4 as (self_c & other_c & b & q & tq & Hcal).
        unfold get_free_qlock_oracle_pool in ×.
        specialize (get_free_qlock_oracle_pool_inv0 _ _ _ Hdestruct1); eauto.
        eapply get_free_qlock_oracle_pool_inv0; eauto.
        × eapply valid_from_inv0; eauto.
        × eapply valid_from_event.
          eapply valid_from_oracle; eauto.

      - simpl. eapply get_free_qlock_pool_gss; eauto.
        eapply Q_CalWait_get_free in Hdestruct9; eauto.
        + eapply get_free_HOLD_LOCK; eauto.
        + eapply get_free_INC_TICKET; eauto.
          unfold get_free_qlock_oracle_pool in ×.
          specialize (get_free_qlock_oracle_pool_inv0 _ _ _ Hdestruct1); eauto.
          eapply get_free_qlock_oracle_pool_inv0; eauto.
          eapply valid_from_inv0; eauto.
        + eapply valid_from_event.
          eapply valid_from_oracle; eauto.

      - eapply qlock_pool_length_gss; eauto.
        eapply qlock_length_cons; eauto.
        eapply Q_CalWait_length; try eapply Hdestruct9; eauto.
        + eapply qlock_length_cons; eauto.
          eapply qlock_length_append; eauto.
          intros. eapply valid_multi_oracle_pool_inv0; eauto.
        + intros. eapply valid_multi_oracle_pool_inv0; eauto.
      - eapply valid_last_pool_gss; eauto.
        simpl. split; trivial.
        intros. inv H.
        + inv H0. eauto.
        + eapply Q_CalWait_valid_last; eauto.
          {
            intros. inv H.
            - inv H1. eauto.
            - eapply in_app_or in H1. inv H1.
              + eapply valid_multi_oracle_pool_inv0; eauto.
              + eapply valid_last_pool_inv0; eauto.
          }
          {
            eapply valid_multi_oracle_pool_inv0; eauto.
          }
      - eapply valid_from_pool_gss; eauto.
        eapply Q_CalWait_valid_from in Hdestruct9; eauto.
        + eapply valid_from_event; eauto.
        + eapply valid_from_event; eauto.
          eapply valid_from_oracle; eauto.
    Qed.


    Global Instance cli_inv: PreservesInvariants cli_spec.
    Proof.
      econstructor; intros; inv H; functional inversion H1; subst.
      - inv H0; constructor; trivial.
      - inv H0; constructor; auto.
      - simpl in *; assumption.
    Qed.

    Global Instance sti_inv: PreservesInvariants sti_spec.
    Proof.
      econstructor; intros; inv H; functional inversion H1; subst.
      - inv H0; constructor; trivial.
      - inv H0; constructor; auto.
      - simpl in *; assumption.
    Qed.

    Require Import FutureTactic.
    Global Instance serial_in_inv: PreservesInvariants serial_in_spec.
    Proof.
      preserves_invariants_subst low_level_invariant high_level_invariant; eauto.
    Qed.

    Global Instance serial_out_inv: PreservesInvariants serial_out_spec.
    Proof.
      preserves_invariants_subst low_level_invariant high_level_invariant; eauto.
    Qed.

    Global Instance serial_hw_intr_inv: PreservesInvariants serial_hw_intr_spec.
    Proof.
      preserves_invariants_nested low_level_invariant high_level_invariant; eauto.
    Qed.

    Global Instance ic_intr_inv: PreservesInvariants ic_intr_spec.
    Proof.
      preserves_invariants_nested low_level_invariant high_level_invariant; eauto.
    Qed.

    Global Instance keyboard_in_inv: PreservesInvariants keyboard_in_spec.
    Proof.
      preserves_invariants_subst low_level_invariant high_level_invariant; eauto.
    Qed.

    Global Instance keyboard_hw_intr_inv: PreservesInvariants keyboard_hw_intr_spec.
    Proof.
      preserves_invariants_nested low_level_invariant high_level_invariant; eauto.
    Qed.

    Global Instance ioapic_read_inv: PreservesInvariants ioapic_read_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant; eauto.
    Qed.

    Global Instance ioapic_write_inv: PreservesInvariants ioapic_write_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant; eauto.
    Qed.

    Global Instance lapic_read_inv: PreservesInvariants lapic_read_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant; eauto.
    Qed.

    Global Instance lapic_write_inv: PreservesInvariants lapic_write_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant; eauto.
    Qed.

    Global Instance iret_inv: PreservesInvariants iret_spec.
    Proof.
      preserves_invariants_direct low_level_invariant high_level_invariant; eauto 2.
    Qed.

    Global Instance local_irq_save_inv:
      PreservesInvariants local_irq_save_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
    Qed.

    Global Instance local_irq_restore_inv:
      PreservesInvariants local_irq_restore_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
    Qed.

    Global Instance save_context_inv:
      SaveContextInvariants save_context_spec.
    Proof.
      constructor; intros.
      -
        inv H0. unfold save_context_spec; simpl. inv H1.
        constructor; simpl; trivial.
        induction (tf d).
        simpl.
        constructor.
        constructor.
        simpl.
        apply inv_reg_wt.
        simpl.
        inv inv_inject_neutral.
        eapply inv_reg_inject_neutral; auto.
        constructor.
        constructor.
        inv tf_INJECT.
        assumption.
        eapply IHt; auto.
        inv tf_INJECT.
        assumption.
      -
        inv H0. unfold save_context_spec; simpl.
        constructor; simpl; trivial.
    Qed.

    Global Instance restore_context_inv:
      RestoreContextInvariants restore_context_spec.
    Proof.
      constructor; intros.
      -
        inv H0. functional inversion H; inv H1;
        constructor; simpl; trivial.
        rewrite H4 in ×.
        clear H4 H.
        generalize dependent _x.
        induction _x0.
        intros.
        simpl in ×.
        constructor.
        intros.
        simpl in ×.
        constructor.
        inv tf_INJECT.
        assumption.
        eapply IH_x0; auto.
        inv tf_INJECT; assumption.
      -
        inv H0. functional inversion H; subst;
        constructor; simpl; trivial.
    Qed.

    Global Instance proc_create_postinit_inv:
      PreservesInvariants proc_create_postinit_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
    Qed.

  End INV.

  Definition exec_loadex {F V} := exec_loadex1 (F := F) (V := V).

  Definition exec_storeex {F V} := exec_storeex1 (flatmem_store:= flatmem_store´) (F := F) (V := V).

  Global Instance flatmem_store_inv: FlatmemStoreInvariant (flatmem_store:= flatmem_store´).
  Proof.
    split; inversion 1; intros.
    - functional inversion H0; constructor; auto.
    - functional inversion H1; constructor; auto.
  Qed.

  Global Instance trapinfo_set_inv: TrapinfoSetInvariant.
  Proof.
    split; inversion 1; intros; constructor; auto.
  Qed.


  Global Instance page_copy_back_inv: PreservesInvariants page_copy_back´_spec.
  Proof.
    preserves_invariants_simpl_auto.
  Qed.

  Lemma cal_qlock_status_buffer:
     l k c bp
           (Hc: cal_qlock_status l k c),
      cal_qlock_status (TEVENT c (TSHARED (OBUFFERE bp)) :: l) k c.
  Proof.
    unfold cal_qlock_status; intros.
    Local Transparent Q_CalLock.
    simpl in Hcal. subdestruct; subst.
    - inv Hcal. exploit Hc; eauto.
      intros (HT1 & HT2).
      split; intros HP; inv HP.
      + rewrite HT1; trivial.
      + rewrite HT2; trivial.
    - inv Hcal. exploit Hc; eauto.
      intros (HT1 & HT2).
      split; intros HP; inv HP.
      + rewrite HT1; trivial.
      + rewrite HT2; trivial.
        Local Opaque Q_CalLock.
  Qed.

  Global Instance page_copy_inv: PreservesInvariants page_copy´´_spec.
  Proof.
    preserves_invariants_simpl_auto.
    - eapply valid_qlock_pool_gss´; eauto.
    - simpl. eapply correct_qlock_pool_status_gss´; eauto.
      eapply correct_qlock_status_shared; eauto.
    - simpl. eapply cal_qlock_pool_status_gss´; eauto.
      eapply cal_qlock_status_buffer; eauto.
    - simpl. eapply get_free_qlock_pool_gss; eauto.
      eapply Q_CalLock_append in H10.
      destruct H10 as (self_c & other_c & b0 & q & tq & Hcal).
      unfold correct_qlock_pool_status in ×.
      exploit valid_qlock_pool_status_inv0; eauto.
      rewrite H8.
      intros (_ & HT).
      exploit HT; eauto. intros (_ & _ & ? & _); subst.
      eapply get_free_qlock_shared; eauto.
    - eapply qlock_pool_length_gss; eauto.
      eapply qlock_length_cons; eauto.
    - eapply valid_last_pool_gss; eauto.
      eapply valid_last_event; eauto.
    - eapply valid_from_pool_gss; eauto.
      eapply valid_from_event; eauto.
  Qed.

  Definition mqticketlockop_fresh : compatlayer (cdata RData) :=
    .

  Definition mqticketlockop_passthrough : compatlayer (cdata RData) :=
    fload gensem fload´_spec
           fstore gensem fstore´_spec
          
           page_copy gensem page_copy´´_spec
           page_copy_back gensem page_copy_back´_spec

           vmxinfo_get gensem vmxinfo_get_spec
           set_pg gensem setPG_spec
           set_cr3 setCR3_compatsem setCR3_spec
           get_size gensem MMSize
           is_usable gensem is_mm_usable_spec
           get_mms gensem get_mm_s_spec
           get_mml gensem get_mm_l_spec
           get_CPU_ID gensem get_CPU_ID_spec
           release_shared primcall_release_lock_compatsem release_shared release_shared1_spec0
           acquire_shared primcall_acquire_shared_compatsem acquire_shared1_spec0
           get_curid gensem get_curid_spec
           set_curid gensem set_curid_spec
           set_curid_init gensem set_curid_init_spec

           ticket_lock_init gensem ticket_lock_init0_spec
           pass_lock gensem pass_qlock_spec
           wait_lock gensem wait_qlock_spec
          

           trap_in primcall_general_compatsem trapin_spec
           trap_out primcall_general_compatsem trapout´_spec
           host_in primcall_general_compatsem hostin_spec
           host_out primcall_general_compatsem hostout´_spec
           proc_create_postinit gensem proc_create_postinit_spec
           trap_get primcall_trap_info_get_compatsem trap_info_get_spec
           trap_set primcall_trap_info_ret_compatsem trap_info_ret_spec
           serial_irq_check gensem serial_irq_check_spec
           iret gensem iret_spec
           cli gensem cli_spec
           sti gensem sti_spec
           serial_irq_current gensem serial_irq_current_spec
           ic_intr gensem ic_intr_spec
           save_context primcall_save_context_compatsem save_context_spec
           restore_context primcall_restore_context_compatsem restore_context_spec
           local_irq_save gensem local_irq_save_spec
           local_irq_restore gensem local_irq_restore_spec
           serial_in gensem serial_in_spec
serial device
           serial_out gensem serial_out_spec
serial device
           serial_hw_intr gensem serial_hw_intr_spec
serial device
           ioapic_read gensem ioapic_read_spec
ioapic device
           ioapic_write gensem ioapic_write_spec
ioapic device
           lapic_read gensem lapic_read_spec
lapic device
           lapic_write gensem lapic_write_spec
lapic device
           accessors {| exec_load := @exec_loadex; exec_store := @exec_storeex |}.

    Definition mqticketlockop :=
      mqticketlockop_fresh mqticketlockop_passthrough.

End WITHMEM.