-
Notifications
You must be signed in to change notification settings - Fork 1
/
Copy pathSoundBase.v
1793 lines (1680 loc) · 57 KB
/
SoundBase.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 sflib.
Require Import paco.
Import Opsem.
Require Import TODO.
Require Import TODOProof.
Require Import Decs.
Require Import Exprs.
Require Import Validator.
Require Import GenericValues.
Require Import Inject.
Require AssnMem.
Require AssnState.
Require Import Hints.
Require Import memory_props.
Import Memory.
Require Import opsem_wf.
Require Import genericvalues_inject.
Require Import memory_sim.
Require Import MemAux.
Set Implicit Arguments.
Section SEM_VALID_PTR.
Lemma sem_id_preserves_valid_ptr
conf st invst0 assnmem0 gmax pubs inv0
(STATE: AssnState.Unary.sem conf st invst0 assnmem0 gmax pubs inv0)
(MEM : AssnMem.Unary.sem conf gmax pubs st.(Mem) assnmem0)
idt0 gv0
(SEM: AssnState.Unary.sem_idT st invst0 idt0 = Some gv0)
:
<<VALID: MemProps.valid_ptrs st.(Mem).(Mem.nextblock) gv0>>
.
Proof.
destruct idt0; ss.
destruct t; ss.
- inv STATE. eapply WF_LOCAL; eauto.
- inv STATE. eapply WF_PREVIOUS; eauto.
- inv STATE. eapply WF_GHOST; eauto.
Qed.
Lemma sem_value_preserves_valid_ptr
conf st invst0 assnmem0 gmax pubs inv0
(STATE: AssnState.Unary.sem conf st invst0 assnmem0 gmax pubs inv0)
(MEM : AssnMem.Unary.sem conf gmax pubs st.(Mem) assnmem0)
v0 gv0
(SEM: AssnState.Unary.sem_valueT conf st invst0 v0 = Some gv0)
:
<<VALID: MemProps.valid_ptrs st.(Mem).(Mem.nextblock) gv0>>
.
Proof.
destruct v0; ss.
- eapply sem_id_preserves_valid_ptr; eauto.
- exploit wf_globals_const2GV; try apply MEM; eauto.
i.
eapply MemProps.valid_ptrs__trans; eauto.
{ inv MEM.
red in WF. des.
rewrite <- Pplus_one_succ_r.
eapply Pos.le_succ_l; eauto.
}
Qed.
Lemma sem_list_value_preserves_valid_ptr
conf st invst0 assnmem0 gmax pubs inv0
(STATE: AssnState.Unary.sem conf st invst0 assnmem0 gmax pubs inv0)
(MEM : AssnMem.Unary.sem conf gmax pubs st.(Mem) assnmem0)
vs0 gvs0
(SEM: AssnState.Unary.sem_list_valueT conf st invst0 vs0 = Some gvs0)
:
<<VALID: List.Forall (MemProps.valid_ptrs st.(Mem).(Mem.nextblock)) gvs0>>
.
Proof.
ginduction vs0; ii; ss; clarify.
des_ifs.
econs; eauto.
{ eapply sem_value_preserves_valid_ptr; eauto. }
eapply IHvs0; eauto.
Qed.
Lemma gep_preserves_valid_ptrs
conf st invst0 assnmem0 gmax pubs inv0
(STATE: AssnState.Unary.sem conf st invst0 assnmem0 gmax pubs inv0)
(MEM : AssnMem.Unary.sem conf gmax pubs st.(Mem) assnmem0)
v g
(SEM0: AssnState.Unary.sem_valueT conf st invst0 v = Some g)
l0
lsv
(SEM1: AssnState.Unary.sem_list_valueT conf st invst0 lsv = Some l0)
t ib u gv0
(GEP: gep (CurTargetData conf) t l0 ib u g = Some gv0)
:
<<VALID: MemProps.valid_ptrs (Mem.nextblock (Mem st)) gv0>>
.
Proof.
unfold gep in *. unfold genericvalues.LLVMgv.GEP in GEP.
des_ifs; try (by eapply MemProps.undef_valid_ptrs; eauto).
unfold GV2ptr in *. des_ifs.
unfold mgep in *. des_ifs.
exploit sem_value_preserves_valid_ptr; ss; eauto.
Qed.
Lemma mselect_preserves_valid_ptrs
conf st invst0 assnmem0 gmax pubs inv0
(STATE: AssnState.Unary.sem conf st invst0 assnmem0 gmax pubs inv0)
(MEM : AssnMem.Unary.sem conf gmax pubs st.(Mem) assnmem0)
v w z g g0 g1
(SEM0: AssnState.Unary.sem_valueT conf st invst0 v = Some g)
(SEM1: AssnState.Unary.sem_valueT conf st invst0 w = Some g0)
(SEM2: AssnState.Unary.sem_valueT conf st invst0 z = Some g1)
t gv0
(SELET: mselect (CurTargetData conf) t g g0 g1 = Some gv0)
:
<<VALID: MemProps.valid_ptrs (Mem.nextblock (Mem st)) gv0>>
.
Proof.
unfold mselect in *.
unfold fit_chunk_gv in *.
des_ifs; try (by eapply MemProps.undef_valid_ptrs; eauto).
- eapply sem_value_preserves_valid_ptr; eauto.
- eapply sem_value_preserves_valid_ptr; eauto.
Qed.
Lemma sem_expr_preserves_valid_ptr
conf st invst0 assnmem0 gmax pubs inv0
(STATE: AssnState.Unary.sem conf st invst0 assnmem0 gmax pubs inv0)
(MEM : AssnMem.Unary.sem conf gmax pubs st.(Mem) assnmem0)
x0 gv0
(SEM: AssnState.Unary.sem_expr conf st invst0 x0 = Some gv0)
:
<<VALID: MemProps.valid_ptrs st.(Mem).(Mem.nextblock) gv0>>
.
Proof.
red.
destruct x0; ss; des_ifs.
- eapply MemProps.mbop_preserves_valid_ptrs; eauto.
- eapply MemProps.mfbop_preserves_valid_ptrs; eauto.
- eapply MemProps.extractGenericValue_preserves_valid_ptrs; eauto.
eapply sem_value_preserves_valid_ptr; eauto.
- eapply MemProps.insertGenericValue_preserves_valid_ptrs; eauto.
+ eapply sem_value_preserves_valid_ptr; eauto.
+ eapply sem_value_preserves_valid_ptr; eauto.
- eapply gep_preserves_valid_ptrs; eauto.
- eapply MemProps.mtrunc_preserves_valid_ptrs; eauto.
- eapply MemProps.mext_preserves_valid_ptrs; eauto.
- eapply MemProps.mcast_preserves_valid_ptrs; eauto.
eapply sem_value_preserves_valid_ptr; eauto.
- eapply MemProps.micmp_preserves_valid_ptrs; eauto.
- eapply MemProps.mfcmp_preserves_valid_ptrs; eauto.
- eapply mselect_preserves_valid_ptrs; try exact SEM; eauto.
- eapply sem_value_preserves_valid_ptr; eauto.
- inv MEM.
eapply WF. eauto.
Qed.
End SEM_VALID_PTR.
(* TODO: position *)
Definition get_blocks (f:fdef): blocks :=
let '(fdef_intro _ blocks) := f in
blocks.
(* TODO: position *)
Lemma lookupBlockViaLabelFromFdef_spec
fdef l:
lookupBlockViaLabelFromFdef fdef l =
lookupAL _ fdef.(get_blocks) l.
Proof. destruct fdef. ss. Qed.
(* TODO: position *)
Definition ite A (c:bool) (a b:A): A :=
if c then a else b.
(* TODO: position *)
Lemma ite_spec A c (a b:A):
ite c a b = if c then a else b.
Proof. ss. Qed.
(* TODO: position *)
Opaque ite.
(* TODO: position *)
Lemma lookupAL_ite
X (l:AssocList X) decision l1 l2 v1 v2
(V1: lookupAL _ l l1 = Some v1)
(V2: lookupAL _ l l2 = Some v2):
lookupAL _ l (ite decision l1 l2) = Some (ite decision v1 v2).
Proof. destruct decision; ss. Qed.
(* TODO: position *)
Definition return_locals
(TD:TargetData)
(retval:option GenericValue)
(id:id) (noret:noret) (ct:typ)
(lc:GVMap): option GVsMap :=
match retval, noret with
| Some retval, false =>
match (fit_gv TD ct) retval with
| Some retval' => Some (updateAddAL _ lc id retval')
| _ => None
end
| _, true => Some lc
| _, _ => None
end.
Lemma return_locals_fully_inject_locals
TD id noret ty inv
retval_src locals1_src locals2_src
retval_tgt locals1_tgt
conf_src conf_tgt mem_src mem_tgt
(RETVAL: lift2_option (genericvalues_inject.gv_inject inv.(AssnMem.Rel.inject)) retval_src retval_tgt)
(LOCAL: fully_inject_locals inv.(AssnMem.Rel.inject) locals1_src locals1_tgt)
(MEM: AssnMem.Rel.sem conf_src conf_tgt mem_src mem_tgt inv)
(SRC: return_locals TD retval_src id noret ty locals1_src = Some locals2_src):
exists locals2_tgt,
<<TGT: return_locals TD retval_tgt id noret ty locals1_tgt = Some locals2_tgt>> /\
<<LOCAL: fully_inject_locals inv.(AssnMem.Rel.inject) locals2_src locals2_tgt>>
.
Proof.
unfold return_locals in *.
destruct noret; ss.
{ assert(locals1_src = locals2_src).
{ des_ifs. }
clarify. clear SRC.
esplits; eauto.
- des_ifs. }
des_ifs_safe ss. clarify.
exploit genericvalues_inject.simulation__fit_gv; eauto.
{ apply MEM. }
intro FIT_GV; des.
rewrite FIT_GV.
esplits; eauto.
eapply fully_inject_locals_update; eauto.
Qed.
Lemma return_locals_inject_locals
TD id noret ty inv
retval_src locals1_src locals2_src
retval_tgt locals1_tgt
conf_src conf_tgt mem_src mem_tgt
(RETVAL: lift2_option (genericvalues_inject.gv_inject inv.(AssnMem.Rel.inject)) retval_src retval_tgt)
(LOCAL: inject_locals inv locals1_src locals1_tgt)
(MEM: AssnMem.Rel.sem conf_src conf_tgt mem_src mem_tgt inv)
(SRC: return_locals TD retval_src id noret ty locals1_src = Some locals2_src):
exists locals2_tgt,
<<TGT: return_locals TD retval_tgt id noret ty locals1_tgt = Some locals2_tgt>> /\
<<LOCAL: inject_locals inv locals2_src locals2_tgt>>.
Proof.
unfold return_locals in *.
simtac; try by esplits; eauto.
inv MEM. clear SRC TGT INJECT.
exploit genericvalues_inject.simulation__fit_gv; eauto; []; ii; des.
rewrite x0. esplits; eauto.
apply updateAddAL_inject_locals; ss.
Qed.
Lemma meminj_eq_inject_locals
inv0 inv1 locals_src locals_tgt
(LOCALS: inject_locals inv0 locals_src locals_tgt)
(MEMINJ: inv0.(AssnMem.Rel.inject) = inv1.(AssnMem.Rel.inject))
: inject_locals inv1 locals_src locals_tgt.
Proof.
unfold inject_locals in *.
rewrite <- MEMINJ.
eauto.
Qed.
Lemma returnUpdateLocals_spec
TD id noret clattrs ty va f args Result lc lc' gl:
returnUpdateLocals TD (insn_call id noret clattrs ty va f args) Result lc lc' gl =
match getOperandValue TD Result lc gl with
| Some retval => return_locals TD (Some retval) id noret ty lc'
| None => None
end.
Proof.
unfold returnUpdateLocals.
destruct (getOperandValue TD Result lc gl); ss.
Qed.
Lemma exCallUpdateLocals_spec
TD rt noret rid oResult lc:
exCallUpdateLocals TD rt noret rid oResult lc =
return_locals TD oResult rid noret rt lc.
Proof. destruct oResult; ss. Qed.
Lemma has_false_False
conf_src conf_tgt st_src st_tgt invst assnmem inv
(HAS_FALSE: Hints.Assertion.has_false inv)
(SEM: AssnState.Rel.sem conf_src conf_tgt st_src st_tgt invst assnmem inv):
False.
Proof.
unfold Hints.Assertion.has_false in *.
unfold Hints.Assertion.false_encoding in *.
inv SEM. inv SRC.
apply Exprs.ExprPairSet.mem_2 in HAS_FALSE.
specialize (LESSDEF _ HAS_FALSE).
clear -LESSDEF.
unfold AssnState.Unary.sem_lessdef in *. ss.
compute in LESSDEF.
exploit LESSDEF; eauto; []; ii; des.
inv x. inv x0. inv H4. inv H2. inv H0. inv H.
Qed.
Ltac inject_clarify :=
repeat
match goal with
| [H1: getTypeAllocSize ?TD ?ty = Some ?tsz1,
H2: getTypeAllocSize ?TD ?ty = Some ?tsz2 |- _] =>
rewrite H1 in H2; inv H2
| [H: proj_sumbool ?dec = true |- _] =>
destruct dec; ss; subst
| [H1: getOperandValue (CurTargetData ?conf) ?v (Locals (EC ?st)) ?GL = Some ?gv1,
H2: AssnState.Unary.sem_valueT ?conf ?st ?invst (Exprs.ValueT.lift Exprs.Tag.physical ?v) =
Some ?gv2 |- _] =>
let Hnew := fresh in
assert (Hnew: AssnState.Unary.sem_valueT conf st invst (Exprs.ValueT.lift Exprs.Tag.physical v) = Some gv1);
[ destruct v; [ss; unfold Exprs.IdT.lift; unfold AssnState.Unary.sem_idT in *; eauto | ss] | ];
rewrite Hnew in H2; clear Hnew; inv H2
| [H1: getOperandValue (CurTargetData ?conf) (value_id ?x) (Locals (EC ?st)) ?GL = Some ?gv1 |-
AssnState.Unary.sem_idT ?st ?invst (Exprs.Tag.physical, ?x) = Some ?gv2] =>
let Hnew := fresh in
assert (Hnew: AssnState.Unary.sem_idT st invst (Exprs.Tag.physical, x) = Some gv1); [ss|];
apply Hnew; clear Hnew
end.
Lemma Subset_unary_sem
conf st
invst assnmem gmax public
inv0 inv1
(STATE: AssnState.Unary.sem conf st invst assnmem gmax public inv1)
(SUBSET: Hints.Assertion.Subset_unary inv0 inv1)
: AssnState.Unary.sem conf st invst assnmem gmax public inv0.
Proof.
inv STATE. inv SUBSET.
econs; eauto.
- ii. exploit LESSDEF; eauto.
- inv NOALIAS. inv SUBSET_NOALIAS.
econs.
+ ii. unfold sflib.is_true in *.
exploit DIFFBLOCK; rewrite <- ValueTPairSetFacts.mem_iff in *; eauto.
+ ii. unfold sflib.is_true in *.
exploit NOALIAS0; rewrite <- PtrPairSetFacts.mem_iff in *; eauto.
- ii. exploit UNIQUE; eauto.
- ii. exploit PRIVATE; eauto.
Qed.
Lemma Subset_sem
conf_src st_src
conf_tgt st_tgt
invst assnmem inv0 inv1
(STATE: AssnState.Rel.sem conf_src conf_tgt st_src st_tgt invst assnmem inv1)
(SUBSET: Hints.Assertion.Subset inv0 inv1)
: AssnState.Rel.sem conf_src conf_tgt st_src st_tgt invst assnmem inv0.
Proof.
inv STATE. inv SUBSET.
econs; try (eapply Subset_unary_sem; eauto).
- eapply AtomSetFacts.Empty_s_m_Proper; eauto. unfold flip. inv SUBSET_TGT. ss.
- i. apply MAYDIFF.
destruct (IdTSet.mem id0 (Hints.Assertion.maydiff inv1)) eqn:MEM1; ss.
rewrite <- IdTSetFacts.not_mem_iff in *.
rewrite <- IdTSetFacts.mem_iff in *.
exploit SUBSET_MAYDIFF; eauto. i. congruence.
- eauto.
Qed.
Lemma is_known_nonzero_by_src_Subset
inv0 inv1 v0
(NONZERO: Postcond.is_known_nonzero_by_src inv0 v0 = true)
(SUBSET: Hints.Assertion.Subset inv0 inv1)
:
<<NONZERO: Postcond.is_known_nonzero_by_src inv1 v0 = true>>
.
Proof.
red.
unfold Postcond.is_known_nonzero_by_src in *.
apply ExprPairSetFacts.exists_iff; try by solve_compat_bool.
apply ExprPairSetFacts.exists_iff in NONZERO; try by solve_compat_bool.
unfold ExprPairSet.Exists in *.
des. des_bool. des.
exists x.
split; ss.
- eapply ExprPairSetFacts.In_s_m; eauto.
apply ExprPairSet.E.eq_refl.
apply SUBSET.
- des_ifs_safe. ss.
erewrite AssnState.Subset.inject_value_Subset; eauto.
Qed.
Lemma is_known_nonzero_unary_Subset
inv0 inv1 v0
(NONZERO: Postcond.is_known_nonzero_unary inv0 v0 = true)
(SUBSET: Hints.Assertion.Subset_unary inv0 inv1)
:
<<NONZERO: Postcond.is_known_nonzero_unary inv1 v0 = true>>
.
Proof.
red.
unfold Postcond.is_known_nonzero_unary in *.
apply ExprPairSetFacts.exists_iff; try by solve_compat_bool.
apply ExprPairSetFacts.exists_iff in NONZERO; try by solve_compat_bool.
unfold ExprPairSet.Exists in *.
des. des_bool. des.
exists x.
split; ss.
- eapply ExprPairSetFacts.In_s_m; eauto.
apply ExprPairSet.E.eq_refl.
apply SUBSET.
- des_ifs_safe.
Qed.
Lemma is_known_nonzero_Subset
inv0 inv1 value2
(NONZERO: Postcond.is_known_nonzero_by_src inv0 value2
|| Postcond.is_known_nonzero_unary (Assertion.src inv0) value2)
(SUBSET: Assertion.Subset inv0 inv1)
:
<<NONZERO: Postcond.is_known_nonzero_by_src inv1 value2
|| Postcond.is_known_nonzero_unary (Assertion.src inv1) value2>>
.
Proof.
red.
unfold is_true in *.
des_bool. des.
- erewrite is_known_nonzero_by_src_Subset; try eassumption; ss.
- erewrite is_known_nonzero_unary_Subset; try eassumption; ss.
+ apply orb_true_r.
+ apply SUBSET.
Qed.
Lemma postcond_cmd_inject_event_Subset cmd_src cmd_tgt inv0 inv1
(INJECT_EVENT: Postcond.postcond_cmd_inject_event cmd_src cmd_tgt inv0)
(SUBSET: Hints.Assertion.Subset inv0 inv1)
:
<<INJECT_EVENT: Postcond.postcond_cmd_inject_event cmd_src cmd_tgt inv1>>
.
Proof.
red.
{
destruct cmd_src; destruct cmd_tgt; ss;
des_ifs; ss; try (by eapply is_known_nonzero_Subset; eauto);
unfold proj_sumbool, is_true in *; des_ifs; ss; des_bool; des;
try (eapply AssnState.Subset.inject_value_Subset; eauto); ss.
- apply andb_true_iff.
split; eapply AssnState.Subset.inject_value_Subset; eauto.
- unfold Hints.Assertion.is_private in *. des_ifs.
inv SUBSET. inv SUBSET_SRC.
unfold is_true in *.
AssnState.Subset.conv_mem2In.
exploit SUBSET_PRIVATE; eauto.
- apply andb_true_iff. des_bool. des.
split; eapply AssnState.Subset.inject_value_Subset; eauto.
- apply andb_true_iff.
split; ss.
+ eapply AssnState.Subset.inject_value_Subset; eauto.
+ eapply TODO.list_forallb2_implies; eauto.
i. ss.
repeat match goal with
| [a: ?t * ?s |- _] => destruct a
end.
des_bool; des. clarify. ss.
eapply AssnState.Subset.inject_value_Subset; eauto.
}
Qed.
(* tactics from yoonseung *)
Ltac solve_match_bool :=
repeat (let MATCH := fresh "MATCH" in
match goal with
| [H:match ?c with _ => _ end = true |- _] =>
destruct c eqn:MATCH; try done
| [H:match ?c with _ => _ end = false |- _] =>
destruct c eqn:MATCH; try done
| [H:match ?c with | Some _ => _ | None => _ end =
Some _ |- _] =>
destruct c eqn:MATCH; try done
| [H:match ?c with | Some _ => _ | None => _ end =
None |- _] =>
destruct c eqn:MATCH; try done
| [H:if ?c then Some _ else None = Some _ |- _] =>
destruct c eqn:MATCH; try done
| [H:if ?c then Some _ else None = None |- _] =>
destruct c eqn:MATCH; try done
| [H:if ?c then None else Some _ = Some _ |- _] =>
destruct c eqn:MATCH; try done
| [H:if ?c then None else Some _ = None |- _] =>
destruct c eqn:MATCH; try done
end).
Ltac solve_des_bool :=
match goal with
| [H: _ || _ = false |- _] =>
apply orb_false_iff in H; des
| [H: _ || _ = true |- _] =>
apply orb_true_iff in H; des
| [H: _ && _ = true |- _] =>
apply andb_true_iff in H; des
| [H: _ && _ = false |- _] =>
apply andb_false_iff in H; des
end.
Ltac solve_set_union :=
match goal with
| [H: Exprs.ExprPairSet.In _ (Exprs.ExprPairSet.union _ _) |- _] =>
let IN := fresh "IN" in
apply Exprs.ExprPairSet.union_1 in H; destruct H as [IN|IN]
| [H: ?is_true (Exprs.ValueTPairSet.mem _ (Exprs.ValueTPairSet.union _ _)) |- _] =>
unfold is_true in H;
rewrite Exprs.ValueTPairSetFacts.union_b in H; solve_des_bool
| [H: ?is_true (Exprs.PtrPairSet.mem _ (Exprs.PtrPairSet.union _ _)) |- _] =>
unfold is_true in H;
rewrite Exprs.PtrPairSetFacts.union_b in H; solve_des_bool
end.
Ltac solve_sem_idT :=
try match goal with
| [H: AssnState.Unary.sem_idT _ _ (_, _) = _ |- _] =>
unfold AssnState.Unary.sem_idT in *; ss
| [_:_ |- AssnState.Unary.sem_idT _ _ (_, _) = _] =>
unfold AssnState.Unary.sem_idT in *; ss
end.
Ltac solve_in_filter :=
match goal with
| [H: Exprs.ExprPairSet.In _ (Exprs.ExprPairSet.filter _ _) |- _] =>
let IN := fresh "IN" in
let FILTER := fresh "FILTER" in
apply Exprs.ExprPairSetFacts.filter_iff in H; try (solve_compat_bool; fail); destruct H as [IN FILTER]
end.
Ltac solve_negb_liftpred :=
match goal with
| [H: (negb <*> Postcond.LiftPred.ExprPair _) (_, _) = _ |- _] =>
unfold compose, Postcond.LiftPred.ExprPair in H; simtac; solve_des_bool
| [H: (negb <*> Postcond.LiftPred.ValueTPair _) (_, _) = _ |- _] =>
unfold compose, Postcond.LiftPred.ValueTPair in H; simtac; solve_des_bool
| [H: (negb <*> Postcond.LiftPred.PtrPair _) (_, _) = _ |- _] =>
unfold compose, Postcond.LiftPred.PtrPair in H; simtac; solve_des_bool
end.
Ltac solve_sem_valueT :=
repeat match goal with
| [v: Exprs.ValueT.t |- _] =>
match goal with
| [Hv: AssnState.Unary.sem_valueT _ _ _ v = _ |- _] =>
destruct v
end
end;
ss;
repeat match goal with
| [x:Exprs.IdT.t |- _] =>
let xtag := fresh "xtag" in
destruct x as [xtag x]; destruct xtag; ss
end.
(* TODO: position *)
(* spec for AtomSetImpl_from_list *)
Lemma AtomSetImpl_spec_aux x l s
: x `in` fold_left (flip add) l s <-> In x l \/ x `in` s.
Proof.
split.
- revert x s.
induction l; eauto.
i. ss.
exploit IHl; eauto. i.
unfold flip in *.
des; eauto.
rewrite -> AtomSetFacts.add_iff in *. des; eauto.
- revert x s.
induction l; i.
+ ss. des; done.
+ ss. des; (exploit IHl; [|eauto]); eauto;
right; apply AtomSetFacts.add_iff; eauto.
Qed.
Lemma AtomSetImpl_from_list_spec1 x l
: AtomSetImpl.In x (AtomSetImpl_from_list l) <-> In x l.
Proof.
assert (EQUIV: In x l <-> In x l \/ x `in` empty).
{ split; eauto.
i. des; eauto.
apply AtomSetFacts.empty_iff in H. done.
}
rewrite EQUIV.
apply AtomSetImpl_spec_aux.
Qed.
Lemma AtomSetImpl_from_list_spec2 x l
: ~ AtomSetImpl.In x (AtomSetImpl_from_list l) <-> ~ In x l.
Proof.
split; ii; apply AtomSetImpl_from_list_spec1 in H0; done.
Qed.
Lemma AtomSetImpl_singleton_mem_false x y
: AtomSetImpl.mem x (AtomSetImpl_from_list [y]) = false -> x <> y.
Proof.
i.
apply AtomSetFacts.not_mem_iff in H.
apply AtomSetImpl_from_list_spec2 in H.
apply elim_not_In_cons in H. eauto.
Qed.
(* specs for equiv_except *)
Lemma sem_idT_eq_locals
st0 st1 invst0 x
(LOCALS_EQ : Locals (EC st0) = Locals (EC st1))
: AssnState.Unary.sem_idT st0 invst0 x = AssnState.Unary.sem_idT st1 invst0 x.
Proof.
destruct x as [[] x]; unfold AssnState.Unary.sem_idT in *; ss.
rewrite LOCALS_EQ; eauto.
Qed.
Lemma sem_valueT_eq_locals
conf st0 st1 invst0 v
(LOCALS_EQ: Locals (EC st0) = Locals (EC st1))
: AssnState.Unary.sem_valueT conf st1 invst0 v = AssnState.Unary.sem_valueT conf st0 invst0 v.
Proof.
destruct v; eauto.
eapply sem_idT_eq_locals; eauto.
Qed.
Lemma sem_list_valueT_eq_locals
conf st0 st1 invst0 lsv
(LOCALS_EQ: Locals (EC st0) = Locals (EC st1))
: AssnState.Unary.sem_list_valueT conf st1 invst0 lsv = AssnState.Unary.sem_list_valueT conf st0 invst0 lsv.
Proof.
induction lsv; ss; i.
destruct a as [s v].
exploit sem_valueT_eq_locals; eauto. intro EQ_VAL.
rewrite EQ_VAL. rewrite IHlsv. eauto.
Qed.
Lemma sem_expr_eq_locals_mem
conf st0 st1 invst0 e
(LOCALS_EQ: Locals (EC st0) = Locals (EC st1))
(MEM_EQ: Mem st0 = Mem st1)
: AssnState.Unary.sem_expr conf st1 invst0 e = AssnState.Unary.sem_expr conf st0 invst0 e.
Proof.
Ltac sem_value_st :=
let H' := fresh in
repeat
match goal with
| [LOCALS_EQ: Locals (EC ?st0) = Locals (EC ?st1) |-
context[ AssnState.Unary.sem_valueT ?conf ?st1 ?invst0 ?v ] ] =>
exploit sem_valueT_eq_locals; try exact LOCALS_EQ; intro H';
rewrite H'; clear H'
| [LOCALS_EQ: Locals (EC ?st0) = Locals (EC ?st1) |-
context[ AssnState.Unary.sem_list_valueT ?conf ?st1 ?invst0 ?lsv ] ] =>
exploit sem_list_valueT_eq_locals; try exact LOCALS_EQ; intro H';
rewrite H'; clear H'
end.
destruct e; ss; try rewrite <- MEM_EQ; sem_value_st; eauto.
Qed.
Definition memory_blocks_of (conf: Config) lc ids : list mblock :=
List.flat_map (fun x =>
match lookupAL _ lc x with
| Some gv => GV2blocks gv
| _ => []
end)
(AtomSetImpl.elements ids).
Definition memory_blocks_of_t (conf: Config) st invst idts : list mblock :=
(List.flat_map (fun x =>
match AssnState.Unary.sem_idT st invst x with
| Some gv => GV2blocks gv
| _ => []
end)
(IdTSet.elements idts)).
Definition unique_is_private_unary inv : Prop :=
forall x (UNIQUE: AtomSetImpl.mem x inv.(Hints.Assertion.unique) = true),
IdTSet.mem (Tag.physical, x) inv.(Hints.Assertion.private) = true.
Lemma lift_unlift_le_unary
inv0 inv1 mem
uniqs privs
(LE: AssnMem.Unary.le (AssnMem.Unary.lift mem uniqs privs inv0) inv1)
:
AssnMem.Unary.le inv0 (AssnMem.Unary.unlift inv0 inv1)
.
Proof.
inv LE.
econs; eauto.
Qed.
(* "AssnMem.Rel" is repeated a lot; how about moving this into AssnMem? *)
Lemma lift_unlift_le
inv0 inv1
mem_src uniqs_src privs_src
mem_tgt uniqs_tgt privs_tgt
(NB_LE_SRC: (Mem.nextblock (AssnMem.Unary.mem_parent inv0.(AssnMem.Rel.src)) <=
Mem.nextblock mem_src)%positive)
(NB_LE_TGT: (Mem.nextblock (AssnMem.Unary.mem_parent inv0.(AssnMem.Rel.tgt)) <=
Mem.nextblock mem_tgt)%positive)
(* above two can be achieved from AssnMem.Unary.sem NEXTBLOCK_PARENT *)
(LE: AssnMem.Rel.le (AssnMem.Rel.lift mem_src mem_tgt uniqs_src uniqs_tgt privs_src privs_tgt inv0) inv1)
: AssnMem.Rel.le inv0 (AssnMem.Rel.unlift inv0 inv1).
Proof.
inv LE. ss.
econs; eauto.
- eapply lift_unlift_le_unary; eauto.
- eapply lift_unlift_le_unary; eauto.
- eapply AssnMem.Rel.frozen_shortened; eauto.
Qed.
Ltac des_matchH H :=
repeat
match goal with
| [ H' : context[match ?X with _ => _ end] |- _ ] => check_equal H' H; destruct X
end.
Lemma assnmem_unlift
conf_src mem_src uniqs_src privs_src mem1_src
conf_tgt mem_tgt uniqs_tgt privs_tgt mem1_tgt
inv0 inv1
(MEM_CALLER : AssnMem.Rel.sem conf_src conf_tgt mem_src mem_tgt inv0)
(LIFT : AssnMem.Rel.le (AssnMem.Rel.lift mem_src mem_tgt uniqs_src uniqs_tgt privs_src privs_tgt inv0) inv1)
(MEM_CALLEE : AssnMem.Rel.sem conf_src conf_tgt mem1_src mem1_tgt inv1)
: AssnMem.Rel.sem conf_src conf_tgt mem1_src mem1_tgt (AssnMem.Rel.unlift inv0 inv1).
Proof.
inv MEM_CALLEE.
rename SRC into CALLEE_SRC. rename TGT into CALLEE_TGT.
rename INJECT into CALLEE_INJECT. rename WF into CALLEE_WF.
inv MEM_CALLER.
rename SRC into CALLER_SRC. rename TGT into CALLER_TGT.
rename INJECT into CALLER_INJECT. rename WF into CALLER_WF.
inv LIFT.
rename SRC into LE_SRC. rename TGT into LE_TGT.
econs; eauto; ss.
- (* src *)
inv CALLEE_SRC. inv CALLER_SRC.
econs; eauto; ss.
+ rewrite GMAX. eauto.
+ i. apply PRIVATE_PARENT.
inv LE_SRC.
rewrite <- PRIVATE_PARENT_EQ. ss.
apply in_app. right. eauto.
+ i. exploit MEM_PARENT0; eauto.
intro MLOAD_EQ. rewrite MLOAD_EQ.
inv LE_SRC. ss. subst.
apply MEM_PARENT.
rewrite <- PRIVATE_PARENT_EQ.
apply in_app. right. eauto.
+ ii. exploit UNIQUE_PARENT_MEM; eauto.
des. esplits; eauto.
inv LE_SRC; ss.
rewrite <- UNIQUE_PARENT_EQ.
apply in_app. right. eauto.
+ etransitivity.
{ instantiate (1:= mem_src.(Mem.nextblock)). eauto. }
inv LE_SRC. ss. clarify. Undo 1.
{
subst mem_src.
rewrite NEXTBLOCK_PARENT. reflexivity.
}
- (* tgt *)
inv CALLEE_TGT. inv CALLER_TGT.
econs; eauto; ss.
+ rewrite GMAX. eauto.
+ i. apply PRIVATE_PARENT.
inv LE_TGT.
rewrite <- PRIVATE_PARENT_EQ. ss.
apply in_app. right. eauto.
+ i. exploit MEM_PARENT0; eauto.
intro MLOAD_EQ. rewrite MLOAD_EQ.
inv LE_TGT. ss. subst.
apply MEM_PARENT.
rewrite <- PRIVATE_PARENT_EQ.
apply in_app. right. eauto.
+ ii. exploit UNIQUE_PARENT_MEM; eauto.
des. esplits; eauto.
inv LE_TGT; ss.
rewrite <- UNIQUE_PARENT_EQ.
apply in_app. right. eauto.
+ etransitivity.
{ instantiate (1:= mem_tgt.(Mem.nextblock)). eauto. }
inv LE_TGT. ss. clarify.
- inv CALLEE_WF.
econs; eauto.
rewrite GMAX.
eauto.
Qed.
Lemma assnmem_lift
conf_src mem_src uniqs_src privs_src
conf_tgt mem_tgt privs_tgt
inv
(MEM: AssnMem.Rel.sem conf_src conf_tgt mem_src mem_tgt inv)
(UNIQS_SRC : forall (mptr : mptr) (typ : typ) (align : align) (val : GenericValue),
mload conf_src.(CurTargetData) mem_src mptr typ align = Some val ->
AssnMem.gv_diffblock_with_blocks conf_src val uniqs_src)
(UNIQS_GLOBALS_SRC: forall b, In b uniqs_src -> (inv.(AssnMem.Rel.gmax) < b)%positive)
(PRIVS_SRC: forall b, In b privs_src -> AssnMem.private_block mem_src (AssnMem.Rel.public_src inv.(AssnMem.Rel.inject)) b)
(PRIVS_TGT: forall b, In b privs_tgt -> AssnMem.private_block mem_tgt (AssnMem.Rel.public_tgt inv.(AssnMem.Rel.inject)) b)
: AssnMem.Rel.sem conf_src conf_tgt mem_src mem_tgt
(AssnMem.Rel.lift mem_src mem_tgt uniqs_src [] (* uniqs_tgt *) privs_src privs_tgt inv).
Proof.
inv MEM.
econs; eauto.
- inv SRC.
econs; eauto; ss.
+ i. apply in_app in IN. des.
* apply PRIVS_SRC; eauto.
* exploit PRIVATE_PARENT; eauto.
+ ii. apply in_app in INB. des.
* apply filter_In in INB. des.
exploit PRIVS_SRC; eauto. i. des.
exploit UNIQS_SRC; eauto.
rewrite existsb_exists in *. des.
destruct (Values.eq_block b x0); ss.
subst. eauto.
* exploit UNIQUE_PARENT_MEM; eauto.
+ inv WF0.
i. apply in_app in IN_UNIQUE_PARENT. des.
* apply filter_In in IN_UNIQUE_PARENT. des.
apply UNIQS_GLOBALS_SRC.
rewrite existsb_exists in *. des.
destruct (Values.eq_block b x); ss.
subst. eauto.
* exploit UNIQUE_PARENT_GLOBALS; eauto.
+ apply sublist_app; eauto.
apply filter_sublist.
+ reflexivity.
- inv TGT.
econs; eauto; ss.
+ i. apply in_app in IN. des.
* apply PRIVS_TGT; eauto.
* exploit PRIVATE_PARENT; eauto.
+ ii. apply in_app in INB. des.
* apply filter_In in INB. des. ss.
* exploit UNIQUE_PARENT_MEM; eauto.
+ inv WF0.
i. apply in_app in IN_UNIQUE_PARENT. des.
* apply filter_In in IN_UNIQUE_PARENT. des.
ss.
* exploit UNIQUE_PARENT_GLOBALS; eauto.
+ apply sublist_app; eauto.
apply filter_sublist.
+ reflexivity.
- ss. rewrite TGT_NOUNIQ. rewrite app_nil_r.
reductio_ad_absurdum.
destruct (filter (fun _ : positive => false) privs_tgt) eqn:T; ss.
assert(In p (filter (fun _ : positive => false) privs_tgt)).
{ rewrite T. left; ss. }
apply filter_In in H. des; ss.
Qed.
Lemma positive_lt_plus_one
y gmax
(POS1: (y < gmax + 1)%positive)
(POS2: (gmax < y)%positive)
:
False.
Proof.
replace (gmax + 1)%positive with (Pos.succ gmax)%positive in *; cycle 1.
{ rewrite Pos.add_comm. ss. destruct gmax; ss. }
rewrite Pos.lt_succ_r in POS1.
apply Pos.le_lteq in POS1; eauto.
des.
+ exploit Pos.lt_trans; eauto. ii.
apply Pos.lt_irrefl in x0; ss.
+ clarify.
apply Pos.lt_irrefl in POS2; ss.
Qed.
Lemma unique_const_diffblock
gval1 gval2 conf gmax st i0 cnst
(UNIQUE: AssnState.Unary.sem_unique conf st gmax i0)
(GLOBALS: genericvalues_inject.wf_globals gmax (Globals conf))
(VAL1: lookupAL GenericValue (Locals (EC st)) i0 = Some gval1)
(VAL2: const2GV (CurTargetData conf) (Globals conf) cnst = Some gval2)
:
<<DIFFBLOCK: AssnState.Unary.sem_diffblock conf gval1 gval2>>
.
Proof.
red.
eapply MemAux.wf_globals_const2GV in VAL2; eauto. des.
inv UNIQUE. clear LOCALS MEM. clarify.
{
generalize dependent gval1.
induction gval2; i; ss.
hexploit IHgval2; eauto.
{ des_ifs; ss. des; ss. }
i.
des_ifs.
ii. clarify.
cbn in H1.
des.
- clarify. exploit GLOBALS0; eauto; []; ii; des.
eapply positive_lt_plus_one; eauto.
- exploit H; eauto.
}
Qed.
Lemma valid_ptr_globals_diffblock
conf gmax val val'
(GLOBALS : forall b : Values.block, In b (GV2blocks val) -> (gmax < b)%positive)
(VALID_PTR : memory_props.MemProps.valid_ptrs (gmax + 1)%positive val')
:
<<DIFFBLOCK: AssnState.Unary.sem_diffblock conf val val' >>
.
Proof.
ii.
exploit GLOBALS; eauto; []; intro GMAX; des.
{
clarify.
induction val'; ss.
exploit IHval'; eauto.
{ des_ifs. des. ss. }
des_ifs.
des.
cbn in *. des.
- clarify. exfalso. eapply positive_lt_plus_one; eauto.
- ss.
}
Qed.
Lemma valid_ptr_globals_diffblock2
conf gmax val val'
(GLOBALS : forall b : Values.block, In b (GV2blocks val') -> (gmax < b)%positive)
(VALID_PTR : memory_props.MemProps.valid_ptrs (gmax + 1)%positive val')
:
<<DIFFBLOCK: AssnState.Unary.sem_diffblock conf val val' >>
.
Proof.
ii.
exploit GLOBALS; eauto; []; intro GMAX; des.
{
clarify.
induction val'; ss.
destruct a; ss. destruct v; ss; try (by exploit IHval'; eauto).
des.
- clarify. eapply positive_lt_plus_one; eauto.
- ss. exploit IHval'; eauto.
}
Qed.
Lemma valid_ptr_globals_diffblock_with_blocks
conf gmax val blocks
(GLOBALS : forall b : Values.block, In b blocks -> (gmax < b)%positive)
(VALID_PTR : memory_props.MemProps.valid_ptrs (gmax + 1)%positive val)
:
<<DIFFBLOCK: AssnMem.gv_diffblock_with_blocks conf val blocks>>
.
Proof.
ii.
exploit GLOBALS; eauto; []; intro GMAX; des.
induction val; ss.
cbn in *. destruct a; ss. cbn in *.
unfold compose in *. ss.
destruct v; ss; try (eapply IHval; eauto; fail).
des; clarify.
+ rewrite <- Pplus_one_succ_r in VALID_PTR.
apply Plt_succ_inv in VALID_PTR.
des.
* exploit Plt_trans; eauto. ii.
exploit dom_libs.PositiveSet.MSet.Raw.L.MO.lt_irrefl; eauto.
* clarify.
exploit dom_libs.PositiveSet.MSet.Raw.L.MO.lt_irrefl; eauto.
+ eapply IHval; eauto.
Qed.
Lemma const2GV_mc2undefs
TD gl ty val
(CONST2GV_UNDEF: const2GV TD gl (const_undef ty) = Some val)
: exists mcs, flatten_typ TD ty = Some mcs /\ val = mc2undefs mcs.
Proof.
unfold const2GV, _const2GV, gundef in *. des_ifs.
esplits; eauto.
Qed.
Lemma mc2undefs_vundef
mcs val
(VAL: val = mc2undefs mcs)
: List.Forall (eq Values.Vundef) (List.map fst val) /\
List.map snd val = mcs.
Proof.