-
Notifications
You must be signed in to change notification settings - Fork 0
/
proph_erasure.v
848 lines (793 loc) · 33.5 KB
/
proph_erasure.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
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
376
377
378
379
380
381
382
383
384
385
386
387
388
389
390
391
392
393
394
395
396
397
398
399
400
401
402
403
404
405
406
407
408
409
410
411
412
413
414
415
416
417
418
419
420
421
422
423
424
425
426
427
428
429
430
431
432
433
434
435
436
437
438
439
440
441
442
443
444
445
446
447
448
449
450
451
452
453
454
455
456
457
458
459
460
461
462
463
464
465
466
467
468
469
470
471
472
473
474
475
476
477
478
479
480
481
482
483
484
485
486
487
488
489
490
491
492
493
494
495
496
497
498
499
500
501
502
503
504
505
506
507
508
509
510
511
512
513
514
515
516
517
518
519
520
521
522
523
524
525
526
527
528
529
530
531
532
533
534
535
536
537
538
539
540
541
542
543
544
545
546
547
548
549
550
551
552
553
554
555
556
557
558
559
560
561
562
563
564
565
566
567
568
569
570
571
572
573
574
575
576
577
578
579
580
581
582
583
584
585
586
587
588
589
590
591
592
593
594
595
596
597
598
599
600
601
602
603
604
605
606
607
608
609
610
611
612
613
614
615
616
617
618
619
620
621
622
623
624
625
626
627
628
629
630
631
632
633
634
635
636
637
638
639
640
641
642
643
644
645
646
647
648
649
650
651
652
653
654
655
656
657
658
659
660
661
662
663
664
665
666
667
668
669
670
671
672
673
674
675
676
677
678
679
680
681
682
683
684
685
686
687
688
689
690
691
692
693
694
695
696
697
698
699
700
701
702
703
704
705
706
707
708
709
710
711
712
713
714
715
716
717
718
719
720
721
722
723
724
725
726
727
728
729
730
731
732
733
734
735
736
737
738
739
740
741
742
743
744
745
746
747
748
749
750
751
752
753
754
755
756
757
758
759
760
761
762
763
764
765
766
767
768
769
770
771
772
773
774
775
776
777
778
779
780
781
782
783
784
785
786
787
788
789
790
791
792
793
794
795
796
797
798
799
800
801
802
803
804
805
806
807
808
809
810
811
812
813
814
815
816
817
818
819
820
821
822
823
824
825
826
827
828
829
830
831
832
833
834
835
836
837
838
839
840
841
842
843
844
845
846
847
848
From iris.program_logic Require Export adequacy.
From smr.lang Require Export lang notation tactics.
From iris.prelude Require Import options.
(** This file contains the proof that prophecies can be safely erased
from programs. We erase a program by replacing prophecy identifiers
with the unit values and respectively adapt the [NewProph] and [Resolve]
expressions. We prove that if a program [e] is safe with respect to a (pure)
postcondition [φ], then program [erase e] is also safe with respect to [φ]. *)
Implicit Types e : expr.
Implicit Types v w : val.
Implicit Types l : loc.
Implicit Types n m : Z.
Implicit Types i : nat.
Definition erase_base_lit (l : base_lit) : base_lit :=
match l with
| LitProphecy p => LitPoison
| _ => l
end.
Definition erase_resolve (e0 e1 e2 : expr) : expr := Fst (Fst (e0, e1, e2)).
Definition erased_new_proph : expr := (λ: <>, #LitPoison)%V #().
Fixpoint erase_expr (e : expr) : expr :=
match e with
| Val v => Val (erase_val v)
| Var x => Var x
| Rec f x e => Rec f x (erase_expr e)
| App e1 e2 => App (erase_expr e1) (erase_expr e2)
| UnOp op e => UnOp op (erase_expr e)
| BinOp op e1 e2 => BinOp op (erase_expr e1) (erase_expr e2)
| If e0 e1 e2 => If (erase_expr e0) (erase_expr e1) (erase_expr e2)
| Pair e1 e2 => Pair (erase_expr e1) (erase_expr e2)
| Fst e => Fst (erase_expr e)
| Snd e => Snd (erase_expr e)
| InjL e => InjL (erase_expr e)
| InjR e => InjR (erase_expr e)
| Case e0 e1 e2 => Case (erase_expr e0) (erase_expr e1) (erase_expr e2)
| Fork e => Fork (erase_expr e)
| AllocN e1 e2 => AllocN (erase_expr e1) (erase_expr e2)
| Free e1 e2 => Free (erase_expr e1) (erase_expr e2)
| Load e => Load (erase_expr e)
| Xchg e1 e2 => Xchg (erase_expr e1) (erase_expr e2)
| Store e1 e2 => Store (erase_expr e1) (erase_expr e2)
| CmpXchg e0 e1 e2 => CmpXchg (erase_expr e0) (erase_expr e1) (erase_expr e2)
| FAA e1 e2 => FAA (erase_expr e1) (erase_expr e2)
| NewProph => erased_new_proph
| Resolve e0 e1 e2 =>
erase_resolve (erase_expr e0) (erase_expr e1) (erase_expr e2)
end
with
erase_val (v : val) : val :=
match v with
| LitV l => LitV (erase_base_lit l)
| RecV f x e => RecV f x (erase_expr e)
| PairV v1 v2 => PairV (erase_val v1) (erase_val v2)
| InjLV v => InjLV (erase_val v)
| InjRV v => InjRV (erase_val v)
end.
Lemma erase_expr_subst x v e :
erase_expr (subst x v e) = subst x (erase_val v) (erase_expr e)
with
erase_val_subst x v (w : val) :
erase_expr (subst x v w) = subst x (erase_val v) (erase_val w).
Proof.
- destruct e; simpl; try case_decide;
rewrite ?erase_expr_subst ?erase_val_subst; auto.
- by destruct v.
Qed.
Lemma erase_expr_subst' x v e :
erase_expr (subst' x v e) = subst' x (erase_val v) (erase_expr e).
Proof. destruct x; eauto using erase_expr_subst. Qed.
Lemma erase_val_subst' x v (w : val) :
erase_expr (subst x v w) = subst x (erase_val v) (erase_val w).
Proof. destruct x; eauto using erase_val_subst. Qed.
Fixpoint erase_ectx_item (Ki : ectx_item) : list ectx_item :=
match Ki with
| AppLCtx v2 => [AppLCtx (erase_val v2)]
| AppRCtx e1 => [AppRCtx (erase_expr e1)]
| UnOpCtx op => [UnOpCtx op]
| BinOpLCtx op v2 => [BinOpLCtx op (erase_val v2)]
| BinOpRCtx op e1 => [BinOpRCtx op (erase_expr e1)]
| IfCtx e1 e2 => [IfCtx (erase_expr e1) (erase_expr e2)]
| PairLCtx v2 => [PairLCtx (erase_val v2)]
| PairRCtx e1 => [PairRCtx (erase_expr e1)]
| FstCtx => [FstCtx]
| SndCtx => [SndCtx]
| InjLCtx => [InjLCtx]
| InjRCtx => [InjRCtx]
| CaseCtx e1 e2 => [CaseCtx (erase_expr e1) (erase_expr e2)]
| AllocNLCtx v2 => [AllocNLCtx (erase_val v2)]
| AllocNRCtx e1 => [AllocNRCtx (erase_expr e1)]
| FreeLCtx v2 => [FreeLCtx (erase_val v2)]
| FreeRCtx e1 => [FreeRCtx (erase_expr e1)]
| LoadCtx => [LoadCtx]
| XchgLCtx v2 => [XchgLCtx (erase_val v2)]
| XchgRCtx e1 => [XchgRCtx (erase_expr e1)]
| StoreLCtx v2 => [StoreLCtx (erase_val v2)]
| StoreRCtx e1 => [StoreRCtx (erase_expr e1)]
| CmpXchgLCtx v1 v2 => [CmpXchgLCtx (erase_val v1) (erase_val v2)]
| CmpXchgMCtx e0 v2 => [CmpXchgMCtx (erase_expr e0) (erase_val v2)]
| CmpXchgRCtx e0 e1 => [CmpXchgRCtx (erase_expr e0) (erase_expr e1)]
| FaaLCtx v2 => [FaaLCtx (erase_val v2)]
| FaaRCtx e1 => [FaaRCtx (erase_expr e1)]
| ResolveLCtx ctx v1 v2 =>
erase_ectx_item ctx ++
[PairLCtx (erase_val v1); PairLCtx (erase_val v2); FstCtx; FstCtx]
| ResolveMCtx e0 v2 =>
[PairRCtx (erase_expr e0); PairLCtx (erase_val v2); FstCtx; FstCtx]
| ResolveRCtx e0 e1 =>
[PairRCtx (erase_expr e0, erase_expr e1); FstCtx; FstCtx]
end.
Definition erase_ectx (K : ectx heap_ectx_lang) : ectx heap_ectx_lang :=
mbind erase_ectx_item K.
Definition erase_tp (tp : list expr) : list expr := erase_expr <$> tp.
Definition erase_heap (h : memory) : memory := erase_val <$> h.
Definition erase_state (σ : state) : state :=
{| heap := erase_heap (heap σ); used_proph_id := ∅ |}.
Definition erase_cfg (ρ : cfg heap_lang) : cfg heap_lang :=
(erase_tp ρ.1, erase_state ρ.2).
Lemma erase_to_val e v :
to_val (erase_expr e) = Some v → ∃ v', to_val e = Some v' ∧ erase_val v' = v.
Proof. destruct e; naive_solver. Qed.
Lemma erase_not_val e : to_val e = None → to_val (erase_expr e) = None.
Proof. by destruct e. Qed.
Lemma erase_ectx_app K K' : erase_ectx (K ++ K') = erase_ectx K ++ erase_ectx K'.
Proof. by rewrite /erase_ectx bind_app. Qed.
Lemma erase_ectx_expr K e :
erase_expr (fill K e) = fill (erase_ectx K) (erase_expr e).
Proof.
revert e.
induction K as [|Ki K IHK] using rev_ind; simplify_eq/=; first done.
intros e.
rewrite !erase_ectx_app !fill_app /= -IHK {IHK}.
induction Ki; rewrite /= ?fill_app /= /erase_resolve; eauto with f_equal.
Qed.
Lemma val_is_unboxed_erased v :
val_is_unboxed (erase_val v) ↔ val_is_unboxed v.
Proof.
destruct v; rewrite /= /lit_is_unboxed; repeat (done || simpl; case_match).
Qed.
Lemma vals_compare_safe_erase v1 v2 :
vals_compare_safe (erase_val v1) (erase_val v2) ↔
vals_compare_safe v1 v2.
Proof. rewrite /vals_compare_safe !val_is_unboxed_erased. done. Qed.
Lemma erase_val_inj_iff v1 v2 :
vals_compare_safe v1 v2 → erase_val v1 = erase_val v2 ↔ v1 = v2.
Proof.
destruct v1, v2; rewrite /= /lit_is_unboxed;
repeat (done || (by intros [[] | []]) || simpl; case_match).
Qed.
(** if [un_op_eval] succeeds on erased value,
the so should it on the original value. *)
Lemma un_op_eval_erase op v v' :
un_op_eval op (erase_val v) = Some v' ↔
∃ w, un_op_eval op v = Some w ∧ erase_val w = v'.
Proof.
destruct op; simpl; repeat case_match; naive_solver.
Qed.
(** if [bin_op_eval] succeeds on erased value,
then so should it on the original value. *)
Lemma bin_op_eval_erase op v1 v2 v' :
bin_op_eval op (erase_val v1) (erase_val v2) = Some v' ↔
∃ w, bin_op_eval op v1 v2 = Some w ∧ erase_val w = v'.
Proof.
rewrite /bin_op_eval /bin_op_eval_int /bin_op_eval_bool /bin_op_eval_loc;
split; [intros ?|intros (?&?&?)];
repeat (case_match; simplify_eq/=); eauto.
- eexists _; split; eauto; simpl.
erewrite bool_decide_ext; first by eauto.
rewrite erase_val_inj_iff; done.
- by assert (vals_compare_safe v1 v2) by by apply vals_compare_safe_erase.
- by erewrite bool_decide_ext; last apply erase_val_inj_iff.
- by assert (vals_compare_safe (erase_val v1) (erase_val v2))
by by apply vals_compare_safe_erase.
Qed.
Lemma lookup_erase_heap h l : erase_heap h !! l = erase_val <$> h !! l.
Proof. by rewrite lookup_fmap. Qed.
Lemma erase_heap_insert h l v :
erase_heap (<[l := v]> h) = <[l := erase_val v]> (erase_heap h).
Proof.
by rewrite /erase_heap fmap_insert.
Qed.
Lemma erase_heap_delete h l :
erase_heap (delete l h) = delete l (erase_heap h).
Proof.
by rewrite /erase_heap fmap_delete.
Qed.
Lemma erase_heap_init_mem h l n v :
erase_heap (init_mem l (Z.to_nat n) v h) = init_mem l (Z.to_nat n) (erase_val v) (erase_heap h).
Proof.
move: (Z.to_nat n) l => {}n.
elim: n => [//|n IH l] /=; rewrite erase_heap_insert IH //.
Qed.
Lemma erase_heap_free_mem h l n :
erase_heap (free_mem l (Z.to_nat n) h) = free_mem l (Z.to_nat n) (erase_heap h).
Proof.
move: (Z.to_nat n) l => {}n.
elim: n => [//|n IH l] /=; rewrite erase_heap_delete IH //.
Qed.
Definition head_steps_to_erasure_of (e1 : expr) (σ1 : state) (e2 : expr)
(σ2 : state) (efs : list expr) :=
∃ κ' e2' σ2' efs',
head_step e1 σ1 κ' e2' σ2' efs' ∧
erase_expr e2' = e2 ∧ erase_state σ2' = σ2 ∧ erase_tp efs' = efs.
Lemma erased_head_step_head_step_rec f x e v σ :
head_steps_to_erasure_of ((rec: f x := e)%V v) σ
(subst' x (erase_val v)
(subst' f (rec: f x := erase_expr e) (erase_expr e)))
(erase_state σ) [].
Proof. by repeat econstructor; rewrite !erase_expr_subst'. Qed.
Lemma erased_head_step_head_step_NewProph σ :
head_steps_to_erasure_of NewProph σ #LitPoison (erase_state σ) [].
Proof. eexists _, _, _, _; split; first eapply new_proph_id_fresh; done. Qed.
Lemma erased_head_step_head_step_AllocN n v σ (l:blk) :
(0 < n)%Z →
(∀ m, erase_heap σ.(heap) !! (l,m) = None) →
head_steps_to_erasure_of
(AllocN #n v) σ #l (state_upd_heap (init_mem l (Z.to_nat n) (erase_val v)) (erase_state σ)) [].
Proof.
setoid_rewrite lookup_erase_heap. setoid_rewrite fmap_None.
eexists _, _, _, _; simpl; split; first econstructor; repeat split; try done.
rewrite /state_upd_heap /erase_state /= erase_heap_init_mem //.
Qed.
Lemma erased_head_step_head_step_Free n l σ :
(0 < n)%Z →
(∀ m, is_Some (erase_heap σ.(heap) !! (l +ₗ m)) ↔ 0 ≤ m < n)%Z →
head_steps_to_erasure_of (Free #n #l) σ #()
{| heap := free_mem l (Z.to_nat n) (erase_heap (heap σ)); used_proph_id := ∅ |} [].
Proof.
setoid_rewrite lookup_erase_heap. setoid_rewrite fmap_is_Some.
intros ? ?.
eexists _, _, _, _; simpl; split; first econstructor; [done|done|]; repeat split; eauto.
rewrite /state_upd_heap /erase_state /= erase_heap_free_mem //.
Qed.
Lemma erased_head_step_head_step_Load l σ v :
erase_heap (heap σ) !! l = Some v →
head_steps_to_erasure_of (! #l) σ v (erase_state σ) [].
Proof.
intros Hl.
rewrite lookup_erase_heap in Hl.
destruct (heap σ !! l) as [|] eqn:?; simplify_eq/=.
eexists _, _, _, _; simpl; split; first econstructor; eauto.
Qed.
Lemma erased_head_step_head_step_Xchg l v w σ :
erase_heap (heap σ) !! l = Some v →
head_steps_to_erasure_of (Xchg #l w) σ v
{| heap := <[l:=erase_val w]> (erase_heap (heap σ)); used_proph_id := ∅ |} [].
Proof.
intros Hl.
rewrite lookup_erase_heap in Hl.
destruct (heap σ !! l) as [|] eqn:?; simplify_eq/=.
eexists _, _, _, _; simpl; split; first econstructor; repeat split; eauto.
rewrite /state_upd_heap /erase_state /= erase_heap_insert //.
Qed.
Lemma erased_head_step_head_step_Store l v w σ :
erase_heap (heap σ) !! l = Some v →
head_steps_to_erasure_of (#l <- w) σ #()
{| heap := <[l:=erase_val w]> (erase_heap (heap σ)); used_proph_id := ∅ |} [].
Proof.
intros Hl.
rewrite lookup_erase_heap in Hl.
destruct (heap σ !! l) as [|] eqn:?; simplify_eq/=.
eexists _, _, _, _; simpl; split; first econstructor; repeat split; eauto.
rewrite /state_upd_heap /erase_state /= erase_heap_insert //.
Qed.
Lemma erased_head_step_head_step_CmpXchg l v w σ vl :
erase_heap (heap σ) !! l = Some vl →
vals_compare_safe vl (erase_val v) →
head_steps_to_erasure_of
(CmpXchg #l v w) σ
(vl, #(bool_decide (vl = erase_val v)))%V
(if bool_decide (vl = erase_val v)
then {| heap := <[l:=erase_val w]> (erase_heap (heap σ));
used_proph_id := ∅ |}
else erase_state σ) [].
Proof.
intros Hl Hvl.
rewrite lookup_erase_heap in Hl.
destruct (heap σ !! l) as [u|] eqn:?; simplify_eq/=.
rewrite -> vals_compare_safe_erase in Hvl.
destruct (decide (u = v)) as [->|Hneq].
- eexists _, _, _, _; simpl; split.
{ econstructor; eauto. }
rewrite !bool_decide_eq_true_2; eauto using erase_val_inj_iff; [].
rewrite -?erase_heap_insert.
split_and!; auto.
- eexists _, _, _, _; simpl; split.
{ econstructor; eauto. }
rewrite !bool_decide_eq_false_2; eauto; [].
by rewrite erase_val_inj_iff.
Qed.
Lemma erased_head_step_head_step_FAA l n m σ :
erase_heap (heap σ) !! l = Some #n →
head_steps_to_erasure_of
(FAA #l #m) σ #n
{| heap := <[l:= #(n + m)]> (erase_heap (heap σ));
used_proph_id := ∅ |} [].
Proof.
intros Hl.
rewrite lookup_erase_heap in Hl.
destruct (heap σ !! l) as [[[]| | | |]|] eqn:?; simplify_eq/=.
repeat econstructor; first by eauto.
by rewrite /state_upd_heap /erase_state /= erase_heap_insert.
Qed.
(** When the erased program makes a head step, so does the original program. *)
Lemma erased_head_step_head_step e1 σ1 κ e2 σ2 efs:
head_step (erase_expr e1) (erase_state σ1) κ e2 σ2 efs →
head_steps_to_erasure_of e1 σ1 e2 σ2 efs.
Proof.
intros Hhstep.
inversion Hhstep; simplify_eq/=;
repeat match goal with
| H : _ = erase_expr ?e |- _ => destruct e; simplify_eq/=
| H : _ = erase_val ?v |- _ => destruct v; simplify_eq/=
| H : _ = erase_base_lit ?l |- _ => destruct l; simplify_eq/=
| H : context [erased_new_proph] |- _ => unfold erased_new_proph in H
| H : un_op_eval _ (erase_val _) = Some _ |- _ =>
apply un_op_eval_erase in H as [? [? ?]]
| H : bin_op_eval _ (erase_val _) (erase_val _) = Some _ |- _ =>
apply bin_op_eval_erase in H as [? [? ?]]
| H : val_is_unboxed (erase_val _) |- _ =>
apply -> val_is_unboxed_erased in H
end; simplify_eq/=;
try (by repeat econstructor);
eauto using erased_head_step_head_step_rec,
erased_head_step_head_step_NewProph,
erased_head_step_head_step_Load,
erased_head_step_head_step_Xchg,
erased_head_step_head_step_Store,
erased_head_step_head_step_CmpXchg,
erased_head_step_head_step_FAA.
- by eapply erased_head_step_head_step_AllocN.
- by eapply erased_head_step_head_step_Free.
Qed.
Lemma fill_to_resolve e v1 v2 K e' :
to_val e' = None →
Resolve e v1 v2 = fill K e' →
K = [] ∨ ∃ K' Ki, K = K' ++ [ResolveLCtx Ki v1 v2].
Proof.
intros Hnv Hrs; simpl in *.
assert (∀ v K, fill K e' ≠ v) as Hcontr.
{ intros w K' Hw.
assert (to_val (of_val w) = to_val (fill K' e')) as He'
by (by rewrite Hw).
rewrite fill_not_val in He'; by eauto. }
destruct K as [|Ki K _] using rev_ind; first by left.
rewrite fill_app in Hrs.
destruct Ki; simplify_eq/=; eauto;
try exfalso; eapply Hcontr; eauto.
Qed.
Lemma projs_pure_steps (v0 v1 v2 : val) :
rtc pure_step (Fst (Fst (v0, v1, v2))) v0.
Proof.
etrans; first apply (rtc_pure_step_ctx (fill [PairLCtx _; FstCtx; FstCtx])).
{ apply rtc_once.
apply pure_head_step_pure_step.
split; first repeat econstructor.
intros ????? Hhstep; inversion Hhstep; simplify_eq/=; eauto. }
simpl.
etrans; first apply (rtc_pure_step_ctx (fill [FstCtx; FstCtx])).
{ apply rtc_once.
apply pure_head_step_pure_step.
split; first repeat econstructor.
intros ????? Hhstep; inversion Hhstep; simplify_eq/=; eauto. }
simpl.
etrans; first apply (rtc_pure_step_ctx (fill [FstCtx])).
{ apply rtc_once.
apply pure_head_step_pure_step.
split; first repeat econstructor.
intros ????? Hhstep; inversion Hhstep; simplify_eq/=; eauto. }
simpl.
apply rtc_once.
apply pure_head_step_pure_step.
split; first repeat econstructor.
intros ????? Hhstep; inversion Hhstep; simplify_eq/=; eauto.
Qed.
Lemma Resolve_3_vals_head_stuck v0 v1 v2 σ κ e σ' efs :
¬ head_step (Resolve v0 v1 v2) σ κ e σ' efs.
Proof.
intros Hhstep.
inversion Hhstep; simplify_eq/=.
apply (eq_None_not_Some (to_val (Val v0))); last eauto.
by eapply val_head_stuck.
Qed.
Lemma Resolve_3_vals_unsafe (v0 v1 v2 : val) σ :
¬ not_stuck (Resolve v0 v1 v2) σ.
Proof.
assert(∀ w K e, Val w = fill K e → is_Some (to_val e)) as Hvfill.
{ intros ? K ? Heq; eapply (fill_val K); rewrite /= -Heq; eauto. }
apply not_not_stuck.
split; first done.
intros ???? [K e1' e2' Hrs Hhstep]; simplify_eq/=.
destruct K as [|Ki K _] using rev_ind.
{ simplify_eq/=.
eapply Resolve_3_vals_head_stuck; eauto. }
rewrite fill_app in Hrs.
destruct Ki; simplify_eq/=.
- rename select ectx_item into Ki.
pose proof (fill_item_val Ki (fill K e1')) as Hnv.
apply fill_val in Hnv as [? Hnv]; last by rewrite -Hrs; eauto.
by erewrite val_head_stuck in Hnv.
- edestruct Hvfill as [? Heq]; eauto.
by erewrite val_head_stuck in Heq.
- edestruct Hvfill as [? Heq]; eauto.
by erewrite val_head_stuck in Heq.
Qed.
(** [(e1, σ1)] takes a [prim_step] to [(e2', σ2')] forking threads [efs']
such that [σ2] is the erasure of [σ2'] and [efs] is the erasure of
[efs']. Furthermore, [e2] takes [pure_steps] to match up with [e2]. It is
crucial for us that [e2] takes [pure_step]s because we need to know
that [e2] does not get stuck and that the steps are deterministic.
Essentially, the main part of the erasure proof's argument is that if
the erased program takes steps, then the original program also takes
matching steps. This however, does not entirely hold. In cases where
the erasure of [Resovle] takes a step, the original program
immediately produces the value while the erased program has to still
perform projections [Fst] to get the result (see the [Resolve] case
of [erase_expr]). For this purpose, we prove that in those cases (and
also in general) the erased program also takes a number of (possibly
zero) steps so that the original and the erased programs are matched
up again. *)
Definition prim_step_matched_by_erased_steps e1 σ1 e2 σ2 efs :=
∃ e2' σ2' κ' efs' e2'',
prim_step e1 σ1 κ' e2' σ2' efs' ∧ rtc pure_step e2 e2'' ∧
erase_expr e2' = e2'' ∧ erase_state σ2' = σ2 ∧ erase_tp efs' = efs.
Lemma prim_step_matched_by_erased_steps_ectx K e1 σ1 e2 σ2 efs :
prim_step_matched_by_erased_steps e1 σ1 e2 σ2 efs →
prim_step_matched_by_erased_steps (fill K e1) σ1 (fill (erase_ectx K) e2) σ2 efs.
Proof.
intros (?&?&?&?&?&?&?&?&?&?); simplify_eq/=.
eexists _, _, _, _, _; repeat split.
- by apply fill_prim_step.
- rewrite erase_ectx_expr.
by eapply (rtc_pure_step_ctx (fill (erase_ectx K))).
Qed.
Definition is_Resolve (e : expr) :=
match e with Resolve _ _ _ => True | _ => False end.
Global Instance is_Resolve_dec e : Decision (is_Resolve e).
Proof. destruct e; solve_decision. Qed.
Lemma non_resolve_prim_step_matched_by_erased_steps_ectx_item
Ki e1 e1' σ1 e2 σ2 efs :
to_val e1' = None →
¬ is_Resolve e1 →
not_stuck e1 σ1 →
erase_expr e1 = fill_item Ki e1' →
(∀ e1, erase_expr e1 = e1' → not_stuck e1 σ1 →
prim_step_matched_by_erased_steps e1 σ1 e2 σ2 efs) →
prim_step_matched_by_erased_steps e1 σ1 (fill_item Ki e2) σ2 efs.
Proof.
intros Hnv Hnr Hsf He1 IH.
destruct Ki; simplify_eq/=;
repeat
match goal with
| H : erase_expr ?e = _ |- _ => destruct e; simplify_eq/=; try done
| H : context [erased_new_proph] |- _ =>
rewrite /erased_new_proph in H; simplify_eq/=
| |- prim_step_matched_by_erased_steps ?e _ _ _ _ =>
let tac K e :=
lazymatch K with
| [] => fail
| _ => apply (prim_step_matched_by_erased_steps_ectx K);
apply IH; [done| by eapply (not_stuck_fill_inv (fill K))]
end
in
reshape_expr e tac
end.
Qed.
Lemma prim_step_matched_by_erased_steps_ectx_item Ki K e1 e1' σ1 e2 σ2 efs κ :
head_step e1' (erase_state σ1) κ e2 σ2 efs →
not_stuck e1 σ1 →
erase_expr e1 = fill_item Ki (fill K e1') →
(∀ K' e1, length K' ≤ length K →
erase_expr e1 = (fill K' e1') → not_stuck e1 σ1 →
prim_step_matched_by_erased_steps e1 σ1 (fill K' e2) σ2 efs) →
prim_step_matched_by_erased_steps e1 σ1 (fill_item Ki (fill K e2)) σ2 efs.
Proof.
intros Hhstp Hsf He1 IH; simpl in *.
(** Case split on whether e1 is a [Resolve] expression. *)
destruct (decide (is_Resolve e1)); last first.
{ (** e1 is not a [Resolve] expression. *)
eapply non_resolve_prim_step_matched_by_erased_steps_ectx_item; [|by eauto..].
by eapply fill_not_val, val_head_stuck. }
(** e1 is a [Resolve] expression. *)
destruct Ki; simplify_eq/=;
repeat
match goal with
| H : erase_expr ?e = ?e' |- _ =>
progress
match e' with
| fill _ _ => idtac
| _ => destruct e; simplify_eq/=
end
end; try done.
destruct K as [|Ki K _] using rev_ind; simplify_eq/=; [|].
{ (* case where (Fst (erase_expr e1_1, erase_expr e1_2, erase_expr e1_3)) *)
(* takes a head_step; it is impossible! *)
by inversion Hhstp; simplify_eq. }
rewrite /erase_resolve fill_app /= in He1; simplify_eq/=.
destruct Ki; simplify_eq/=; rewrite fill_app /=.
destruct K as [|Ki K _] using rev_ind; simplify_eq/=; [|].
{ (* case where (erase_expr e1_1, erase_expr e1_2, erase_expr e1_3) *)
(* takes a head_step; it is impossible! *)
inversion Hhstp. }
rewrite fill_app /= in He1.
destruct Ki; simplify_eq/=; rewrite fill_app /=.
- destruct K as [|Ki K _] using rev_ind; simplify_eq/=; [|].
{ (** [Resolve v0 v1 v2] is not safe! *)
inversion Hhstp; simplify_eq/=.
repeat
match goal with
| H : erase_expr ?e = _ |- _ => destruct e; simplify_eq/=
| H : _ = erase_expr ?e |- _ => destruct e; simplify_eq/=
end.
by exfalso; eapply Resolve_3_vals_unsafe. }
rewrite fill_app /= in He1.
destruct Ki; simplify_eq/=; rewrite fill_app /=.
+ (** e1 is of the form ([Resolve] e10 e11 v0) and e11 takes a prim_step. *)
destruct Hsf as [[? ?]| (?&?&?&?&Hrpstp)]; first done; simpl in *.
inversion Hrpstp as [??? Hrs ? Hhstp']; simplify_eq/=.
repeat
match goal with
| H : erase_expr ?e = ?e' |- _ =>
progress
match e' with
| fill _ _ => idtac
| _ => destruct e; simplify_eq/=
end
end.
edestruct fill_to_resolve as [?|[K' [Ki HK]]]; eauto;
[by eapply val_head_stuck| |]; simplify_eq/=.
* (** e1 is of the form ([Resolve] e10 e11 v0) and e11 takes a head_step. *)
inversion Hhstp'; simplify_eq.
edestruct (IH K) as (?&?&?&?&?&Hpstp&?&?&?&?);
[rewrite !app_length /=; lia|done|by eapply head_step_not_stuck|];
simplify_eq/=.
apply head_reducible_prim_step in Hpstp; simpl in *;
last by rewrite /head_reducible /=; eauto 10.
epose (λ H, head_step_to_val _ _ _ (Val _) _ _ _ _ _ _ _ H Hpstp)
as Hhstv; edestruct Hhstv as [? ?%of_to_val]; [done|eauto|];
simplify_eq.
eexists _, _, _, _, _; repeat split;
first (by apply head_prim_step; econstructor; eauto); auto.
etrans.
{ by apply (rtc_pure_step_ctx
(fill [PairLCtx _; PairLCtx _; FstCtx; FstCtx])). }
apply projs_pure_steps.
* (** e1 is of the form ([Resolve] e10 v v0) and e10 takes a
(non-head) prim_step. *)
rewrite fill_app in Hrs; simplify_eq/=.
edestruct (IH K) as (?&?&?&?&?&Hpstp&Hprstps&?&?&?);
[rewrite !app_length; lia|done| |].
{ change (fill_item Ki) with (fill [Ki]).
by rewrite -fill_app; eapply prim_step_not_stuck, Ectx_step. }
simplify_eq/=.
change (fill_item Ki) with (fill [Ki]) in Hpstp.
rewrite -fill_app in Hpstp.
eapply head_reducible_prim_step_ctx in Hpstp as [e2'' [He2'' Hpstp]];
last by eexists _; eauto.
simplify_eq/=.
eexists _, _, _, _, _; repeat split.
-- apply (fill_prim_step [ResolveLCtx _ _ _]); eapply Ectx_step; eauto.
-- simpl; rewrite fill_app in Hprstps.
by apply (rtc_pure_step_ctx
(fill [PairLCtx _; PairLCtx _; FstCtx; FstCtx])).
+ (** e1 is of the form ([Resolve] e1_ e1_2 v) and e1_2 takes a prim_step. *)
repeat
match goal with
| H : erase_expr ?e = ?e' |- _ =>
progress
match e' with
| fill _ _ => idtac
| _ => destruct e; simplify_eq/=
end
end.
apply (prim_step_matched_by_erased_steps_ectx [ResolveMCtx _ _]).
apply IH; [rewrite !app_length /=; lia|done|
by eapply (not_stuck_fill_inv (fill [ResolveMCtx _ _])); simpl].
- (** e1 is of the form ([Resolve] e1_ e1_2 e13) and e1_3 takes a prim_step. *)
apply (prim_step_matched_by_erased_steps_ectx [ResolveRCtx _ _]).
apply IH; [rewrite !app_length /=; lia|done|
by eapply (not_stuck_fill_inv (fill [ResolveRCtx _ _])); simpl].
Qed.
Lemma erased_prim_step_prim_step e1 σ1 κ e2 σ2 efs:
prim_step (erase_expr e1) (erase_state σ1) κ e2 σ2 efs → not_stuck e1 σ1 →
prim_step_matched_by_erased_steps e1 σ1 e2 σ2 efs.
Proof.
intros Hstp He1sf.
inversion Hstp as [K e1' e2' He1 ? Hhstp]; clear Hstp; simplify_eq/=.
set (len := length K); assert (length K = len) as Hlen by done; clearbody len.
revert K Hlen e1 He1 He1sf.
induction len as [m IHm]using lt_wf_ind; intros K Hlen e1 He1 He1sf;
simplify_eq.
destruct K as [|Ki K _] using rev_ind; simplify_eq/=.
{ apply erased_head_step_head_step in Hhstp as (?&?&?&?&?&<-&?&<-).
eexists _, _, _, _, _; repeat split;
first (by apply head_prim_step); auto using rtc_refl. }
rewrite app_length in IHm; simpl in *.
rewrite fill_app /=; rewrite fill_app /= in He1.
eapply prim_step_matched_by_erased_steps_ectx_item; eauto; [].
{ intros K' **; simpl in *. apply (IHm (length K')); auto with lia. }
Qed.
Lemma head_step_erased_prim_step_CmpXchg v1 v2 σ l vl:
heap σ !! l = Some vl →
vals_compare_safe vl v1 →
∃ e2' σ2' ef', prim_step (CmpXchg #l (erase_val v1)
(erase_val v2)) (erase_state σ) [] e2' σ2' ef'.
Proof.
intros Hl Hv.
destruct (bool_decide (vl = v1)) eqn:Heqvls.
- do 3 eexists; apply head_prim_step;
econstructor; [|by apply vals_compare_safe_erase|by eauto].
by rewrite /erase_state /state_upd_heap /= lookup_erase_heap Hl.
- do 3 eexists; apply head_prim_step;
econstructor; [|by apply vals_compare_safe_erase|by eauto].
by rewrite /erase_state /state_upd_heap /= lookup_erase_heap Hl.
Qed.
Lemma head_step_erased_prim_step_resolve e w σ :
(∃ e2' σ2' ef', prim_step (erase_expr e) (erase_state σ) [] e2' σ2' ef') →
∃ e2' σ2' ef',
prim_step (erase_resolve (erase_expr e) #LitPoison (erase_val w))
(erase_state σ) [] e2' σ2' ef'.
Proof.
intros (?&?&?&?).
by eexists _, _, _;
apply (fill_prim_step [PairLCtx _; PairLCtx _;FstCtx; FstCtx]).
Qed.
Lemma head_step_erased_prim_step_un_op σ op v v':
un_op_eval op v = Some v' →
∃ e2' σ2' ef', prim_step (UnOp op (erase_val v)) (erase_state σ) [] e2' σ2' ef'.
Proof.
do 3 eexists; apply head_prim_step; econstructor.
apply un_op_eval_erase; eauto.
Qed.
Lemma head_step_erased_prim_step_bin_op σ op v1 v2 v':
bin_op_eval op v1 v2 = Some v' →
∃ e2' σ2' ef', prim_step (BinOp op (erase_val v1) (erase_val v2))
(erase_state σ) [] e2' σ2' ef'.
Proof.
do 3 eexists; apply head_prim_step; econstructor.
apply bin_op_eval_erase; eauto.
Qed.
Lemma head_step_erased_prim_step_allocN σ (l:blk) n v:
(0 < n)%Z →
(∀ m, σ.(heap) !! (l,m) = None) →
∃ e2' σ2' ef',
prim_step (AllocN #n (erase_val v)) (erase_state σ) [] e2' σ2' ef'.
Proof.
do 3 eexists; apply head_prim_step; econstructor; eauto.
intros; rewrite lookup_erase_heap fmap_None; eauto.
Qed.
Lemma head_step_erased_prim_step_free σ l n :
(0 < n)%Z →
(∀ m, is_Some (σ.(heap) !! (l +ₗ m)) ↔ 0 ≤ m < n)%Z →
∃ e2' σ2' ef', prim_step (Free #n #l) (erase_state σ) [] e2' σ2' ef'.
Proof.
intros Hw. do 3 eexists; apply head_prim_step; econstructor; auto.
intros. rewrite /erase_state /state_upd_heap /= /erase_state /erase_heap lookup_fmap fmap_is_Some //.
Qed.
Lemma head_step_erased_prim_step_load σ l v:
heap σ !! l = Some v →
∃ e2' σ2' ef', prim_step (! #l) (erase_state σ) [] e2' σ2' ef'.
Proof.
do 3 eexists; apply head_prim_step; econstructor.
rewrite /erase_state /state_upd_heap /= lookup_erase_heap.
by destruct lookup; simplify_eq.
Qed.
Lemma head_step_erased_prim_step_xchg σ l v w :
heap σ !! l = Some v →
∃ e2' σ2' ef', prim_step (Xchg #l (erase_val w)) (erase_state σ) [] e2' σ2' ef'.
Proof.
intros Hl. do 3 eexists; apply head_prim_step; econstructor.
rewrite /erase_state /state_upd_heap /= lookup_erase_heap Hl; eauto.
Qed.
Lemma head_step_erased_prim_step_store σ l v w :
heap σ !! l = Some v →
∃ e2' σ2' ef', prim_step (#l <- erase_val w) (erase_state σ) [] e2' σ2' ef'.
Proof.
intros Hw. do 3 eexists; apply head_prim_step; econstructor.
rewrite /erase_state /state_upd_heap /= lookup_erase_heap Hw; eauto.
Qed.
Lemma head_step_erased_prim_step_FAA σ l n n':
heap σ !! l = Some #n →
∃ e2' σ2' ef', prim_step (FAA #l #n') (erase_state σ) [] e2' σ2' ef'.
Proof.
intros Hl.
do 3 eexists; apply head_prim_step. econstructor.
by rewrite /erase_state /state_upd_heap /= lookup_erase_heap Hl.
Qed.
(**
[Resolve] is translated as a projection out of a triple.
Therefore, when resolve takes a head step, the erasure of [Resolve] takes a
prim step inside the triple.
*)
Lemma head_step_erased_prim_step e1 σ1 κ e2 σ2 ef:
head_step e1 σ1 κ e2 σ2 ef →
∃ e2' σ2' ef', prim_step (erase_expr e1) (erase_state σ1) [] e2' σ2' ef'.
Proof.
induction 1; simplify_eq/=;
eauto using head_step_erased_prim_step_CmpXchg,
head_step_erased_prim_step_resolve,
head_step_erased_prim_step_un_op,
head_step_erased_prim_step_bin_op,
head_step_erased_prim_step_allocN,
head_step_erased_prim_step_free,
head_step_erased_prim_step_load,
head_step_erased_prim_step_store,
head_step_erased_prim_step_xchg,
head_step_erased_prim_step_FAA;
by do 3 eexists; apply head_prim_step; econstructor.
Qed.
Lemma reducible_erased_reducible e σ :
reducible e σ → reducible (erase_expr e) (erase_state σ).
Proof.
intros (?&?&?&?&Hpstp); simpl in *.
inversion Hpstp; simplify_eq/=.
rewrite erase_ectx_expr.
edestruct head_step_erased_prim_step as (?&?&?&?); first done; simpl in *.
eexists _, _, _, _; eapply fill_prim_step; eauto.
Qed.
Lemma pure_step_tp_safe t1 t2 e1 σ :
(∀ e2, e2 ∈ t2 → not_stuck e2 σ) → pure_steps_tp t1 (erase_tp t2) →
e1 ∈ t1 → not_stuck e1 (erase_state σ).
Proof.
intros Ht2 Hpr [i He1]%elem_of_list_lookup_1.
eapply Forall2_lookup_l in Hpr as [e2' [He2' Hpr]]; simpl in *; eauto.
rewrite /erase_tp list_lookup_fmap in He2'.
destruct (t2 !! i) eqn:He2; simplify_eq/=.
apply elem_of_list_lookup_2, Ht2 in He2.
clear -Hpr He2.
inversion Hpr as [|??? [? _]]; simplify_eq.
- destruct He2 as [[? ?%of_to_val]|]; simplify_eq/=; first by left; eauto.
by right; apply reducible_erased_reducible.
- right; eauto using reducible_no_obs_reducible.
Qed.
(** This is the top-level erasure theorem: erasure preserves adequacy. *)
Theorem erasure e σ φ :
adequate NotStuck e σ φ →
adequate NotStuck (erase_expr e) (erase_state σ)
(λ v σ, ∃ v' σ', erase_val v' = v ∧ erase_state σ' = σ ∧ φ v' σ').
Proof.
simpl; intros Hade; simpl in *.
cut (∀ t2 σ2,
rtc erased_step ([erase_expr e], erase_state σ) (t2, σ2) →
(∃ t2' t2'' σ2',
rtc erased_step ([e], σ) (t2'', σ2') ∧
t2' = erase_tp t2'' ∧ σ2 = erase_state σ2' ∧
pure_steps_tp t2 t2')).
{ intros Hreach; split; simpl in *.
- intros ? ? ? Hrtc; edestruct (Hreach _ _ Hrtc) as
(t2'&t2''&σ2'&Hos&Ht2'&Hσ2&Hptp); simplify_eq/=.
apply Forall2_cons_inv_l in Hptp as (oe&t3&Hoe%rtc_pure_step_val&_&?);
destruct t2''; simplify_eq/=.
apply erase_to_val in Hoe as (?&?%of_to_val&?); simplify_eq.
pose proof (adequate_result _ _ _ _ Hade _ _ _ Hos); eauto.
- intros ? ? ? Hs Hrtc He2; edestruct (Hreach _ _ Hrtc) as
(t2'&t2''&σ2'&Hos&Ht2'&Hσ2&Hptp); simplify_eq/=.
eapply pure_step_tp_safe; [|done..].
intros e2' He2'.
apply (adequate_not_stuck _ _ _ _ Hade _ _ _ eq_refl Hos He2'). }
intros t2 σ2 [n Hstps]%rtc_nsteps; simpl in *; revert t2 σ2 Hstps.
induction n as [|n IHn].
{ intros t2 σ2 Hstps; inversion Hstps; simplify_eq /=.
repeat econstructor. }
intros t2 σ2 Hstps.
apply nsteps_inv_r in Hstps as [[t3 σ3] [Hstps Hρ]]; simpl in *.
destruct (IHn _ _ Hstps) as (t2'&t2''&σ2'&Hostps&?&?&Hprstps); simplify_eq.
edestruct @erased_step_pure_step_tp as [[? Hint]|Hext]; simplify_eq/=;
eauto 10; [|done..].
destruct Hext as (i&ei&t2'&efs&e'&κ&Hi1&Ht2&Hpstp);
simplify_eq/=.
rewrite /erase_tp list_lookup_fmap in Hi1.
destruct (t2'' !! i) as [eio|] eqn:Heq; simplify_eq/=.
edestruct erased_prim_step_prim_step as
(eio' & σ3 & κ' & efs' & ee & Heiopstp & Hprstps' & ?&?&?); first done;
last simplify_eq/=.
{ eapply adequate_not_stuck; eauto using elem_of_list_lookup_2. }
eexists _, _, _; repeat split.
{ etrans; first done.
apply rtc_once; eexists.
eapply step_insert; eauto. }
rewrite /erase_tp fmap_app.
rewrite list_fmap_insert/=.
apply Forall2_app; last done.
apply Forall2_same_length_lookup; split.
{ apply Forall2_length in Hprstps; rewrite fmap_length in Hprstps.
by rewrite !insert_length fmap_length. }
intros j x y.
destruct (decide (i = j)); simplify_eq.
{ rewrite !list_lookup_insert ?fmap_length; eauto using lookup_lt_Some; [].
by intros ? ?; simplify_eq. }
rewrite !list_lookup_insert_ne // list_lookup_fmap.
intros ? ?.
eapply Forall2_lookup_lr; eauto.
by rewrite /erase_tp list_lookup_fmap.
Qed.