1
+ Require Import Coq.ZArith.ZArith.
1
2
Require Import coqutil.sanity coqutil.Macros.subst coqutil.Macros.unique coqutil.Byte.
2
3
Require Import coqutil.Datatypes.PrimitivePair coqutil.Datatypes.HList.
3
4
Require Import coqutil.Decidable.
4
5
Require Import coqutil.Tactics.fwd coqutil.Tactics.Tactics.
5
- Require Import bedrock2.Notations bedrock2.Syntax coqutil.Map.Interface coqutil.Map.OfListWord.
6
- Require Import BinIntDef coqutil.Word.Interface coqutil.Word.Bitwidth.
6
+ Require Import bedrock2.Notations bedrock2.Syntax.
7
+ Require Import coqutil.Map.Interface coqutil.Map.Properties coqutil.Map.OfListWord.
8
+ Require Import coqutil.Word.Interface coqutil.Word.Bitwidth.
7
9
Require Import bedrock2.MetricLogging.
8
- Require Import bedrock2.Memory.
10
+ Require Import bedrock2.Memory bedrock2.ptsto_bytes bedrock2.Map.Separation .
9
11
Require Import bedrock2.Semantics.
12
+ Require Import bedrock2.Map.DisjointUnion bedrock2.Map.split_alt.
10
13
11
14
Require Import Coq.Lists.List.
12
15
@@ -15,23 +18,96 @@ Section semantics.
15
18
Context {locals: map.map String.string word}.
16
19
Context {env: map.map String.string (list String.string * list String.string * cmd)}.
17
20
Context {ext_spec: ExtSpec}.
21
+ Context {mem_ok: map.ok mem} {word_ok: word.ok word}.
18
22
19
23
Lemma frame_load: forall mSmall mBig mAdd a r (v: word),
20
- map .split mBig mSmall mAdd ->
24
+ mmap .split mBig mSmall mAdd ->
21
25
load a mSmall r = Some v ->
22
26
load a mBig r = Some v.
23
27
Proof .
24
- Admitted .
28
+ unfold load, load_Z. intros. fwd. eapply sep_of_load_bytes in E0.
29
+ 2: destruct a; simpl; destruct width_cases as [W | W]; rewrite W; cbv; discriminate.
30
+ fwd. unfold sep in E0. fwd.
31
+ eapply map.split_alt in E0p0.
32
+ unfold mmap.split in *.
33
+ rewrite E0p0 in H.
34
+ rewrite mmap.du_assoc in H. unfold mmap.du in H at 1. fwd.
35
+ erewrite load_bytes_of_sep. 1: reflexivity.
36
+ unfold sep. do 2 eexists.
37
+ rewrite map.split_alt.
38
+ unfold mmap.split.
39
+ ssplit. 2: eassumption. all: simpl; exact H.
40
+ Qed .
41
+
42
+ (* TODO share with FlatToRiscvCommon *)
43
+
44
+ Lemma store_bytes_preserves_footprint: forall n a v (m m': mem),
45
+ Memory.store_bytes n m a v = Some m' ->
46
+ map.same_domain m m'.
47
+ Proof using word_ok mem_ok.
48
+ intros. unfold store_bytes, load_bytes, unchecked_store_bytes in *. fwd.
49
+ eapply map.putmany_of_tuple_preserves_domain; eauto .
50
+ Qed .
51
+
52
+ Lemma disjoint_putmany_preserves_store_bytes: forall n a vs (m1 m1' mq: mem),
53
+ store_bytes n m1 a vs = Some m1' ->
54
+ map.disjoint m1 mq ->
55
+ store_bytes n (map.putmany m1 mq) a vs = Some (map.putmany m1' mq).
56
+ Proof using word_ok mem_ok.
57
+ intros.
58
+ unfold store_bytes, load_bytes, unchecked_store_bytes in *. fwd.
59
+ erewrite map.getmany_of_tuple_in_disjoint_putmany by eassumption.
60
+ f_equal.
61
+ set (ks := (footprint a n)) in *.
62
+ rename mq into m2.
63
+ rewrite map.putmany_of_tuple_to_putmany.
64
+ rewrite (map.putmany_of_tuple_to_putmany n m1 ks vs).
65
+ apply map.disjoint_putmany_commutes.
66
+ pose proof map.getmany_of_tuple_to_sub_domain _ _ _ _ E as P.
67
+ apply map.sub_domain_value_indep with (vs2 := vs) in P.
68
+ set (mp := (map.putmany_of_tuple ks vs map.empty)) in *.
69
+ apply map.disjoint_comm.
70
+ eapply map.sub_domain_disjoint; eassumption.
71
+ Qed .
72
+
73
+ Lemma store_bytes_frame: forall {n: nat} {m1 m1' m: mem} {a: word} {v: HList.tuple byte n} {F},
74
+ Memory.store_bytes n m1 a v = Some m1' ->
75
+ (eq m1 * F)%sep m ->
76
+ exists m', (eq m1' * F)%sep m' /\ Memory.store_bytes n m a v = Some m'.
77
+ Proof using word_ok mem_ok.
78
+ intros.
79
+ unfold sep in H0.
80
+ destruct H0 as (mp & mq & A & B & C).
81
+ subst mp.
82
+ unfold map.split in A. destruct A as [A1 A2].
83
+ eexists (map.putmany m1' mq).
84
+ split.
85
+ - unfold sep.
86
+ exists m1', mq. repeat split; trivial.
87
+ apply store_bytes_preserves_footprint in H.
88
+ clear -H A2.
89
+ unfold map.disjoint, map.same_domain, map.sub_domain in *. destruct H as [P Q].
90
+ intros.
91
+ edestruct Q; eauto .
92
+ - subst m.
93
+ eauto using disjoint_putmany_preserves_store_bytes.
94
+ Qed .
25
95
26
96
Lemma frame_store: forall sz (mSmall mSmall' mBig mAdd: mem) a v,
27
- map .split mBig mSmall mAdd ->
97
+ mmap .split mBig mSmall mAdd ->
28
98
store sz mSmall a v = Some mSmall' ->
29
- exists mBig', map .split mBig' mSmall' mAdd /\ store sz mBig a v = Some mBig'.
99
+ exists mBig', mmap .split mBig' mSmall' mAdd /\ store sz mBig a v = Some mBig'.
30
100
Proof .
31
- Admitted .
101
+ intros. eapply (store_bytes_frame (F := eq mAdd)) in H0.
102
+ 2: {
103
+ unfold sep. do 2 eexists. ssplit. 2,3: reflexivity. eapply map.split_alt; exact H.
104
+ }
105
+ fwd. unfold store, store_Z. rewrite H0p1. eexists. split. 2: reflexivity.
106
+ unfold sep in H0p0. fwd. eapply map.split_alt. assumption.
107
+ Qed .
32
108
33
109
Lemma frame_eval_expr: forall l e mSmall mBig mAdd mc (v: word) mc',
34
- map .split mBig mSmall mAdd ->
110
+ mmap .split mBig mSmall mAdd ->
35
111
eval_expr mSmall l e mc = Some (v, mc') ->
36
112
eval_expr mBig l e mc = Some (v, mc').
37
113
Proof .
@@ -50,7 +126,7 @@ Section semantics.
50
126
51
127
Lemma frame_evaluate_call_args_log: forall l mSmall mBig mAdd arges
52
128
mc (args: list word) mc',
53
- map .split mBig mSmall mAdd ->
129
+ mmap .split mBig mSmall mAdd ->
54
130
evaluate_call_args_log mSmall l arges mc = Some (args, mc') ->
55
131
evaluate_call_args_log mBig l arges mc = Some (args, mc').
56
132
Proof .
@@ -60,14 +136,12 @@ Section semantics.
60
136
1: reflexivity. all: assumption.
61
137
Qed .
62
138
63
- Axiom TODO: False.
64
-
65
139
Lemma frame_exec: forall e c t mSmall l mc P,
66
140
exec e c t mSmall l mc P ->
67
141
forall mBig mAdd,
68
- map .split mBig mSmall mAdd ->
142
+ mmap .split mBig mSmall mAdd ->
69
143
exec e c t mBig l mc (fun t' mBig' l' mc' =>
70
- exists mSmall', map .split mBig' mSmall' mAdd /\ P t' mSmall' l' mc').
144
+ exists mSmall', mmap .split mBig' mSmall' mAdd /\ P t' mSmall' l' mc').
71
145
Proof .
72
146
induction 1; intros;
73
147
try match goal with
@@ -79,19 +153,36 @@ Section semantics.
79
153
intros.
80
154
rename mCombined into mCombinedBig.
81
155
specialize H1 with (1 := H3).
82
- assert (map.split (map.putmany mSmall mStack) mSmall mStack) as Sp by case TODO.
83
- specialize H1 with (1 := Sp).
84
- specialize (H1 mCombinedBig mAdd).
85
- assert (map.split mCombinedBig (map.putmany mSmall mStack) mAdd) as Sp' by case TODO.
86
- specialize (H1 Sp').
156
+ specialize (H1 (mmap.force (mmap.du (Some mSmall) (Some mStack)))).
157
+ eapply map.split_alt in H4. pose proof H4 as D0. unfold mmap.split in D0.
158
+ rewrite H2 in D0.
159
+ rewrite (mmap.du_comm (Some mSmall) (Some mAdd)) in D0.
160
+ rewrite mmap.du_assoc in D0. unfold mmap.du at 1 in D0. fwd.
161
+ change (Some mCombinedBig = mmap.du (Some mAdd) (Some r)) in D0.
87
162
eapply exec.weaken. 1: eapply H1.
163
+ { simpl. eapply map.split_alt. unfold mmap.split. symmetry. assumption. }
164
+ { unfold mmap.split. simpl. rewrite map.du_comm. exact D0. }
88
165
cbv beta. intros. fwd.
89
- assert (map.split m' (map.putmany mSmall'0 mAdd) mStack') as Sp'' by case TODO.
90
- eexists. exists mStack'. ssplit.
91
- 1,2: eassumption.
166
+ move D0 at bottom.
167
+ repeat match reverse goal with
168
+ | H: map.split _ _ _ |- _ => eapply map.split_alt in H
169
+ | H: mmap.split _ _ _ |- _ =>
170
+ unfold mmap.split in H;
171
+ let F := fresh "D0" in
172
+ rename H into F;
173
+ move F at bottom
174
+ end.
175
+ rewrite D1 in D2. clear D1 mBig. rewrite D4 in D3. clear D4 mSmall'.
176
+ rewrite mmap.du_assoc in D3. rewrite (mmap.du_comm (Some mStack')) in D3.
177
+ rewrite <- mmap.du_assoc in D3.
178
+ eexists (mmap.force (mmap.du (Some mSmall'0) (Some mAdd))). exists mStack'. ssplit.
179
+ 1: eassumption.
180
+ { simpl. eapply map.split_alt. unfold mmap.split.
181
+ rewrite D3. f_equal. unfold mmap.du at 1 in D3. fwd. simpl in E0. rewrite E0.
182
+ reflexivity. }
92
183
eexists; split. 2: eassumption.
93
- assert (map .split (map.putmany mSmall'0 mAdd) mSmall'0 mAdd) by case TODO .
94
- assumption . }
184
+ unfold mmap .split. simpl .
185
+ unfold mmap.du at 1 in D3. fwd. simpl in E0. rewrite E0. reflexivity . }
95
186
{ eapply exec.seq. 1: eapply IHexec; eassumption.
96
187
cbv beta. intros. fwd. eapply H1. 1: eassumption. assumption. }
97
188
{ eapply exec.while_true.
@@ -100,22 +191,47 @@ Section semantics.
100
191
{ eapply IHexec. 1: eassumption. }
101
192
cbv beta. intros. fwd. eauto. }
102
193
{ (* call *)
103
- case TODO. }
194
+ econstructor. 1: eassumption.
195
+ { eauto using frame_evaluate_call_args_log. }
196
+ 1: eassumption.
197
+ 1: eapply IHexec. 1: eassumption.
198
+ cbv beta. intros. fwd.
199
+ specialize H3 with (1 := H5p1). fwd. eauto 10. }
104
200
{ (* interact *)
105
- assert (map.split mBig (map.putmany mKeep mAdd) mGive) as Sp by case TODO.
106
- econstructor. 1: exact Sp.
201
+ eapply map.split_alt in H. pose proof H3 as A.
202
+ unfold mmap.split in H3, H. rewrite H in H3.
203
+ rewrite mmap.du_assoc in H3. rewrite (mmap.du_comm (Some mGive)) in H3.
204
+ rewrite <- mmap.du_assoc in H3.
205
+ assert (exists mKeepBig, Some mKeepBig = mmap.du (Some mKeep) (Some mAdd)) as Sp0. {
206
+ exists (mmap.force (map.du mKeep mAdd)).
207
+ unfold mmap.du in H3 at 1. fwd. simpl in E. rewrite E. reflexivity.
208
+ }
209
+ destruct Sp0 as (mKeepBig & Sp0).
210
+ assert (mmap.split mBig mKeepBig mGive) as Sp.
211
+ { unfold mmap.split. rewrite Sp0. assumption. }
212
+ econstructor. 1: eapply map.split_alt; exact Sp.
107
213
{ eauto using frame_evaluate_call_args_log. }
108
214
1: eassumption.
109
215
intros.
110
216
specialize H2 with (1 := H4). fwd.
111
217
eexists. split. 1: eassumption.
112
218
intros.
113
- assert (exists mSmall', map.split mSmall' mKeep mReceive) by case TODO.
114
- fwd. exists mSmall'.
115
- assert (map.split m' mSmall' mAdd) by case TODO.
116
- split.
117
- 1: assumption.
118
- eapply H2p1. assumption.
219
+ eapply map.split_alt in H2. unfold mmap.split in *.
220
+ assert (exists mSmall', mmap.split mSmall' mKeep mReceive) as Sp1. {
221
+ exists (mmap.force (map.du mKeep mReceive)).
222
+ eapply map.split_alt. rewrite Sp0 in H2.
223
+ rewrite mmap.du_assoc in H2. rewrite (mmap.du_comm (Some mAdd)) in H2.
224
+ rewrite <- mmap.du_assoc in H2.
225
+ unfold mmap.du at 1 in H2. fwd.
226
+ eapply map.split_alt. unfold mmap.split. simpl in E. simpl. rewrite E. reflexivity.
227
+ }
228
+ destruct Sp1 as (mSmall' & Sp1).
229
+ exists mSmall'. split. 2: eapply H2p1.
230
+ 2: { eapply map.split_alt. exact Sp1. }
231
+ rewrite Sp0 in H2.
232
+ unfold mmap.split in Sp1. rewrite Sp1. rewrite mmap.du_assoc in H2.
233
+ rewrite (mmap.du_comm (Some mAdd)) in H2. rewrite <- mmap.du_assoc in H2.
234
+ exact H2.
119
235
}
120
236
Qed .
121
237
0 commit comments