-
Notifications
You must be signed in to change notification settings - Fork 1
/
table_unmap2.v
74 lines (59 loc) · 3.19 KB
/
table_unmap2.v
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
Require Import CodeProofDeps.
Require Import Ident.
Require Import Constants.
Require Import RData.
Require Import EventReplay.
Require Import MoverTypes.
Require Import CommonLib.
Require Import TableDataOpsIntro.Spec.
Require Import TableDataOpsRef1.Spec.
Require Import TableDataOpsRef1.Layer.
Require Import TableDataOpsRef2.Code.table_unmap2.
Require Import TableDataOpsRef2.LowSpecs.table_unmap2.
Local Open Scope Z_scope.
Section CodeProof.
Context `{real_params: RealParams}.
Context {memb} `{Hmemx: Mem.MemoryModelX memb}.
Context `{Hmwd: UseMemWithData memb}.
Let mem := mwd (cdata RData).
Context `{Hstencil: Stencil}.
Context `{make_program_ops: !MakeProgramOps Clight.function type Clight.fundef type}.
Context `{Hmake_program: !MakeProgram Clight.function type Clight.fundef type}.
Let L : compatlayer (cdata RData) :=
_table_unmap ↦ gensem table_unmap_spec
⊕ _table_unmap1 ↦ gensem table_unmap1_spec
.
Local Instance: ExternalCallsOps mem := CompatExternalCalls.compatlayer_extcall_ops L.
Local Instance: CompilerConfigOps mem := CompatExternalCalls.compatlayer_compiler_config_ops L.
Section BodyProof.
Context `{Hwb: WritableBlockOps}.
Variable (sc: stencil).
Variables (ge: genv) (STENCIL_MATCHES: stencil_matches sc ge).
Variable b_table_unmap: block.
Hypothesis h_table_unmap_s : Genv.find_symbol ge _table_unmap = Some b_table_unmap.
Hypothesis h_table_unmap_p : Genv.find_funct_ptr ge b_table_unmap
= Some (External (EF_external _table_unmap
(signature_of_type (Tcons Tptr (Tcons tulong (Tcons tulong Tnil))) tulong cc_default))
(Tcons Tptr (Tcons tulong (Tcons tulong Tnil))) tulong cc_default).
Local Opaque table_unmap_spec.
Variable b_table_unmap1: block.
Hypothesis h_table_unmap1_s : Genv.find_symbol ge _table_unmap1 = Some b_table_unmap1.
Hypothesis h_table_unmap1_p : Genv.find_funct_ptr ge b_table_unmap1
= Some (External (EF_external _table_unmap1
(signature_of_type (Tcons Tptr (Tcons tulong (Tcons tulong Tnil))) tulong cc_default))
(Tcons Tptr (Tcons tulong (Tcons tulong Tnil))) tulong cc_default).
Local Opaque table_unmap1_spec.
Lemma table_unmap2_body_correct:
forall m d d' env le g_rd_base g_rd_offset map_addr level res
(Henv: env = PTree.empty _)
(Hinv: high_level_invariant d)
(HPTg_rd: PTree.get _g_rd le = Some (Vptr g_rd_base (Int.repr g_rd_offset)))
(HPTmap_addr: PTree.get _map_addr le = Some (Vlong map_addr))
(HPTlevel: PTree.get _level le = Some (Vlong level))
(Hspec: table_unmap2_spec0 (g_rd_base, g_rd_offset) (VZ64 (Int64.unsigned map_addr)) (VZ64 (Int64.unsigned level)) d = Some (d', VZ64 (Int64.unsigned res))),
exists le', (exec_stmt ge env le ((m, d): mem) table_unmap2_body E0 le' (m, d') (Out_return (Some (Vlong res, tulong)))).
Proof.
solve_code_proof Hspec table_unmap2_body; eexists; solve_proof_low.
Qed.
End BodyProof.
End CodeProof.