Library mcertikos.virt.intel.VMCSIntroGenDef
*********************************************************************** * * * The CertiKOS Certified Kit Operating System * * * * The FLINT Group, Yale University * * * * Copyright The FLINT Group, Yale University. All rights reserved. * * This file is distributed under the terms of the Yale University * * Non-Commercial License Agreement. * * * ***********************************************************************
This file provide the contextual refinement proof between PIPC layer and PUCtxtIntro layer
Require Export Coqlib.
Require Export Errors.
Require Export AST.
Require Export Integers.
Require Export Floats.
Require Export Op.
Require Export Asm.
Require Export Events.
Require Export Globalenvs.
Require Export Smallstep.
Require Export Values.
Require Export Memory.
Require Export Maps.
Require Export CommonTactic.
Require Export AuxLemma.
Require Export FlatMemory.
Require Export AuxStateDataType.
Require Export Constant.
Require Export GlobIdent.
Require Export RealParams.
Require Export LoadStoreSem3.
Require Export AsmImplLemma.
Require Export GenSem.
Require Export RefinementTactic.
Require Export PrimSemantics.
Require Export XOmega.
Require Export liblayers.logic.PTreeModules.
Require Export liblayers.logic.LayerLogicImpl.
Require Export liblayers.compcertx.Stencil.
Require Export liblayers.compcertx.MakeProgram.
Require Export liblayers.compat.CompatLayers.
Require Export liblayers.compat.CompatGenSem.
Require Export compcert.cfrontend.Ctypes.
Require Export LAsmModuleSemAux.
Require Export LayerCalculusLemma.
Require Export AbstractDataType.
Require Export VVMCSIntro.
Notation HDATA := RData.
Notation LDATA := RData.
Notation HDATAOps := (cdata (cdata_ops := pproc_data_ops) HDATA).
Notation LDATAOps := (cdata (cdata_ops := pproc_data_ops) LDATA).
Open Scope string_scope.
Open Scope error_monad_scope.
Open Scope Z_scope.
Require Export Errors.
Require Export AST.
Require Export Integers.
Require Export Floats.
Require Export Op.
Require Export Asm.
Require Export Events.
Require Export Globalenvs.
Require Export Smallstep.
Require Export Values.
Require Export Memory.
Require Export Maps.
Require Export CommonTactic.
Require Export AuxLemma.
Require Export FlatMemory.
Require Export AuxStateDataType.
Require Export Constant.
Require Export GlobIdent.
Require Export RealParams.
Require Export LoadStoreSem3.
Require Export AsmImplLemma.
Require Export GenSem.
Require Export RefinementTactic.
Require Export PrimSemantics.
Require Export XOmega.
Require Export liblayers.logic.PTreeModules.
Require Export liblayers.logic.LayerLogicImpl.
Require Export liblayers.compcertx.Stencil.
Require Export liblayers.compcertx.MakeProgram.
Require Export liblayers.compat.CompatLayers.
Require Export liblayers.compat.CompatGenSem.
Require Export compcert.cfrontend.Ctypes.
Require Export LAsmModuleSemAux.
Require Export LayerCalculusLemma.
Require Export AbstractDataType.
Require Export VVMCSIntro.
Notation HDATA := RData.
Notation LDATA := RData.
Notation HDATAOps := (cdata (cdata_ops := pproc_data_ops) HDATA).
Notation LDATAOps := (cdata (cdata_ops := pproc_data_ops) LDATA).
Open Scope string_scope.
Open Scope error_monad_scope.
Open Scope Z_scope.
Section Refinement.
Context `{real_params: RealParams}.
Section WITHMEM.
Context `{Hstencil: Stencil}.
Context `{Hmem: Mem.MemoryModel}.
Context `{Hmwd: UseMemWithData mem}.
Context `{real_params: RealParams}.
Section WITHMEM.
Context `{Hstencil: Stencil}.
Context `{Hmem: Mem.MemoryModel}.
Context `{Hmwd: UseMemWithData mem}.
Relation between the kernel context pool and the underline memory
Inductive match_VMCS: stencil → VMCS → mem → meminj → Prop :=
| MATCH_VMCS:
∀ vvm m b s abrtid revid f,
(∀ i,
0 ≤ i ≤ C_VMCS_HOST_RIP + 1 →
(∃ v,
Mem.load Mint32 m b (i × 4 + 8) = Some v ∧
Mem.valid_access m Mint32 b (i × 4 + 8) Writable ∧
val_inject f (Val.load_result Mint32 (ZMap.get i vvm)) v)) →
Mem.load Mint32 m b 0 = Some (Vint revid) →
Mem.valid_access m Mint32 b 0 Writable →
Mem.load Mint32 m b 4 = Some (Vint abrtid) →
Mem.valid_access m Mint32 b 4 Writable →
find_symbol s VMCS_LOC = Some b →
match_VMCS s (VMCSValid (Vint revid) (Vint abrtid) vvm) m f.
| MATCH_VMCS:
∀ vvm m b s abrtid revid f,
(∀ i,
0 ≤ i ≤ C_VMCS_HOST_RIP + 1 →
(∃ v,
Mem.load Mint32 m b (i × 4 + 8) = Some v ∧
Mem.valid_access m Mint32 b (i × 4 + 8) Writable ∧
val_inject f (Val.load_result Mint32 (ZMap.get i vvm)) v)) →
Mem.load Mint32 m b 0 = Some (Vint revid) →
Mem.valid_access m Mint32 b 0 Writable →
Mem.load Mint32 m b 4 = Some (Vint abrtid) →
Mem.valid_access m Mint32 b 4 Writable →
find_symbol s VMCS_LOC = Some b →
match_VMCS s (VMCSValid (Vint revid) (Vint abrtid) vvm) m f.
Relation between the new raw data at the higher layer with the mememory at lower layer
Inductive match_RData: stencil → HDATA → mem → meminj → Prop :=
| MATCH_RDATA:
∀ hadt m s f,
match_VMCS s (vmcs hadt) m f
→ match_RData s hadt m f.
| MATCH_RDATA:
∀ hadt m s f,
match_VMCS s (vmcs hadt) m f
→ match_RData s hadt m f.
Relation between raw data at two layers
Record relate_RData (f: meminj) (hadt: HDATA) (ladt: LDATA) :=
mkrelate_RData {
flatmem_re: FlatMem.flatmem_inj (HP hadt) (HP ladt);
vmxinfo_re: vmxinfo hadt = vmxinfo ladt;
CR3_re: CR3 hadt = CR3 ladt;
ikern_re: ikern hadt = ikern ladt;
pg_re: pg hadt = pg ladt;
ihost_re: ihost hadt = ihost ladt;
AC_re: AC hadt = AC ladt;
ti_fst_re: (fst (ti hadt)) = (fst (ti ladt));
ti_snd_re: val_inject f (snd (ti hadt)) (snd (ti ladt));
LAT_re: LAT hadt = LAT ladt;
nps_re: nps hadt = nps ladt;
init_re: init hadt = init ladt;
pperm_re: pperm hadt = pperm ladt;
PT_re: PT hadt = PT ladt;
ptp_re: ptpool hadt = ptpool ladt;
idpde_re: idpde hadt = idpde ladt;
ipt_re: ipt hadt = ipt ladt;
smspool_re: smspool hadt = smspool ladt;
pb_re: pb hadt = pb ladt;
kctxt_re: kctxt_inj f num_proc (kctxt hadt) (kctxt ladt);
abtcb_re: abtcb hadt = abtcb ladt;
abq_re: abq hadt = abq ladt;
cid_re: cid hadt = cid ladt;
chpool_re: syncchpool hadt = syncchpool ladt;
uctxt_re: uctxt_inj f (uctxt hadt) (uctxt ladt);
ept_re: ept hadt = ept ladt;
com1_re: com1 hadt = com1 ladt;
console_re: console hadt = console ladt;
console_concrete_re: console_concrete hadt = console_concrete ladt;
ioapic_re: ioapic ladt = ioapic hadt;
lapic_re: lapic ladt = lapic hadt;
intr_flag_re: intr_flag ladt = intr_flag hadt;
curr_intr_num_re: curr_intr_num ladt = curr_intr_num hadt;
in_intr_re: in_intr hadt = in_intr ladt;
drv_serial_re: drv_serial hadt = drv_serial ladt
}.
Global Instance rel_ops: CompatRelOps HDATAOps LDATAOps :=
{
relate_AbData s f d1 d2 := relate_RData f d1 d2;
match_AbData s d1 m f := match_RData s d1 m f;
new_glbl := VMCS_LOC :: nil
}.
End REFINEMENT_REL.
mkrelate_RData {
flatmem_re: FlatMem.flatmem_inj (HP hadt) (HP ladt);
vmxinfo_re: vmxinfo hadt = vmxinfo ladt;
CR3_re: CR3 hadt = CR3 ladt;
ikern_re: ikern hadt = ikern ladt;
pg_re: pg hadt = pg ladt;
ihost_re: ihost hadt = ihost ladt;
AC_re: AC hadt = AC ladt;
ti_fst_re: (fst (ti hadt)) = (fst (ti ladt));
ti_snd_re: val_inject f (snd (ti hadt)) (snd (ti ladt));
LAT_re: LAT hadt = LAT ladt;
nps_re: nps hadt = nps ladt;
init_re: init hadt = init ladt;
pperm_re: pperm hadt = pperm ladt;
PT_re: PT hadt = PT ladt;
ptp_re: ptpool hadt = ptpool ladt;
idpde_re: idpde hadt = idpde ladt;
ipt_re: ipt hadt = ipt ladt;
smspool_re: smspool hadt = smspool ladt;
pb_re: pb hadt = pb ladt;
kctxt_re: kctxt_inj f num_proc (kctxt hadt) (kctxt ladt);
abtcb_re: abtcb hadt = abtcb ladt;
abq_re: abq hadt = abq ladt;
cid_re: cid hadt = cid ladt;
chpool_re: syncchpool hadt = syncchpool ladt;
uctxt_re: uctxt_inj f (uctxt hadt) (uctxt ladt);
ept_re: ept hadt = ept ladt;
com1_re: com1 hadt = com1 ladt;
console_re: console hadt = console ladt;
console_concrete_re: console_concrete hadt = console_concrete ladt;
ioapic_re: ioapic ladt = ioapic hadt;
lapic_re: lapic ladt = lapic hadt;
intr_flag_re: intr_flag ladt = intr_flag hadt;
curr_intr_num_re: curr_intr_num ladt = curr_intr_num hadt;
in_intr_re: in_intr hadt = in_intr ladt;
drv_serial_re: drv_serial hadt = drv_serial ladt
}.
Global Instance rel_ops: CompatRelOps HDATAOps LDATAOps :=
{
relate_AbData s f d1 d2 := relate_RData f d1 d2;
match_AbData s d1 m f := match_RData s d1 m f;
new_glbl := VMCS_LOC :: nil
}.
End REFINEMENT_REL.
Section Rel_Property.
Lemma inject_match_correct:
∀ s d1 m2 f m2´ j,
match_RData s d1 m2 f →
Mem.inject j m2 m2´ →
inject_incr (Mem.flat_inj (genv_next s)) j →
match_RData s d1 m2´ (compose_meminj f j).
Proof.
inversion 1; subst; intros.
inv H0.
assert (HFB0: j b = Some (b, 0)).
{
eapply stencil_find_symbol_inject´; eauto.
}
econstructor; eauto; intros. inv H3.
econstructor; eauto; intros.
- specialize (H4 _ H0).
destruct H4 as [v1[HL1[HV1 HM]]].
specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H1 HL1 HFB0).
repeat rewrite Z.add_0_r.
intros [v1´[HLD1´ HV1´]].
refine_split´; eauto.
specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H1 HV1).
rewrite Z.add_0_r; trivial.
eapply val_inject_compose; eauto.
- specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H1 H5 HFB0).
repeat rewrite Z.add_0_r.
intros [v1´[HLD1´ HV1´]]. inv HV1´. trivial.
- specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H1 H6).
rewrite Z.add_0_r; trivial.
- specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H1 H7 HFB0).
repeat rewrite Z.add_0_r.
intros [v1´[HLD1´ HV1´]]. inv HV1´. trivial.
- specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H1 H8).
rewrite Z.add_0_r; trivial.
Qed.
Lemma store_match_correct:
∀ s abd m0 m0´ f b2 v v´ chunk,
match_RData s abd m0 f →
(∀ i b,
In i new_glbl →
find_symbol s i = Some b → b ≠ b2) →
Mem.store chunk m0 b2 v v´ = Some m0´ →
match_RData s abd m0´ f.
Proof.
intros. inv H. inv H2.
econstructor; eauto. inv H.
econstructor; eauto.
- intros. specialize (H3 _ H).
destruct H3 as [v1[HL1[HV1 HM]]].
eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_store_other _ _ _ _ _ _ H1); auto.
refine_split´; eauto;
eapply Mem.store_valid_access_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_store_other _ _ _ _ _ _ H1); auto.
- eapply Mem.store_valid_access_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_store_other _ _ _ _ _ _ H1); auto.
- eapply Mem.store_valid_access_1; eauto.
Qed.
Lemma storebytes_match_correct:
∀ s abd m0 m0´ f b2 v v´,
match_RData s abd m0 f →
(∀ i b,
In i new_glbl →
find_symbol s i = Some b → b ≠ b2) →
Mem.storebytes m0 b2 v v´ = Some m0´ →
match_RData s abd m0´ f.
Proof.
intros. inv H. inv H2.
econstructor; eauto. inv H.
econstructor; eauto.
- intros. specialize (H3 _ H).
destruct H3 as [v1[HL1[HV1 HM]]].
eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
refine_split´; eauto;
eapply Mem.storebytes_valid_access_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
- eapply Mem.storebytes_valid_access_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
- eapply Mem.storebytes_valid_access_1; eauto.
Qed.
Lemma free_match_correct:
∀ s abd m0 m0´ f ofs sz b2,
match_RData s abd m0 f→
(∀ i b,
In i new_glbl →
find_symbol s i = Some b → b ≠ b2) →
Mem.free m0 b2 ofs sz = Some m0´ →
match_RData s abd m0´ f.
Proof.
intros; inv H; inv H2.
econstructor; eauto. inv H.
econstructor; eauto.
- intros. specialize (H3 _ H).
destruct H3 as [v1[HL1[HV1 HM]]].
eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_free _ _ _ _ _ H1); auto.
refine_split´; eauto;
eapply Mem.valid_access_free_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_free _ _ _ _ _ H1); auto.
- eapply H0 in H9; simpl; eauto.
eapply Mem.valid_access_free_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_free _ _ _ _ _ H1); auto.
- eapply H0 in H9; simpl; eauto.
eapply Mem.valid_access_free_1; eauto.
Qed.
Lemma alloc_match_correct:
∀ s abd m´0 m´1 f f´ ofs sz b0 b´1,
match_RData s abd m´0 f→
Mem.alloc m´0 ofs sz = (m´1, b´1) →
f´ b0 = Some (b´1, 0%Z) →
(∀ b : block, b ≠ b0 → f´ b = f b) →
inject_incr f f´ →
(∀ i b,
In i new_glbl →
find_symbol s i = Some b → b ≠ b0) →
match_RData s abd m´1 f´.
Proof.
intros. rename H1 into HF1, H2 into HB. inv H; inv H1.
econstructor; eauto. inv H.
econstructor; eauto.
- intros. specialize (H2 _ H).
destruct H2 as [v1[HL1[HV1 HM]]].
refine_split´; eauto;
try (apply (Mem.load_alloc_other _ _ _ _ _ H0));
try (eapply Mem.valid_access_alloc_other); eauto.
- apply (Mem.load_alloc_other _ _ _ _ _ H0); eauto.
- eapply Mem.valid_access_alloc_other; eauto.
- apply (Mem.load_alloc_other _ _ _ _ _ H0); eauto.
- eapply Mem.valid_access_alloc_other; eauto.
Qed.
Lemma inject_match_correct:
∀ s d1 m2 f m2´ j,
match_RData s d1 m2 f →
Mem.inject j m2 m2´ →
inject_incr (Mem.flat_inj (genv_next s)) j →
match_RData s d1 m2´ (compose_meminj f j).
Proof.
inversion 1; subst; intros.
inv H0.
assert (HFB0: j b = Some (b, 0)).
{
eapply stencil_find_symbol_inject´; eauto.
}
econstructor; eauto; intros. inv H3.
econstructor; eauto; intros.
- specialize (H4 _ H0).
destruct H4 as [v1[HL1[HV1 HM]]].
specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H1 HL1 HFB0).
repeat rewrite Z.add_0_r.
intros [v1´[HLD1´ HV1´]].
refine_split´; eauto.
specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H1 HV1).
rewrite Z.add_0_r; trivial.
eapply val_inject_compose; eauto.
- specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H1 H5 HFB0).
repeat rewrite Z.add_0_r.
intros [v1´[HLD1´ HV1´]]. inv HV1´. trivial.
- specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H1 H6).
rewrite Z.add_0_r; trivial.
- specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H1 H7 HFB0).
repeat rewrite Z.add_0_r.
intros [v1´[HLD1´ HV1´]]. inv HV1´. trivial.
- specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H1 H8).
rewrite Z.add_0_r; trivial.
Qed.
Lemma store_match_correct:
∀ s abd m0 m0´ f b2 v v´ chunk,
match_RData s abd m0 f →
(∀ i b,
In i new_glbl →
find_symbol s i = Some b → b ≠ b2) →
Mem.store chunk m0 b2 v v´ = Some m0´ →
match_RData s abd m0´ f.
Proof.
intros. inv H. inv H2.
econstructor; eauto. inv H.
econstructor; eauto.
- intros. specialize (H3 _ H).
destruct H3 as [v1[HL1[HV1 HM]]].
eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_store_other _ _ _ _ _ _ H1); auto.
refine_split´; eauto;
eapply Mem.store_valid_access_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_store_other _ _ _ _ _ _ H1); auto.
- eapply Mem.store_valid_access_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_store_other _ _ _ _ _ _ H1); auto.
- eapply Mem.store_valid_access_1; eauto.
Qed.
Lemma storebytes_match_correct:
∀ s abd m0 m0´ f b2 v v´,
match_RData s abd m0 f →
(∀ i b,
In i new_glbl →
find_symbol s i = Some b → b ≠ b2) →
Mem.storebytes m0 b2 v v´ = Some m0´ →
match_RData s abd m0´ f.
Proof.
intros. inv H. inv H2.
econstructor; eauto. inv H.
econstructor; eauto.
- intros. specialize (H3 _ H).
destruct H3 as [v1[HL1[HV1 HM]]].
eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
refine_split´; eauto;
eapply Mem.storebytes_valid_access_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
- eapply Mem.storebytes_valid_access_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
- eapply Mem.storebytes_valid_access_1; eauto.
Qed.
Lemma free_match_correct:
∀ s abd m0 m0´ f ofs sz b2,
match_RData s abd m0 f→
(∀ i b,
In i new_glbl →
find_symbol s i = Some b → b ≠ b2) →
Mem.free m0 b2 ofs sz = Some m0´ →
match_RData s abd m0´ f.
Proof.
intros; inv H; inv H2.
econstructor; eauto. inv H.
econstructor; eauto.
- intros. specialize (H3 _ H).
destruct H3 as [v1[HL1[HV1 HM]]].
eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_free _ _ _ _ _ H1); auto.
refine_split´; eauto;
eapply Mem.valid_access_free_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_free _ _ _ _ _ H1); auto.
- eapply H0 in H9; simpl; eauto.
eapply Mem.valid_access_free_1; eauto.
- eapply H0 in H9; simpl; eauto.
repeat rewrite (Mem.load_free _ _ _ _ _ H1); auto.
- eapply H0 in H9; simpl; eauto.
eapply Mem.valid_access_free_1; eauto.
Qed.
Lemma alloc_match_correct:
∀ s abd m´0 m´1 f f´ ofs sz b0 b´1,
match_RData s abd m´0 f→
Mem.alloc m´0 ofs sz = (m´1, b´1) →
f´ b0 = Some (b´1, 0%Z) →
(∀ b : block, b ≠ b0 → f´ b = f b) →
inject_incr f f´ →
(∀ i b,
In i new_glbl →
find_symbol s i = Some b → b ≠ b0) →
match_RData s abd m´1 f´.
Proof.
intros. rename H1 into HF1, H2 into HB. inv H; inv H1.
econstructor; eauto. inv H.
econstructor; eauto.
- intros. specialize (H2 _ H).
destruct H2 as [v1[HL1[HV1 HM]]].
refine_split´; eauto;
try (apply (Mem.load_alloc_other _ _ _ _ _ H0));
try (eapply Mem.valid_access_alloc_other); eauto.
- apply (Mem.load_alloc_other _ _ _ _ _ H0); eauto.
- eapply Mem.valid_access_alloc_other; eauto.
- apply (Mem.load_alloc_other _ _ _ _ _ H0); eauto.
- eapply Mem.valid_access_alloc_other; eauto.
Qed.
Prove that after taking one step, the refinement relation still holds
Lemma relate_incr:
∀ abd abd´ f f´,
relate_RData f abd abd´
→ inject_incr f f´
→ relate_RData f´ abd abd´.
Proof.
inversion 1; subst; intros; inv H; constructor; eauto.
- eapply kctxt_inj_incr; eauto.
- eapply uctxt_inj_incr; eauto.
Qed.
Global Instance rel_prf: CompatRel HDATAOps LDATAOps.
Proof.
constructor.
- apply inject_match_correct.
- apply store_match_correct.
- apply alloc_match_correct.
- apply free_match_correct.
- apply storebytes_match_correct.
- intros. eapply relate_incr; eauto.
Qed.
End Rel_Property.
End WITHMEM.
End Refinement.
∀ abd abd´ f f´,
relate_RData f abd abd´
→ inject_incr f f´
→ relate_RData f´ abd abd´.
Proof.
inversion 1; subst; intros; inv H; constructor; eauto.
- eapply kctxt_inj_incr; eauto.
- eapply uctxt_inj_incr; eauto.
Qed.
Global Instance rel_prf: CompatRel HDATAOps LDATAOps.
Proof.
constructor.
- apply inject_match_correct.
- apply store_match_correct.
- apply alloc_match_correct.
- apply free_match_correct.
- apply storebytes_match_correct.
- intros. eapply relate_incr; eauto.
Qed.
End Rel_Property.
End WITHMEM.
End Refinement.