-
Notifications
You must be signed in to change notification settings - Fork 1
/
Copy pathSoundForgetMemory.v
1998 lines (1939 loc) · 76.2 KB
/
SoundForgetMemory.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
849
850
851
852
853
854
855
856
857
858
859
860
861
862
863
864
865
866
867
868
869
870
871
872
873
874
875
876
877
878
879
880
881
882
883
884
885
886
887
888
889
890
891
892
893
894
895
896
897
898
899
900
901
902
903
904
905
906
907
908
909
910
911
912
913
914
915
916
917
918
919
920
921
922
923
924
925
926
927
928
929
930
931
932
933
934
935
936
937
938
939
940
941
942
943
944
945
946
947
948
949
950
951
952
953
954
955
956
957
958
959
960
961
962
963
964
965
966
967
968
969
970
971
972
973
974
975
976
977
978
979
980
981
982
983
984
985
986
987
988
989
990
991
992
993
994
995
996
997
998
999
1000
Require Import syntax.
Require Import alist.
Require Import FMapWeakList.
Require Import Classical.
Require Import Coqlib.
Require Import infrastructure.
Require Import Metatheory.
Import LLVMsyntax.
Import LLVMinfra.
Require Import opsem.
Require Import memory_props.
Require Import sflib.
Require Import paco.
Import Opsem.
Require Import TODO.
Require Import Exprs.
Require Import Hints.
Require Import Postcond.
Require Import Validator.
Require Import GenericValues.
Require AssnMem.
Require AssnState.
Require Import Inject.
Require Import SoundBase.
Require Import TODOProof.
Import Memory.
Require Import MemAux.
Set Implicit Arguments.
(* TODO: move *)
Lemma some_injective A (a b:A):
Some a = Some b -> a = b.
Proof.
congruence.
Qed.
Inductive mem_change_inject (conf conf_tgt:Config) assnmem: mem_change -> mem_change -> Prop :=
| mem_change_inject_alloca_alloca
gsz gn0 gn1 a
ty dv
(N_INJECT: genericvalues_inject.gv_inject assnmem.(AssnMem.Rel.inject) gn0 gn1)
: mem_change_inject conf conf_tgt assnmem
(mem_change_alloca dv ty gsz gn0 a)
(mem_change_alloca dv ty gsz gn1 a)
| mem_change_inject_alloca_none
gsz gn a ty dv
: mem_change_inject conf conf_tgt assnmem
(mem_change_alloca dv ty gsz gn a)
mem_change_none
| mem_change_inject_none_alloca
gsz gn a ty dv
: mem_change_inject conf conf_tgt
assnmem mem_change_none
(mem_change_alloca dv ty gsz gn a)
| mem_change_inject_store_store
ptr0 ptr1 gv0 gv1 ty a
(PTR_INJECT: genericvalues_inject.gv_inject assnmem.(AssnMem.Rel.inject) ptr0 ptr1)
(VAL_INJECT: genericvalues_inject.gv_inject assnmem.(AssnMem.Rel.inject) gv0 gv1)
: mem_change_inject conf conf_tgt assnmem
(mem_change_store ptr0 ty gv0 a)
(mem_change_store ptr1 ty gv1 a)
| mem_change_inject_store_nop
ptr gv ty a
(DISJOINT: forall b (GV2BLOCKS: In b (GV2blocks ptr)),
<<NOT_PUBLIC: ~ AssnMem.Rel.public_src assnmem.(AssnMem.Rel.inject) b>> /\
<<PARENT_DISJOINT: ~ In b assnmem.(AssnMem.Rel.src).(AssnMem.Unary.private_parent)>>)
: mem_change_inject conf conf_tgt assnmem
(mem_change_store ptr ty gv a)
mem_change_none
| mem_change_inject_free
ptr0 ptr1
(PTR_INJECT: genericvalues_inject.gv_inject assnmem.(AssnMem.Rel.inject) ptr0 ptr1)
: mem_change_inject conf conf_tgt assnmem
(mem_change_free ptr0)
(mem_change_free ptr1)
| mem_change_inject_none
: mem_change_inject conf conf_tgt assnmem
mem_change_none
mem_change_none
.
Inductive states_mem_change conf mem0 mem1: mem_change -> Prop :=
| states_mem_change_alloca
ty bsz gn a dv mb
(ALLOCA: alloca conf.(CurTargetData) mem0 bsz gn a = Some (mem1, mb))
: states_mem_change conf mem0 mem1 (mem_change_alloca dv ty bsz gn a)
| states_mem_change_store
ptr ty gv a
(VALID_PTRS: MemProps.valid_ptrs mem0.(Mem.nextblock) gv)
(STORE: mstore conf.(CurTargetData) mem0 ptr ty gv a = Some mem1)
: states_mem_change conf mem0 mem1 (mem_change_store ptr ty gv a)
| states_mem_change_free
ptr
(FREE: free conf.(CurTargetData) mem0 ptr = Some mem1)
: states_mem_change conf mem0 mem1 (mem_change_free ptr)
| states_mem_change_none
(MEM_EQ: mem0 = mem1)
: states_mem_change conf mem0 mem1 mem_change_none
.
(* Relation between mem_change and cmd *)
Lemma gv_inject_ptr_public_src
assnmem ptr0 ptr1 b ofs conf_src
(PTR_INJECT : genericvalues_inject.gv_inject (AssnMem.Rel.inject assnmem) ptr0 ptr1)
(PTR : GV2ptr (CurTargetData conf_src) (getPointerSize (CurTargetData conf_src)) ptr0 = Some (Values.Vptr b ofs))
: AssnMem.Rel.public_src (AssnMem.Rel.inject assnmem) b.
Proof.
exploit genericvalues_inject.simulation__GV2ptr; try exact PTR; eauto. i. des.
ii. inv x1. clarify.
Qed.
Lemma simulation__GV2ptr_tgt
: forall (mi : Values.meminj) (TD : TargetData) (gv1 gv1' : GenericValue) (v' : Values.val),
genericvalues_inject.gv_inject mi gv1 gv1' ->
GV2ptr TD (getPointerSize TD) gv1' = Some v' ->
exists v : Values.val, GV2ptr TD (getPointerSize TD) gv1 = Some v /\ memory_sim.MoreMem.val_inject mi v v'.
Proof.
Abort.
(* Subset *)
Lemma forget_memory_unary_Subset
def_mem leaks_mem inv0
: Assertion.Subset_unary (ForgetMemory.unary def_mem leaks_mem inv0) inv0.
Proof.
unfold ForgetMemory.unary.
destruct leaks_mem; destruct def_mem.
- econs; ss; ii; des_ifs; try econs; ss.
+ eapply ExprPairSetFacts.filter_iff in H; try by solve_compat_bool. des. eauto.
+ eapply AtomSetFacts.remove_iff in H; try by solve_compat_bool. des. eauto.
- econs; ss; ii; des_ifs; try econs; ss.
eapply AtomSetFacts.remove_iff in H; try by solve_compat_bool. des. eauto.
- econs; ss; ii; des_ifs; try econs; ss.
eapply ExprPairSetFacts.filter_iff in H; try by solve_compat_bool. des. eauto.
- econs; ss; ii; des_ifs; try econs; ss.
Qed.
Lemma forget_memory_Subset
def_mem_src def_mem_tgt
leaks_mem_src leaks_mem_tgt
inv0
: Assertion.Subset (ForgetMemory.t def_mem_src def_mem_tgt leaks_mem_src leaks_mem_tgt inv0) inv0.
Proof.
unfold ForgetMemory.t; des_ifs;
econs; ss; try reflexivity; apply forget_memory_unary_Subset.
Qed.
(* soundness proof *)
Lemma step_mem_change
st0 st1 invst0 assnmem0 inv0
cmd cmds
conf evt gmax public
(STATE: AssnState.Unary.sem conf st0 invst0 assnmem0 gmax public inv0)
(MEM: AssnMem.Unary.sem conf gmax public st0.(Mem) assnmem0)
(CMD: st0.(EC).(CurCmds) = cmd::cmds)
(NONCALL: Instruction.isCallInst cmd = false)
(NONMALLOC: isMallocInst cmd = false)
(STEP: sInsn conf st0 st1 evt)
: <<UNIQUE_PARENT_MEM:
forall mptr typ align val'
(LOAD: mload conf.(CurTargetData) st1.(Mem) mptr typ align = Some val'),
AssnMem.gv_diffblock_with_blocks conf val' assnmem0.(AssnMem.Unary.unique_parent)>> /\
exists mc,
<<MC_SOME: mem_change_of_cmd conf cmd st0.(EC).(Locals) = Some mc>> /\
<<STATE_EQUIV: states_mem_change conf st0.(Mem) st1.(Mem) mc>>.
Proof.
assert (MEM':=MEM).
inv MEM'.
inv STEP; destruct cmd; ss; clarify;
try by esplits; ss; econs; eauto.
- split.
+ ii. eapply UNIQUE_PARENT_MEM; eauto.
eapply MemProps.free_preserves_mload_inv; eauto.
+ esplits; ss.
* des_ifs.
* econs; eauto.
- split.
+ ii.
exploit MemProps.alloca_preserves_mload_inv; eauto. i. des.
{ eapply UNIQUE_PARENT_MEM; eauto. }
{ ss.
eapply AssnState.Unary.undef_diffblock; eauto.
eapply {|
CurSystem := S;
CurTargetData := TD;
CurProducts := Ps;
Globals := gl;
FunTable := fs |}.
}
+ esplits; ss.
* des_ifs.
* econs; eauto.
- split.
+ ii.
exploit (mstore_never_produce_new_ptr' {| CurSystem := S;
CurTargetData := TD;
CurProducts := Ps;
Globals := gl;
FunTable := fs |}); eauto.
{ i. hexploit UNIQUE_PARENT_MEM; eauto. }
hexploit getOperandValue_not_unique_parent.
{ eauto. }
{ eauto. }
{ inv STATE. ss.
destruct B. destruct s.
hexploit typings_props.wf_fdef__wf_cmd; try apply WF_FDEF.
{ apply WF_EC. }
{
instantiate (1:= insn_store id5 typ5 value1 value2 align5).
inv WF_EC. ss.
unfold OpsemAux.get_cmds_from_block in *. ss.
eapply sublist_In; eauto.
ss. left. eauto.
}
intro WF_INSN. inv WF_INSN. destruct TD. ss. clarify. eauto.
}
{ instantiate (1 := gv1). eauto. }
ss. clarify.
unfold AssnMem.gv_diffblock_with_blocks. eauto.
+ esplits; ss.
* des_ifs.
* econs; eauto.
inv STATE. ss.
{ destruct B. destruct s.
hexploit typings_props.wf_fdef__wf_cmd; try apply WF_FDEF.
{ apply WF_EC. }
{
instantiate (1:= insn_store id5 typ5 value1 value2 align5).
inv WF_EC. ss.
unfold OpsemAux.get_cmds_from_block in *. ss.
eapply sublist_In; eauto.
ss. left. eauto.
}
intro WF_INSN. inv WF_INSN. destruct TD. ss. clarify.
destruct value1; ss.
- eapply WF_LOCAL; eauto.
- inv H10.
exploit MemAux.wf_globals_const2GV; eauto.
i. inv MEM. ss.
inv WF0.
eapply MemProps.valid_ptrs__trans; eauto.
rewrite <- Pplus_one_succ_r.
apply Pos.le_succ_l. eauto.
}
Qed.
Ltac exploit_inject_value :=
repeat (match goal with
| [H1: Assertion.inject_value ?inv ?vt1 ?vt2 = true |- _] =>
exploit AssnState.Rel.inject_value_spec; try exact H1; eauto; clear H1
end;
(try by
match goal with
| [H: getOperandValue (CurTargetData ?conf) ?v (Locals (EC ?st)) (Globals ?conf) = Some ?gv1 |-
AssnState.Unary.sem_valueT ?conf ?st ?invst (ValueT.lift Tag.physical ?v) = Some ?gv2] =>
destruct v; [ss; unfold IdT.lift; solve_sem_idT; eauto | ss]
end); i; des).
Ltac inv_conf :=
match goal with
| [H: AssnState.valid_conf _ _ ?conf_src ?conf_tgt |- _] =>
let TD := fresh in
let GL := fresh in
destruct H as [[TD GL]]; rewrite TD in *; rewrite GL in *
end.
Lemma inject_mem_change
m_src conf_src cmd_src mc_src st0_src
m_tgt conf_tgt cmd_tgt mc_tgt st0_tgt
inv0 assnmem0 invst0
(CONF: AssnState.valid_conf m_src m_tgt conf_src conf_tgt)
(STATE : AssnState.Rel.sem conf_src conf_tgt st0_src st0_tgt invst0 assnmem0 inv0)
(MEM : AssnMem.Rel.sem conf_src conf_tgt st0_src.(Mem) st0_tgt.(Mem) assnmem0)
(INJECT_EVENT : postcond_cmd_inject_event cmd_src cmd_tgt inv0)
(MC_SRC : mem_change_of_cmd conf_src cmd_src st0_src.(EC).(Locals) = Some mc_src)
(MC_TGT : mem_change_of_cmd conf_tgt cmd_tgt st0_tgt.(EC).(Locals) = Some mc_tgt)
: mem_change_inject conf_src conf_tgt assnmem0 mc_src mc_tgt.
Proof.
destruct cmd_src; destruct cmd_tgt; ss; clarify;
(try by simtac; econs); (* cases including none *)
try by simtac;
unfold is_true in *;
repeat (des_bool; des);
inject_clarify;
exploit_inject_value;
inv_conf;
inject_clarify;
econs; eauto.
unfold Assertion.is_private in *. des_ifs.
destruct x as [t x]; unfold ValueT.lift in *. des_ifs.
inv STATE. inv SRC.
unfold is_true in *.
(* rewrite <- IdTSetFacts.mem_iff in *. *)
econs. ii.
exploit PRIVATE; eauto.
{ eapply IdTSet.mem_2; eauto. }
{ ss. }
ii; des.
inv PRIVATE_BLOCK.
splits; ss.
Qed.
Ltac solve_alloc_inject :=
by ii;
match goal with
| [ALLOCA: ?cmd = insn_alloca _ _ _ _,
MC_SOME: mem_change_of_cmd _ ?cmd _ = Some mem_change_none |- _] =>
rewrite ALLOCA in MC_SOME; ss; des_ifs
| [ALLOCA: ?cmd = insn_alloca _ _ _ _,
MC_SOME: mem_change_of_cmd _ ?cmd _ = Some (mem_change_store _ _ _ _) |- _] =>
rewrite ALLOCA in MC_SOME; ss; des_ifs
| [ALLOCA: ?cmd = insn_alloca _ _ _ _,
MC_SOME: mem_change_of_cmd _ ?cmd _ = Some (mem_change_free _) |- _] =>
rewrite ALLOCA in MC_SOME; ss; des_ifs
end.
Lemma inject_assnmem
m_src conf_src st0_src st1_src cmd_src cmds_src evt_src
m_tgt conf_tgt st0_tgt st1_tgt cmd_tgt cmds_tgt evt_tgt
invst0 assnmem0 inv0
(CONF: AssnState.valid_conf m_src m_tgt conf_src conf_tgt)
(STATE : AssnState.Rel.sem conf_src conf_tgt st0_src st0_tgt invst0 assnmem0 inv0)
(MEM : AssnMem.Rel.sem conf_src conf_tgt (Mem st0_src) (Mem st0_tgt) assnmem0)
(CMD_SRC : CurCmds (EC st0_src) = cmd_src :: cmds_src)
(CMD_TGT : CurCmds (EC st0_tgt) = cmd_tgt :: cmds_tgt)
(NONCALL_SRC: Instruction.isCallInst cmd_src = false)
(NONCALL_TGT: Instruction.isCallInst cmd_tgt = false)
(STEP_SRC : sInsn conf_src st0_src st1_src evt_src)
(STEP_TGT : sInsn conf_tgt st0_tgt st1_tgt evt_tgt)
(INJECT_EVENT : postcond_cmd_inject_event cmd_src cmd_tgt inv0)
: exists assnmem1,
<<ALLOC_INJECT: alloc_inject conf_src conf_tgt st0_src st0_tgt
st1_src st1_tgt cmd_src cmd_tgt assnmem1>> /\
<<ALLOC_PRIVATE: alloc_private conf_src conf_tgt cmd_src cmd_tgt st0_src st0_tgt st1_src st1_tgt assnmem1>> /\
<<MEM: AssnMem.Rel.sem conf_src conf_tgt (Mem st1_src) (Mem st1_tgt) assnmem1>> /\
<<MEMLE: AssnMem.Rel.le assnmem0 assnmem1>> /\
<<PRIVATE_PRESERVED_SRC: IdTSet.For_all
(AssnState.Unary.sem_private
conf_src st0_src (AssnState.Rel.src invst0)
(AssnMem.Unary.private_parent (AssnMem.Rel.src assnmem1))
(AssnMem.Rel.public_src (AssnMem.Rel.inject assnmem1)))
(Assertion.private (Assertion.src inv0))>> /\
<<PRIVATE_PRESERVED_TGT: IdTSet.For_all
(AssnState.Unary.sem_private
conf_tgt st0_tgt (AssnState.Rel.tgt invst0)
(AssnMem.Unary.private_parent (AssnMem.Rel.tgt assnmem1))
(AssnMem.Rel.public_tgt (AssnMem.Rel.inject assnmem1)))
(Assertion.private (Assertion.tgt inv0))>> /\
(* INJECT_ALLOCAS is needed for "inv_state_sem_monotone_wrt_assnmem" *)
(* I am not sure this design is good; as INJECT_ALLOCAS belongs to AssnState *)
(* INJECT_ALLOCAS is needed here. && also in SimLocal *)
(* If we put this inside AssnMem (may need to strengthen "frozen" until "nextblock" *)
(* This will not be the case *)
<<INJECT_ALLOCAS:
AssnState.Rel.inject_allocas (AssnMem.Rel.inject assnmem1)
st0_src.(EC).(Allocas) st0_tgt.(EC).(Allocas)>> /\
<<INJECT_ALLOCAS:
AssnState.Rel.inject_allocas (AssnMem.Rel.inject assnmem1)
st1_src.(EC).(Allocas) st1_tgt.(EC).(Allocas)>> /\
<<VALID_ALLOCAS_SRC: Forall (Mem.valid_block (Mem st1_src)) st1_src.(EC).(Allocas)>> /\
<<VALID_ALLOCAS_TGT: Forall (Mem.valid_block (Mem st1_tgt)) st1_tgt.(EC).(Allocas)>>
.
Proof.
exploit postcond_cmd_inject_event_non_malloc; eauto; []; ii; des.
hexploit step_mem_change; try (inv STATE; exact SRC); eauto.
{ inv MEM. exact SRC. }
intro MCS.
destruct MCS as [UNIQUE_PRIVATE_SRC [mc_src [MC_SOME_SRC STATE_EQUIV_SRC]]]. des.
hexploit step_mem_change; try (inv STATE; exact TGT); eauto.
{ inv MEM. exact TGT. }
intro MCS.
destruct MCS as [UNIQUE_PRIVATE_TGT [mc_tgt [MC_SOME_TGT STATE_EQUIV_TGT]]]. des.
exploit inject_mem_change; eauto. intro MC_INJECT.
inv MC_INJECT.
- (* alloc - alloc *)
inv STEP_SRC; inv CMD_SRC; ss; des_ifs.
rename Mem0 into mem0_src. rename Mem' into mem1_src. rename mb into mb_src.
match goal with
| [H: alloca _ _ _ _ _ = _ |- _] => rename H into ALLOCA_SRC
end.
inv STEP_TGT; inv CMD_TGT; ss; try by des; congruence.
rename Mem0 into mem0_tgt. rename Mem' into mem1_tgt. rename mb into mb_tgt.
match goal with
| [H: alloca _ _ _ _ _ = _ |- _] => rename H into ALLOCA_TGT
end.
clear_tac.
dup ALLOCA_SRC.
dup ALLOCA_TGT.
unfold alloca, option_map, flip in ALLOCA_SRC, ALLOCA_TGT. des_ifs_safe.
expl alloca_result (try exact ALLOCA_SRC0; eauto). clarify.
expl alloca_result (try exact ALLOCA_TGT0; eauto). clarify.
expl Mem.alloc_result (try exact Heq1; eauto). clarify.
expl Mem.alloc_result (try exact Heq3; eauto). clarify.
clear_tac.
eexists.
instantiate (1:= AssnMem.Rel.mk _ _ _
(fun b =>
if Values.eq_block b (Mem.nextblock mem0_src)
then Some ((Mem.nextblock mem0_tgt), 0%Z)
else assnmem0.(AssnMem.Rel.inject) b)).
esplits.
+ (* alloc_inject *)
ii. ss.
inv ALLOCA_SRC. inv ALLOCA_TGT.
esplits.
* unfold alloc_inject_unary.
esplits; try apply lookupAL_updateAddAL_eq; ss.
* unfold alloc_inject_unary.
esplits; try apply lookupAL_updateAddAL_eq; ss.
* destruct (Values.eq_block (Mem.nextblock mem0_src)(Mem.nextblock mem0_src)); ss.
+ (* alloc_private *)
econs; ii; ss; des_ifs.
+ (* AssnMem sem *)
inv MEM; ss.
instantiate (3:= AssnMem.Unary.mk _
assnmem0.(AssnMem.Rel.src).(AssnMem.Unary.mem_parent)
assnmem0.(AssnMem.Rel.src).(AssnMem.Unary.unique_parent)
mem1_src.(Mem.nextblock)).
instantiate (2:= AssnMem.Unary.mk _
assnmem0.(AssnMem.Rel.tgt).(AssnMem.Unary.mem_parent)
assnmem0.(AssnMem.Rel.tgt).(AssnMem.Unary.unique_parent)
mem1_tgt.(Mem.nextblock)).
econs; ss; eauto.
{ (* SRC *)
inv SRC.
econs; eauto.
- eapply MemProps.alloca_preserves_wf_Mem; eauto.
- ss. i. exploit PRIVATE_PARENT; eauto. intros [NOT_PUBLIC_B NEXT_B].
split.
+ ii. unfold AssnMem.Rel.public_src in *.
destruct (Values.eq_block _ _); ss.
psimpl.
+ erewrite Mem.nextblock_drop with (m:= m0); try eassumption.
erewrite Mem.nextblock_alloc; try eassumption.
eapply Pos.lt_le_trans; eauto.
eapply Ple_succ; eauto.
- i. exploit MEM_PARENT; eauto. i. ss.
match goal with
| [H: mload_aux (AssnMem.Unary.mem_parent _) _ b _ = _ |- _] =>
rewrite H
end.
exploit PRIVATE_PARENT; eauto. i.
unfold AssnMem.private_block in *. des.
eapply alloca_preserves_mload_aux_other_eq; eauto.
- ss. rewrite NEXT_BLOCK. etransitivity; [|apply Ple_succ]; eauto.
}
{ (* TGT *)
inv TGT.
econs; eauto.
- eapply MemProps.alloca_preserves_wf_Mem; eauto.
- ss. i. exploit PRIVATE_PARENT; eauto.
intros [NOT_PUBLIC_B NEXT_B].
split.
+ ii.
match goal with
| [H: ~ AssnMem.Rel.public_tgt _ _ |- False] =>
apply H
end.
unfold AssnMem.Rel.public_tgt in *. des.
destruct (Values.eq_block _ _).
* clarify. exfalso. psimpl.
* esplits; eauto.
+ erewrite Mem.nextblock_drop with (m:= m); try eassumption.
erewrite Mem.nextblock_alloc; try eassumption.
eapply Pos.lt_le_trans; eauto.
eapply Ple_succ; eauto.
- i. exploit MEM_PARENT; eauto. i.
match goal with
| [H: mload_aux (AssnMem.Unary.mem_parent _) _ b _ = _ |- _] =>
rewrite H
end.
exploit PRIVATE_PARENT; eauto. i.
unfold AssnMem.private_block in *. des.
eapply alloca_preserves_mload_aux_other_eq; eauto.
- ss. rewrite NEXT_BLOCK0. etransitivity; [|apply Ple_succ]; eauto.
}
{ (* inject *)
inv INJECT.
unfold is_true in *.
repeat rewrite andb_true_iff in INJECT_EVENT.
destruct INJECT_EVENT as [[[ID_EQ TYP_EQ] INJECT_VALUE] DEC_EQ].
unfold proj_sumbool in *. des_sumbool. clarify.
econs.
{ (* mi_access *)
ii. exploit valid_access_alloca_inv; try exact ALLOCA_SRC0; eauto.
i.
destruct (Values.eq_block _ _).
- clarify.
assert(Memtype.perm_order Memtype.Writable p).
{ move Heq6 at bottom.
destruct p; try econs.
eapply Mem.valid_access_perm in H0.
des.
hexploit Mem.perm_drop_2; try apply H0; eauto.
split; ss.
expl Memdata.size_chunk_pos. instantiate (1:= chunk) in size_chunk_pos.
apply Z.gt_lt_iff in size_chunk_pos.
eapply Z.lt_le_trans.
{ instantiate (1:= ofs + Memdata.size_chunk chunk). omega. }
unfold get_or_else in *. des_ifs; omega.
}
eapply valid_access_alloca_same; eauto.
repeat rewrite Z.add_0_r.
des. splits; eauto.
exploit genericvalues_inject.simulation__GV2int; eauto. intro GV2INT_INJECT.
assert(TD = TD0).
{ inv CONF. inv INJECT. ss. }
subst.
rewrite <- GV2INT_INJECT in *. clarify.
- exploit mi_access; eauto.
eapply valid_access_alloca_other; eauto.
}
{ (* mi_memval *)
i. destruct (Values.eq_block _ _).
- clarify.
rewrite Z.add_0_r.
erewrite alloca_contents_same; eauto.
erewrite alloca_contents_same; eauto.
apply memory_sim.MoreMem.memval_inject_undef.
- eapply memory_sim.MoreMem.memval_inject_incr.
+ assert (DIFF_BLK_TGT: b2 <> (Mem.nextblock mem0_tgt)).
{ exploit genericvalues_inject.Hmap2; eauto. }
eapply alloca_contents_other in DIFF_BLK_TGT; eauto.
rewrite DIFF_BLK_TGT.
erewrite alloca_contents_other; eauto.
apply mi_memval; eauto.
eapply Mem.perm_drop_4 in H0; [|try exact Heq6].
exploit Mem.perm_alloc_inv.
{ try exact Heq5. }
{ eauto. }
i. des_ifs.
+ ii.
destruct (Values.eq_block _ _).
{ subst. exfalso.
exploit genericvalues_inject.Hmap1; eauto.
{ instantiate (1:=Mem.nextblock mem0_src).
exploit alloca_inv; try exact ALLOCA_SRC0. i. psimpl.
}
i. congruence.
}
eauto.
}
}
{ (* wf_sb_mi *)
inv WF.
econs.
- (* no_overlap *)
ii.
destruct (Values.eq_block _ _);
destruct (Values.eq_block _ _); clarify.
+ exploit Hmap2; eauto. i. psimpl.
+ exploit Hmap2; eauto. i. psimpl.
+ eapply Hno_overlap with (b1:=b1) (b2:=b2); eauto.
- (* Hmap1 *)
intro b_src. i. destruct (Values.eq_block _ _).
+ subst.
rewrite NEXT_BLOCK in *.
exfalso. psimpl.
+ apply Hmap1. psimpl.
- (* Hmap2 *)
intros b_src b_tgt. i. destruct (Values.eq_block _ _).
+ clarify.
subst. rewrite NEXT_BLOCK0 in *.
apply Plt_succ'.
+ exploit Hmap2; eauto. i. psimpl.
- (* mi_freeblocks *)
intros b NOT_VALID_BLOCK.
destruct (Values.eq_block _ _).
+ subst.
exfalso.
apply NOT_VALID_BLOCK.
unfold Mem.valid_block.
psimpl.
+ apply mi_freeblocks. intros VALID_BLOCK.
apply NOT_VALID_BLOCK.
unfold Mem.valid_block in *.
psimpl.
- (* mi_mappedblocks *)
i. destruct (Values.eq_block _ _).
+ clarify.
unfold Mem.valid_block in *.
psimpl.
+ eapply Mem.drop_perm_valid_block_1; eauto.
eapply Mem.valid_block_alloc.
{ eauto. }
eapply mi_mappedblocks; eauto.
- (* mi_range_blocks *)
ii. destruct (Values.eq_block _ _).
+ subst. clarify.
+ eapply mi_range_block; eauto.
- (* mi_bounds *)
ii. destruct (Values.eq_block _ _).
+ clarify.
erewrite Mem.bounds_drop; eauto.
erewrite Mem.bounds_alloc_same; cycle 1.
{ eauto. }
erewrite Mem.bounds_drop; eauto.
erewrite Mem.bounds_alloc_same; cycle 1.
{ eauto. }
apply injective_projections; ss.
solve_match_bool. clarify.
exploit genericvalues_inject.simulation__GV2int; eauto. intro GV2INT_INJECT.
assert(TD = TD0).
{ inv CONF. inv INJECT0. ss. }
subst.
rewrite GV2INT_INJECT in *. clarify.
+ erewrite Mem.bounds_drop; eauto.
erewrite Mem.bounds_alloc_other with (b':=b); eauto; cycle 1.
assert (NEQ_BLK_TGT: b' <> mem0_tgt.(Mem.nextblock)).
{ exploit Hmap2; eauto. }
symmetry. (* TODO: "rewrite at" doesn't work, WHY???????? *)
erewrite Mem.bounds_drop; eauto.
erewrite Mem.bounds_alloc_other with (b':=b'); try exact NEQ_BLK_TGT; cycle 1.
{ eauto. }
symmetry. eapply mi_bounds; eauto.
- (* mi_globals *)
i. destruct (Values.eq_block _ _).
+ subst.
exploit mi_globals; eauto. i.
exploit Hmap1.
{ psimpl. }
i. congruence.
+ exploit mi_globals; eauto.
}
{ (* ftable *)
eapply inject_incr__preserves__ftable_simulation; eauto.
ii. rename H into INJ0.
des_ifs_safe.
inv WF.
exploit Hmap1.
{ ii. rewrite Pos.compare_refl in *. clarify. }
intro INJ1.
rewrite INJ1 in *. clarify.
}
+ (* le *)
econs; try (econs; ss).
{ inv MEM. inv SRC. rewrite <- NEXTBLOCK. psimpl. }
{ inv MEM. inv TGT. rewrite <- NEXTBLOCK. psimpl. }
{
(* incr *)
ii. ss.
destruct (Values.eq_block _ _); eauto.
subst.
inv MEM. inv WF.
exploit Hmap1.
{ psimpl. }
i. congruence.
}
{
ii. des. des_ifsH NEW0.
unfold Mem.valid_block.
split; ss.
- apply MEM.
- apply MEM.
}
+ ss.
inv STATE. inv SRC. ss.
ii. exploit PRIVATE; eauto. i. des.
esplits; eauto. ss.
unfold AssnMem.private_block in *. des.
split.
* unfold AssnMem.Rel.public_src in *.
destruct (Values.eq_block _ _); ss.
psimpl.
* eauto.
+ ss.
inv STATE. inv TGT. ss.
ii. exploit PRIVATE; eauto. i. des.
esplits; eauto. ss.
unfold AssnMem.private_block in *. des.
split.
* unfold AssnMem.Rel.public_tgt in *.
ii. des.
destruct (Values.eq_block _ _); ss.
{ clarify. psimpl. }
apply PRIVATE_BLOCK. esplits; eauto.
* eauto.
+ ss.
inv STATE. clear MAYDIFF.
inv SRC. clear LESSDEF NOALIAS UNIQUE PRIVATE ALLOCAS_PARENT
WF_LOCAL WF_PREVIOUS WF_GHOST UNIQUE_PARENT_LOCAL WF_FDEF WF_EC.
inv TGT. clear LESSDEF NOALIAS UNIQUE PRIVATE ALLOCAS_PARENT
WF_LOCAL WF_PREVIOUS WF_GHOST UNIQUE_PARENT_LOCAL WF_FDEF WF_EC.
ss.
eapply inject_allocas_enhance; eauto.
{ i. des_ifsG. exfalso. eapply Plt_irrefl. eauto. }
{ i. des_ifsG. exfalso. eapply Plt_irrefl. eauto. }
(* ginduction ALLOCAS; ii; ss. *)
(* * econs; eauto. *)
(* * inv ALLOCAS_VALID. *)
(* econs; eauto. des_ifs. *)
(* exfalso. eapply Plt_irrefl. eauto. *)
(* * inv ALLOCAS_VALID0. *)
(* econs; eauto. *)
(* i. des_ifs. *)
(* { ii. clarify. eapply Plt_irrefl. eauto. } *)
(* eapply PRIVATE. *)
(* * inv ALLOCAS_VALID. inv ALLOCAS_VALID0. *)
(* econs 4; eauto. *)
(* des_ifs. *)
(* exfalso. eapply Plt_irrefl. eauto. *)
+ ss.
inv STATE. clear MAYDIFF.
inv SRC. clear LESSDEF NOALIAS UNIQUE PRIVATE ALLOCAS_PARENT
WF_LOCAL WF_PREVIOUS WF_GHOST UNIQUE_PARENT_LOCAL WF_FDEF WF_EC.
inv TGT. clear LESSDEF NOALIAS UNIQUE PRIVATE ALLOCAS_PARENT
WF_LOCAL WF_PREVIOUS WF_GHOST UNIQUE_PARENT_LOCAL WF_FDEF WF_EC.
ss.
econs 4; eauto.
* des_ifsG.
* eapply inject_allocas_enhance; eauto.
{ i. des_ifsG. exfalso. eapply Pos.lt_irrefl. eauto. }
{ i. des_ifsG. exfalso. eapply Pos.lt_irrefl. eauto. }
+ inv STATE. inv SRC.
clear - NEXT_BLOCK ALLOCAS_VALID.
ss.
econs; eauto.
* unfold Mem.valid_block. rewrite NEXT_BLOCK. eapply Plt_succ.
* eapply Forall_harder; eauto.
i.
unfold Mem.valid_block. rewrite NEXT_BLOCK.
eapply Pos.lt_le_trans; eauto. eapply Ple_succ.
+ inv STATE. inv TGT.
clear - NEXT_BLOCK0 ALLOCAS_VALID.
ss.
econs; eauto.
* unfold Mem.valid_block. rewrite NEXT_BLOCK0. eapply Plt_succ.
* eapply Forall_harder; eauto.
i.
unfold Mem.valid_block. rewrite NEXT_BLOCK0.
eapply Pos.lt_le_trans; eauto. eapply Ple_succ.
- (* alloc - none *)
inv STATE_EQUIV_TGT. rewrite <- MEM_EQ in *.
inv STEP_SRC; destruct cmd_src; ss; clarify;
des_matchH MC_SOME_SRC; clarify; ss.
rename Mem0 into mem0_src.
rename Mem' into mem1_src.
inv STATE_EQUIV_SRC. ss. clarify.
exploit alloca_result; eauto. intros [ALLOCA_BLOCK_SRC ALLOCA_NEXT_SRC]. des.
exists (AssnMem.Rel.mk
(AssnMem.Unary.mk
assnmem0.(AssnMem.Rel.src).(AssnMem.Unary.private_parent)
assnmem0.(AssnMem.Rel.src).(AssnMem.Unary.mem_parent)
assnmem0.(AssnMem.Rel.src).(AssnMem.Unary.unique_parent)
mem1_src.(Mem.nextblock))
assnmem0.(AssnMem.Rel.tgt)
assnmem0.(AssnMem.Rel.gmax)
(* mem0_src should be private. *)
(* we can just copy assnmem0's, because wf_sb_mi guarantees mem0_src is priviate *)
(* Without wf_sb_mi, we need to put a function with if-then-else *)
assnmem0.(AssnMem.Rel.inject)
).
esplits; ss; eauto.
+ (* alloc_inject *)
solve_alloc_inject.
+ (* alloc_private *)
econs; ii; ss; try by des_ifs.
clarify.
inv MEM. inv SRC. ss.
esplits.
* apply lookupAL_updateAddAL_eq.
* ii. ss. des; ss.
move b at bottom.
rename b into __b__.
clarify.
unfold AssnMem.private_block in *.
splits; ss.
{
ii.
unfold AssnMem.Rel.public_src in H.
apply H.
(* destruct assnmem0. ss. *)
inv WF.
apply Hmap1. psimpl.
}
{ psimpl. }
{ ii.
exploit PRIVATE_PARENT; eauto; []; ii; des.
psimpl.
}
+ inv MEM.
econs; eauto.
* ss. eapply assnmem_unary_alloca_sem; eauto.
ii. unfold AssnMem.Rel.public_src in *.
apply H.
inv WF.
apply Hmap1. psimpl.
* inv INJECT.
econs.
{ (* mi-access *)
i. exploit mi_access; eauto.
assert (DIFFBLOCK_ALLOC: b1 <> Mem.nextblock mem0_src).
{ inv WF.
ii. exploit Hmap1.
{ instantiate (1:= Mem.nextblock mem0_src).
psimpl. }
i. subst. ss. congruence.
}
exploit valid_access_alloca_inv; eauto.
des_ifs.
}
{ (* mi_memval *)
i.
assert (DIFFBLOCK_ALLOC: b1 <> Mem.nextblock mem0_src).
{ inv WF.
ii. exploit Hmap1.
{ instantiate (1:= Mem.nextblock mem0_src).
psimpl. }
i. subst. ss. congruence.
}
exploit mi_memval; eauto.
{ u_alloca.
eapply Mem.perm_drop_4 in H0; revgoals; eauto.
hexploit Mem.perm_alloc_inv; eauto; []; i.
clear INJECT_EVENT.
des_ifs. eauto.
}
i. exploit alloca_contents_other; eauto.
intro CONTENTS.
rewrite CONTENTS. eauto.
}
* inv WF.
econs; eauto.
++ i. apply Hmap1. psimpl.
++ i. apply Hmap1.
unfold Mem.valid_block in *. psimpl.
++ i.
assert (ALLOC_PRIVATE: b <> Mem.nextblock mem0_src).
{ ii. subst.
exploit Hmap1.
{ psimpl. }
i. ss. congruence. }
u_alloca.
erewrite Mem.bounds_drop; revgoals; eauto.
erewrite Mem.bounds_alloc_other; try exact ALLOC_PRIVATE; cycle 1.
{ eauto. }
eapply mi_bounds; eauto.
+ econs; eauto.
* econs; eauto. ss.
inv MEM. inv SRC.
rewrite <- NEXTBLOCK.
psimpl.
* econs; eauto. ss.
inv MEM. inv TGT.
rewrite <- NEXTBLOCK.
psimpl.
* clarify. ss.
econs; eauto.
ii. des; ss. clarify.
+ inv STATE. inv SRC. eauto.
+ inv STATE. inv TGT. eauto.
+ inv STATE. ss.
+ inv STATE. ss.
assert(EQ_ALLOC: Allocas (EC st1_tgt) = Allocas (EC st0_tgt)).
{ inv STEP_TGT; ss; des_ifs.
- inv MC_SOME_TGT. des_ifs. }
des. rewrite EQ_ALLOC.
econs; eauto.
inv MEM.
inv WF.
apply Hmap1. psimpl.
+ inv STATE. inv SRC. ss.
clear - ALLOCAS_VALID NEXT_BLOCK.
econs; eauto.
* unfold Mem.valid_block. rewrite NEXT_BLOCK. eapply Plt_succ.
* eapply Forall_harder; eauto.
i.
unfold Mem.valid_block. rewrite NEXT_BLOCK.
eapply Pos.lt_le_trans; eauto.
eapply Ple_succ.
+ assert(EQ_ALLOC: Allocas (EC st1_tgt) = Allocas (EC st0_tgt) /\ ECS st1_tgt = ECS st0_tgt).
{ inv STEP_TGT; ss; des_ifs.
- inv MC_SOME_TGT. des_ifs. } des.
inv STATE. inv TGT. ss.
rewrite EQ_ALLOC. ss.
- (* none - alloc *)
inv STATE_EQUIV_SRC.
rewrite <- MEM_EQ in *.
inv STEP_TGT; destruct cmd_tgt; ss; clarify;
des_matchH MC_SOME_TGT; clarify; ss.
rename Mem0 into mem0_tgt.
rename Mem' into mem1_tgt.
inv STATE_EQUIV_TGT. ss. clarify.
exploit alloca_result; eauto. intros [ALLOCA_BLOCK_TGT ALLOCA_NEXT_TGT]. des.
exists (AssnMem.Rel.mk
assnmem0.(AssnMem.Rel.src)
(AssnMem.Unary.mk
assnmem0.(AssnMem.Rel.tgt).(AssnMem.Unary.private_parent)
assnmem0.(AssnMem.Rel.tgt).(AssnMem.Unary.mem_parent)
assnmem0.(AssnMem.Rel.tgt).(AssnMem.Unary.unique_parent)
mem1_tgt.(Mem.nextblock))
assnmem0.(AssnMem.Rel.gmax)
assnmem0.(AssnMem.Rel.inject)).
esplits; ss; eauto.
+ (* alloc_inject *)
solve_alloc_inject.
+ (* alloc_private *)
econs; ii; ss; try by des_ifs.
clarify.
inv MEM. inv TGT. ss.
esplits; try apply lookupAL_updateAddAL_eq.
* ii. ss.
des; ss.
unfold AssnMem.private_block.
splits.
{
ii.
subst.
unfold AssnMem.Rel.public_tgt in H.
des.
inv WF.
exploit Hmap2; eauto; []; ii; des.
psimpl.
}
{ psimpl. }
{ ii.
subst.
exploit PRIVATE_PARENT; eauto; []; ii; des.
unfold AssnMem.private_block in *.
des.
psimpl.
}
+ inv MEM.
econs; eauto.
* eapply assnmem_unary_alloca_sem; eauto.
ss. ii. unfold AssnMem.Rel.public_tgt in *. des.
inv WF.
exploit Hmap2; eauto. i.
psimpl.
* inv INJECT.
econs.
{ (* mi-access *)
i. exploit mi_access; eauto. i.
assert (DIFFBLOCK_ALLOC: b2 <> Mem.nextblock mem0_tgt).
{ inv WF.
ii. exploit Hmap2; eauto.
i. psimpl. }
exploit valid_access_alloca_other; eauto.
}
{ (* mi_memval *)
i.
assert (DIFFBLOCK_ALLOC: b2 <> Mem.nextblock mem0_tgt).
{ inv WF.
ii. exploit Hmap2; eauto.
i. psimpl. }
exploit mi_memval; eauto.
i. exploit alloca_contents_other; eauto.
intro CONTENTS.
rewrite CONTENTS. eauto.
}
* inv WF.
econs; eauto.
++ i. exploit Hmap2; eauto. i. psimpl.
++ i. exploit Hmap2; eauto. i.
unfold Mem.valid_block. psimpl.
++ i.
assert (ALLOC_PRIVATE: b' <> Mem.nextblock mem0_tgt).
{ ii. subst.
exploit Hmap2; eauto. i. psimpl. }
u_alloca.
symmetry. (* TODO: WHY???????? "rewrite at" dosen't work ???????????????? *)
erewrite Mem.bounds_drop; revgoals; eauto.
erewrite Mem.bounds_alloc_other with (b':=b'); try exact ALLOC_PRIVATE; cycle 1.
{ eauto. }
symmetry.
eapply mi_bounds; eauto.
+ econs; eauto.
* econs; eauto. ss.
inv MEM. inv SRC. rewrite <- NEXTBLOCK. psimpl.
* econs; eauto. ss.
inv MEM. inv TGT. rewrite <- NEXTBLOCK. psimpl.
* ss. econs; eauto.
ii.
des; ss. clarify.
+ inv STATE. inv SRC. eauto.
+ inv STATE. inv TGT. eauto.
+ inv STATE. ss.
+ inv STATE. ss.
assert(EQ_ALLOC: Allocas (EC st1_src) = Allocas (EC st0_src)).
{ inv STEP_SRC; ss; des_ifs.
- inv MC_SOME_SRC. des_ifs. }
des. rewrite EQ_ALLOC.
econs; eauto.
inv MEM.
inv WF.