Library mcertikos.proc.QueueIntroGenLemma


This file provide the contextual refinement proof between MBoot layer and MALInit layer
Require Import QueueIntroGenDef.
Require Import XOmega.
Require Export ByteListProp.

Definition of the refinement relation

Section Refinement.

  Section WITHMEM.

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

    Lemma H_CalLock_relate_TCB_log:
       l s p
             (Hcal: H_CalLock l = Some p)
             (Hrel: relate_TCB_Log s l ),
        H_CalLock = Some p.
    Proof.
      induction l; intros; simpl.
      - inv Hrel. eauto.
      - inv Hrel. inv Hrelate_event.
        + Local Transparent H_CalLock.
          simpl in ×.
          subdestruct. inv Hcal.
          exploit IHl; eauto.
          intros Hcal´. rewrite Hcal´.
          rewrite zeq_true. trivial.
        + simpl in ×.
          subdestruct; inv Hcal.
          × exploit IHl; eauto.
            intros Hcal´. rewrite Hcal´.
            rewrite zeq_true. trivial.
          × exploit IHl; eauto.
            intros Hcal´. rewrite Hcal´. trivial.
        + simpl in ×.
          subdestruct; inv Hcal.
          × exploit IHl; eauto.
            intros Hcal´. rewrite Hcal´.
            rewrite zeq_true. trivial.
    Qed.

    Lemma relate_TCB_Log_cons:
       l1 l1´ l2 l2´ s
             (Hrel1: relate_TCB_Log s l1 l2)
             (Hrel2: relate_TCB_Log s l1´ l2´),
        relate_TCB_Log s (l1 ++ l1´) (l2 ++ l2´).
    Proof.
      induction l1; simpl; intros.
      - inv Hrel1. eapply Hrel2.
      - inv Hrel1. exploit IHl1; eauto.
        intros Hrel3. simpl.
        econstructor; eauto.
    Qed.

    Lemma relate_TCB_Log_oracle_cons:
       c o1 o2 l1 l2 s i
             (Hrel1: relate_TCB_Log s l1 l2)
             (Hrel2: relate_TCB_Oracle_Pool s o1 o2)
             (Hrange: 0 i < num_chan),
        relate_TCB_Log s (ZMap.get (i + ID_AT_range) o1 c l1 ++ l1) (ZMap.get (i + ID_AT_range) o2 c l2 ++ l2).
    Proof.
      intros. eapply relate_TCB_Log_cons; eauto.
      inv Hrel2. specialize (H0 _ Hrange).
      inv H0. eauto.
    Qed.

    Lemma relate_TCB_Log_GetSharedTCB:
       l1 l2 s
             (Hrel1: relate_TCB_Log s l1 l2),
        ( tb tq, GetSharedTCB l1 = Some (tb, tq)
                        m, GetSharedMemEvent l2 = Some m
                                  relate_TCB_Event s (OTCBE tb tq) (OMEME m))
         (GetSharedTCB l1 = NoneGetSharedMemEvent l2 = None).
    Proof.
      induction l1; simpl; intros.
      - split; intros; try congruence. inv Hrel1. trivial.
      - split; intros.
        + inv Hrel1. inv Hrelate_event.
          × inv H. simpl.
            refine_split´; trivial.
          × simpl. exploit IHl1; eauto.
            intros (HP & _). eapply HP; eauto.
          × simpl. exploit IHl1; eauto.
            intros (HP & _). eapply HP; eauto.
        + inv Hrel1. inv Hrelate_event; try congruence.
          × simpl. exploit IHl1; eauto.
            intros (_ & HP). eapply HP; eauto.
          × simpl. exploit IHl1; eauto.
            intros (_ & HP). eapply HP; eauto.
    Qed.

    Lemma relate_TCB_Log_GetSharedTCB_Some:
       l1 l2 s tb tq
             (Hget: GetSharedTCB l1 = Some (tb, tq))
             (Hrel1: relate_TCB_Log s l1 l2),
       m, GetSharedMemEvent l2 = Some m
                 relate_TCB_Event s (OTCBE tb tq) (OMEME m).
    Proof.
      intros. exploit relate_TCB_Log_GetSharedTCB; eauto.
      intros (HP & _). exploit HP; eauto.
    Qed.

    Lemma relate_TCB_Log_GetSharedTCB_None:
       l1 l2 s
             (Hget: GetSharedTCB l1 = None)
             (Hrel1: relate_TCB_Log s l1 l2),
        GetSharedMemEvent l2 = None.
    Proof.
      intros. exploit relate_TCB_Log_GetSharedTCB; eauto.
      intros (_ & HP). exploit HP; eauto.
    Qed.

    Lemma relate_TCB_Log_ticket:
       l1 l2 s c e
             (Hrel1: relate_TCB_Log s l1 l2),
        relate_TCB_Log s (TEVENT c (TTICKET e) :: l1) (TEVENT c (TTICKET e) :: l2).
    Proof.
      intros. econstructor; trivial. econstructor.
    Qed.

    Lemma relate_TCB_Log_pull:
       l1 l2 s c
             (Hrel1: relate_TCB_Log s l1 l2),
        relate_TCB_Log s (TEVENT c (TSHARED OPULL) :: l1) (TEVENT c (TSHARED OPULL) :: l2).
    Proof.
      intros. econstructor; trivial. econstructor.
    Qed.

    Local Opaque H_CalLock.

    Lemma index2Z_eq:
       i (Hrange: 0 i < num_chan),
        index2Z lock_TCB_start i = Some (i + ID_AT_range).
    Proof.
      intros. unfold index2Z.
      Local Opaque Z.add. simpl.
      unfold ID_TCB_range. rewrite zle_lt_true; trivial.
      f_equal. omega.
      Local Transparent Z.add.
    Qed.

    Lemma index2Z_neq:
       i ofs r k
             (Hneq: i 1)
             (Hrange: 0 k < num_chan)
             (Hin: index2Z i ofs = Some r),
        r k + ID_AT_range.
    Proof.
      intros. functional inversion Hin.
      functional inversion H0.
      - subst. functional inversion H1; subst.
        unfold ID_AT_range in ×. omega.
      - omega.
      - subst. functional inversion H1; subst.
        unfold lock_TCB_range, ID_AT_range, ID_TCB_range. omega.
    Qed.

    Lemma neq_add:
       a b c,
        a b
        a + c b + c.
    Proof.
      intros. omega.
    Qed.

  End WITHMEM.

End Refinement.