-
Notifications
You must be signed in to change notification settings - Fork 108
/
Copy pathGraphProof.thy
2418 lines (2214 loc) · 112 KB
/
GraphProof.thy
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
(*
* Copyright 2020, Data61, CSIRO (ABN 41 687 119 230)
*
* SPDX-License-Identifier: BSD-2-Clause
*)
theory GraphProof
imports TailrecPre GraphLangLemmas "Lib.SplitRule"
begin
declare sep_false_def[symmetric, simp del]
lemma exec_graph_step_Gamma_deterministic:
assumes stacks: "(stack, stack') \<in> exec_graph_step Gamma"
"(stack, stack'') \<in> exec_graph_step (adds ++ Gamma)"
and adds: "\<forall>gf \<in> ran Gamma. \<forall>nn fname inps outps.
(Call nn fname inps outps) \<in> ran (function_graph gf) \<longrightarrow> fname \<notin> dom adds"
shows "stack'' = stack'"
proof -
have adds': "\<And>fname fn fname' fn'. adds fname = Some fn \<Longrightarrow> Gamma fname' = Some fn'
\<Longrightarrow> \<forall>x nn inps outps. function_graph fn' x \<noteq> Some (Call nn fname inps outps)"
using adds
by (metis ranI domI)
show ?thesis using stacks
by (auto simp: all_exec_graph_step_cases
split: graph_function.split_asm dest: adds')
qed
lemmas exec_graph_step_deterministic
= exec_graph_step_Gamma_deterministic[where adds=Map.empty, simplified]
lemma exec_graph_n_Gamma_deterministic:
"(stack, stack') \<in> exec_graph_n Gamma n
\<Longrightarrow> (stack, stack'') \<in> exec_graph_n (adds ++ Gamma) n
\<Longrightarrow> \<forall>gf \<in> ran Gamma. \<forall>nn fname inps outps.
(Call nn fname inps outps) \<in> ran (function_graph gf) \<longrightarrow> fname \<notin> dom adds
\<Longrightarrow> stack'' = stack'"
using exec_graph_step_Gamma_deterministic
apply (induct n arbitrary: stack' stack'', simp_all add: exec_graph_n_def)
apply blast
done
lemmas exec_graph_n_deterministic
= exec_graph_n_Gamma_deterministic[where adds=Map.empty, simplified]
lemma step_implies_continuing:
"(stack, stack') \<in> exec_graph_step Gamma
\<Longrightarrow> continuing stack"
by (simp add: exec_graph_step_def continuing_def
split: list.split_asm prod.split_asm next_node.split_asm)
lemma exec_trace_Gamma_deterministic:
"trace \<in> exec_trace Gamma f
\<Longrightarrow> trace' \<in> exec_trace (adds ++ Gamma) f
\<Longrightarrow> \<forall>gf \<in> ran Gamma. \<forall>nn fname inps outps.
(Call nn fname inps outps) \<in> ran (function_graph gf) \<longrightarrow> fname \<notin> dom adds
\<Longrightarrow> trace 0 = trace' 0
\<Longrightarrow> trace n = trace' n"
apply (induct n)
apply simp
apply (clarsimp simp: exec_trace_def nat_trace_rel_def)
apply (drule_tac x=n in spec)+
apply (case_tac "trace' n", clarsimp+)
apply (case_tac "trace (Suc n)", simp_all)
apply (auto dest: step_implies_continuing)[1]
apply (clarsimp simp: step_implies_continuing)
apply (metis exec_graph_step_Gamma_deterministic)
done
lemmas exec_trace_deterministic
= exec_trace_Gamma_deterministic[where adds=Map.empty, simplified]
lemma exec_trace_nat_trace:
"tr \<in> exec_trace Gamma f \<Longrightarrow> tr \<in> nat_trace_rel continuing (exec_graph_step Gamma)"
by (clarsimp simp add: exec_trace_def)
abbreviation(input)
"exec_double_trace Gamma1 f1 Gamma2 f2 trace1 trace2
\<equiv> (trace1 \<in> exec_trace Gamma1 f1 \<and> trace2 \<in> exec_trace Gamma2 f2)"
definition
trace_refine :: "bool \<Rightarrow> (state \<Rightarrow> state \<Rightarrow> bool) \<Rightarrow> trace \<Rightarrow> trace \<Rightarrow> bool"
where
"trace_refine precise rel tr tr' = (case (trace_end tr, trace_end tr') of
(None, None) \<Rightarrow> True
| (_, None) \<Rightarrow> \<not> precise
| (_, Some [(Err, _, _)]) \<Rightarrow> True
| (Some [(Ret, st, _)], Some [(Ret, st', _)]) \<Rightarrow> rel st st'
| _ \<Rightarrow> False)"
definition
"switch_val f x y = (\<lambda>z. if f z then y z else x z)"
definition
"tuple_switch vs = switch_val id (\<lambda>_. fst vs) (\<lambda>_. snd vs)"
lemma tuple_switch_simps[simp]:
"tuple_switch vs True = snd vs"
"tuple_switch vs False = fst vs"
by (auto simp: tuple_switch_def switch_val_def)
definition
"function_outputs_st f st = acc_vars (function_outputs f) st"
definition
trace_refine2 :: "bool \<Rightarrow> graph_function \<times> graph_function
\<Rightarrow> ((bool \<times> bool \<Rightarrow> variable list) \<Rightarrow> bool) \<Rightarrow> trace \<Rightarrow> trace \<Rightarrow> bool"
where
"trace_refine2 precise gfs orel tr tr' = trace_refine precise (\<lambda>st st'.
orel (switch_val fst
(tuple_switch (acc_vars (function_inputs (fst gfs)) (fst (snd (hd (the (tr 0))))),
function_outputs_st (fst gfs) st) o snd)
(tuple_switch (acc_vars (function_inputs (snd gfs)) (fst (snd (hd (the (tr' 0))))),
function_outputs_st (snd gfs) st') o snd)
)) tr tr'"
definition
graph_function_refine :: "bool \<Rightarrow> (string \<Rightarrow> graph_function option) \<Rightarrow> string
\<Rightarrow> (string \<Rightarrow> graph_function option) \<Rightarrow> string
\<Rightarrow> ((bool \<Rightarrow> variable list) \<Rightarrow> bool)
\<Rightarrow> ((bool \<times> bool \<Rightarrow> variable list) \<Rightarrow> bool)
\<Rightarrow> bool"
where
"graph_function_refine precise Gamma1 f1 Gamma2 f2 irel orel
= (\<forall>xs ys tr tr' gf gf'. tr \<in> exec_trace Gamma1 f1 \<and> tr' \<in> exec_trace Gamma2 f2
\<and> Gamma1 f1 = Some gf \<and> tr 0 = Some [init_state f1 gf xs]
\<and> Gamma2 f2 = Some gf' \<and> tr' 0 = Some [init_state f2 gf' ys]
\<and> length xs = length (function_inputs gf) \<and> length ys = length (function_inputs gf')
\<and> irel (tuple_switch (xs, ys))
\<longrightarrow> trace_refine precise (\<lambda>st st'. orel (switch_val fst
(tuple_switch (xs, function_outputs_st gf st) o snd)
(tuple_switch (ys, function_outputs_st gf' st') o snd)
)) tr tr')"
lemma graph_function_refine_trace:
"graph_function_refine prec Gamma1 f1 Gamma2 f2 irel orel
= (\<forall>tr tr' xs ys gf gf'. exec_double_trace Gamma1 f1 Gamma2 f2 tr tr'
\<and> tr 0 \<noteq> None \<and> tr' 0 \<noteq> None \<and> Gamma1 f1 = Some gf \<and> Gamma2 f2 = Some gf'
\<and> the (tr 0) = [init_state f1 gf xs] \<and> the (tr' 0) = [init_state f2 gf' ys]
\<and> length xs = length (function_inputs gf) \<and> length ys = length (function_inputs gf')
\<and> irel (tuple_switch (xs, ys))
\<longrightarrow> trace_refine prec (\<lambda>st st'. orel (switch_val fst
(tuple_switch (xs, function_outputs_st gf st) o snd)
(tuple_switch (ys, function_outputs_st gf' st') o snd)
)) tr tr')"
apply (clarsimp simp: graph_function_refine_def)
apply (simp only: ex_simps[symmetric] all_simps[symmetric]
cong: conj_cong)
apply auto
done
definition
trace_addr :: "trace \<Rightarrow> nat \<Rightarrow> next_node option"
where
"trace_addr tr n = (case tr n of Some [(nn, _, _)] \<Rightarrow> Some nn | _ \<Rightarrow> None)"
lemma trace_addr_SomeD:
"trace_addr tr n = Some nn \<Longrightarrow> \<exists>st g. tr n = Some [(nn, st, g)]"
by (simp add: trace_addr_def split: option.split_asm list.split_asm prod.split_asm)
lemma trace_addr_SomeI:
"\<exists>x. tr n = Some [(nn, x)] \<Longrightarrow> trace_addr tr n = Some nn"
by (clarsimp simp add: trace_addr_def)
type_synonym restrs = "nat \<Rightarrow> nat set"
definition
restrs_condition :: "trace \<Rightarrow> restrs \<Rightarrow> nat \<Rightarrow> bool"
where
"restrs_condition tr restrs n = (\<forall>m. card {i. i < n \<and> trace_addr tr i = Some (NextNode m)} \<in> restrs m)"
definition
succ_restrs :: "next_node \<Rightarrow> restrs \<Rightarrow> restrs"
where
"succ_restrs nn rs = (case nn of NextNode n \<Rightarrow> rs (n := {x. x - 1 \<in> rs n}) | _ \<Rightarrow> rs)"
abbreviation
"succ_restrs' n \<equiv> succ_restrs (NextNode n)"
definition
pred_restrs :: "next_node \<Rightarrow> restrs \<Rightarrow> restrs"
where
"pred_restrs nn rs = (case nn of NextNode n \<Rightarrow> rs (n := {x. Suc x \<in> rs n}) | _ \<Rightarrow> rs)"
abbreviation
"pred_restrs' n \<equiv> pred_restrs (NextNode n)"
definition
trace_bottom_addr :: "trace \<Rightarrow> nat \<Rightarrow> next_node option"
where
"trace_bottom_addr tr n = (case tr n of Some [] \<Rightarrow> None
| Some xs \<Rightarrow> Some (fst (List.last xs)) | None \<Rightarrow> None)"
definition
double_trace_imp :: "(trace \<times> trace \<Rightarrow> bool) \<Rightarrow> (trace \<times> trace \<Rightarrow> bool)
\<Rightarrow> (trace \<times> trace \<Rightarrow> bool)"
where "double_trace_imp P Q = (\<lambda>(tr, tr'). P (tr, tr') \<longrightarrow> Q (tr, tr'))"
type_synonym visit_addr = "bool \<times> next_node \<times> restrs"
definition
restrs_eventually_condition :: "trace \<Rightarrow> restrs \<Rightarrow> bool"
where
"restrs_eventually_condition tr restrs = (\<exists>n. tr n \<noteq> None
\<and> (\<forall>m \<ge> n. restrs_condition tr restrs m))"
definition
visit :: "trace \<Rightarrow> next_node \<Rightarrow> restrs \<Rightarrow> state option"
where
"visit tr n restrs = (if \<exists>i. restrs_condition tr restrs i \<and> trace_addr tr i = Some n
then Some (fst (snd (hd (the (tr (LEAST i. restrs_condition tr restrs i \<and> trace_addr tr i = Some n))))))
else None)"
definition
pc :: "next_node \<Rightarrow> restrs \<Rightarrow> trace \<Rightarrow> bool"
where
"pc n restrs tr = (visit tr n restrs \<noteq> None)"
abbreviation
"pc' n \<equiv> pc (NextNode n)"
definition
double_visit :: "visit_addr \<Rightarrow> trace \<times> trace \<Rightarrow> state option"
where
"double_visit = (\<lambda>(side, nn, restrs) (tr, tr').
visit (if side then tr' else tr) nn restrs)"
definition
double_pc :: "visit_addr \<Rightarrow> trace \<times> trace \<Rightarrow> bool"
where
"double_pc = (\<lambda>(side, nn, restrs) (tr, tr').
(visit (if side then tr' else tr) nn restrs) \<noteq> None)"
definition
related_pair :: "visit_addr \<Rightarrow> (state \<Rightarrow> 'a)
\<Rightarrow> visit_addr \<Rightarrow> (state \<Rightarrow> 'b)
\<Rightarrow> ('a \<times> 'b \<Rightarrow> bool)
\<Rightarrow> (trace \<times> trace) \<Rightarrow> bool"
where
"related_pair v1 expr1 v2 expr2 rel trs
= rel (expr1 (the (double_visit v1 trs)), expr2 (the (double_visit v2 trs)))"
abbreviation
"equals v1 expr1 v2 expr2
\<equiv> related_pair v1 expr1 v2 expr2 (split (=))"
abbreviation
"equals' n1 restrs1 expr1 n2 restrs2 expr2
\<equiv> equals (False, NextNode n1, restrs1) expr1 (True, NextNode n2, restrs2) expr2"
definition
restr_trace_refine :: "bool \<Rightarrow> (string \<Rightarrow> graph_function option) \<Rightarrow> string
\<Rightarrow> (string \<Rightarrow> graph_function option) \<Rightarrow> string
\<Rightarrow> restrs \<Rightarrow> restrs
\<Rightarrow> ((bool \<times> bool \<Rightarrow> variable list) \<Rightarrow> bool) \<Rightarrow> trace \<Rightarrow> trace \<Rightarrow> bool"
where
"restr_trace_refine precise Gamma1 f1 Gamma2 f2 restrs1 restrs2 orel
tr tr' = (\<forall>gf gf'. exec_double_trace Gamma1 f1 Gamma2 f2 tr tr'
\<and> Gamma1 f1 = Some gf \<and> Gamma2 f2 = Some gf'
\<and> (\<exists>xs. the (tr 0) = [init_state f1 gf xs]
\<and> length xs = length (function_inputs gf))
\<and> (\<exists>ys. the (tr' 0) = [init_state f2 gf' ys]
\<and> length ys = length (function_inputs gf'))
\<and> restrs_eventually_condition tr restrs1
\<and> restrs_eventually_condition tr' restrs2
\<longrightarrow> trace_refine2 precise (gf, gf') orel tr tr')"
definition
restrs_list :: "(nat \<times> nat list) list \<Rightarrow> restrs"
where
"restrs_list rs = (\<lambda>i. fold (\<lambda>(n, xs) S.
if i = n then set xs \<inter> S else S) rs UNIV)"
definition
fill_in_below :: "nat list \<Rightarrow> nat list"
where
"fill_in_below xs = [0 ..< fold max (map Suc xs) 0]"
definition
restrs_visit :: "(nat \<times> nat list) list
\<Rightarrow> next_node \<Rightarrow> graph_function
\<Rightarrow> (nat \<times> nat list) list"
where
"restrs_visit xs nn gf = map (\<lambda>(m, xsf). if (nn, NextNode m) \<in> rtrancl (reachable_step' gf)
then (m, (fill_in_below xsf)) else (m, xsf)) xs"
definition
eqs :: "(('a \<times> (variable list \<Rightarrow> variable)) \<times> ('a \<times> (variable list \<Rightarrow> variable))) list
\<Rightarrow> ('a \<Rightarrow> nat option) \<Rightarrow> ('a \<Rightarrow> variable list) \<Rightarrow> bool"
where
"eqs xs lens vs = ((\<forall>((laddr, lval), (raddr, rval)) \<in> set xs. lval (vs laddr) = rval (vs raddr))
\<and> (\<forall>addr. lens addr \<noteq> None \<longrightarrow> length (vs addr) = the (lens addr)))"
definition
visit_restrs_preds_raw :: "next_node \<Rightarrow> restrs
\<Rightarrow> (next_node \<times> restrs) list \<Rightarrow> bool"
where
"visit_restrs_preds_raw nn restrs preds =
(\<forall>tr i. trace_addr tr i = Some nn \<and> restrs_condition tr restrs i
\<longrightarrow> (\<forall>j. trace_addr tr j = Some nn \<and> restrs_condition tr restrs j \<longrightarrow> j = i)
\<and> (\<exists>j nn' restrs'. j < i
\<and> (nn', restrs') \<in> set preds \<and> trace_addr tr j = Some nn'
\<and> restrs_condition tr restrs' j))"
definition
visit_restrs_preds :: "visit_addr \<Rightarrow> visit_addr list \<Rightarrow> bool"
where
"visit_restrs_preds vis preds = (case vis of (side, nn, restrs)
\<Rightarrow> (fst ` set preds \<subseteq> {side})
\<and> visit_restrs_preds_raw nn restrs (map snd preds))"
lemma visit_known:
"tr i = Some [(nn, st, g)]
\<Longrightarrow> restrs_condition tr restrs i
\<Longrightarrow> (\<forall>j < i. trace_addr tr j = Some nn \<longrightarrow> \<not> restrs_condition tr restrs j)
\<Longrightarrow> visit tr nn restrs = Some st"
apply (clarsimp simp: visit_def)
apply (subst Least_equality, blast intro: trace_addr_SomeI)
apply (blast intro: linorder_not_le[THEN iffD1])
apply (auto simp: trace_addr_SomeI)
done
lemma fold_invariant:
"(\<forall>x \<in> set xs. \<forall> s. g (f x s) = g s) \<Longrightarrow> g s = v \<Longrightarrow> g (fold f xs s) = v"
by (induct xs arbitrary: s, simp_all)
lemma var_acc_var_upd:
"var_acc s (var_upd s' v st)
= (if s = s' then v else var_acc s st)"
by (cases st, simp add: var_acc_def var_upd_def)
lemma var_acc_save_vals_distinct:
"distinct xs \<Longrightarrow> length xs = length vs
\<Longrightarrow> map (\<lambda>i. var_acc i (save_vals xs vs st)) xs = vs"
apply (induct xs arbitrary: vs st)
apply simp
apply (case_tac vs, simp_all add: save_vals_def)
apply (rule fold_invariant)
apply (clarsimp simp: var_acc_var_upd elim!: in_set_zipE)
apply (simp add: var_acc_var_upd)
done
definition
wf_graph_function :: "graph_function \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
where
"wf_graph_function f ilen olen = (case f of GraphFunction inputs outputs graph ep
\<Rightarrow> distinct inputs \<and> distinct outputs \<and> length inputs = ilen \<and> length outputs = olen
\<and> (\<forall>nn i. (nn, NextNode i) \<in> reachable_step' f \<longrightarrow> i \<in> dom graph)
\<and> ep \<in> dom graph)"
lemma wf_graph_function_acc_vars_save_vals:
"\<lbrakk> wf_graph_function f ilen olen; xs = function_inputs f; length vs = ilen \<rbrakk>
\<Longrightarrow> acc_vars xs (save_vals xs vs st) = vs"
by (cases f, simp add: acc_vars_def wf_graph_function_def
var_acc_save_vals_distinct)
lemma wf_graph_function_length_function_inputs:
"wf_graph_function f ilen olen \<Longrightarrow> length (function_inputs f) = ilen"
by (cases f, simp_all add: wf_graph_function_def)
lemma graph_function_refine_trace2:
"Gamma1 f1 = Some gf \<Longrightarrow> wf_graph_function gf ilen1 olen1
\<Longrightarrow> Gamma2 f2 = Some gf' \<Longrightarrow> wf_graph_function gf' ilen2 olen2
\<Longrightarrow> graph_function_refine prec Gamma1 f1 Gamma2 f2 irel orel
= (\<forall>tr tr'. exec_double_trace Gamma1 f1 Gamma2 f2 tr tr'
\<and> Gamma1 f1 = Some gf \<and> Gamma2 f2 = Some gf'
\<and> (\<exists>xs ys. the (tr 0) = [init_state f1 gf xs]
\<and> length xs = length (function_inputs gf)
\<and> the (tr' 0) = [init_state f2 gf' ys]
\<and> length ys = length (function_inputs gf')
\<and> irel (tuple_switch (xs, ys)))
\<longrightarrow> trace_refine2 prec (gf, gf') orel tr tr')"
apply (clarsimp simp: graph_function_refine_trace trace_refine2_def)
apply (frule wf_graph_function_length_function_inputs[where f=gf])
apply (frule wf_graph_function_length_function_inputs[where f=gf'])
apply (simp add: imp_conjL)
apply (rule arg_cong[where f=All] ext imp_cong[OF refl])+
apply (clarsimp dest!: exec_trace_init)
apply (safe, simp_all add: init_state_def wf_graph_function_acc_vars_save_vals)
done
lemma exec_trace_step_reachable_induct:
"tr \<in> exec_trace Gamma fn \<Longrightarrow> Gamma fn = Some gf
\<Longrightarrow> trace_bottom_addr tr (Suc i) = Some addr
\<Longrightarrow> (\<forall>n. trace_bottom_addr tr i = Some (NextNode n) \<longrightarrow> n \<in> dom (function_graph gf))
\<longrightarrow> trace_bottom_addr tr i \<noteq> None
\<and> (trace_bottom_addr tr i = Some addr \<and> trace_addr tr (Suc i) = None
\<or> trace_addr tr (Suc i) = Some addr
\<and> (the (trace_bottom_addr tr i), addr) \<in> reachable_step' gf)"
apply (frule_tac i=i in exec_trace_invariant')
apply (frule_tac i=i in exec_trace_step_cases)
apply (simp add: all_exec_graph_step_cases, safe dest!: trace_addr_SomeD,
simp_all add: trace_bottom_addr_def del: last.simps)
apply (auto simp: exec_graph_invariant_Cons reachable_step_def
get_state_function_call_def
split: graph_function.split_asm,
auto split: next_node.splits option.split node.splits if_split_asm
simp: trace_addr_def neq_Nil_conv)
done
lemma exec_trace_step_reachable_induct2:
"tr \<in> exec_trace Gamma fn \<Longrightarrow> Gamma fn = Some gf
\<Longrightarrow> wf_graph_function gf ilen olen
\<Longrightarrow> trace_bottom_addr tr (Suc i) = Some addr
\<Longrightarrow> (\<forall>n. trace_bottom_addr tr i = Some (NextNode n) \<longrightarrow> n \<in> dom (function_graph gf))
\<and> trace_bottom_addr tr i \<noteq> None
\<and> (trace_bottom_addr tr i = Some addr \<and> trace_addr tr (Suc i) = None
\<or> trace_addr tr (Suc i) = Some addr
\<and> (the (trace_bottom_addr tr i), addr) \<in> reachable_step' gf)"
apply (induct i arbitrary: addr)
apply (frule(2) exec_trace_step_reachable_induct)
apply (clarsimp simp: exec_trace_def trace_bottom_addr_def
wf_graph_function_def split: graph_function.split_asm)
apply (frule(2) exec_trace_step_reachable_induct)
apply atomize
apply (case_tac "trace_bottom_addr tr (Suc i)", simp_all)
apply clarsimp
apply (frule_tac i=i in exec_trace_step_reachable_induct, simp)
apply clarsimp
apply (erule disjE, simp_all)
apply (simp add: wf_graph_function_def split: graph_function.split_asm)
apply auto
done
lemma Collect_less_Suc:
"{i. i < Suc n \<and> P i} = {i. i < n \<and> P i} \<union> (if P n then {n} else {})"
by (auto simp: less_Suc_eq)
lemma trace_addr_trace_bottom_addr_eq:
"trace_addr tr i = Some addr \<Longrightarrow> trace_bottom_addr tr i = Some addr"
by (clarsimp simp: trace_bottom_addr_def dest!: trace_addr_SomeD)
lemma reachable_trace_induct:
"tr \<in> exec_trace Gamma fname \<Longrightarrow> Gamma fname = Some f
\<Longrightarrow> wf_graph_function f ilen olen
\<Longrightarrow> trace_addr tr i = Some nn
\<Longrightarrow> i + k \<le> j
\<Longrightarrow> (\<forall>nn'. trace_bottom_addr tr (i + k) = Some nn'
\<longrightarrow> (Some nn' \<in> trace_addr tr ` {i .. i + k})
\<and> (nn, nn') \<in> rtrancl (reachable_step' f \<inter> {(x, y). Some x \<in> trace_addr tr ` {i ..< j}})
\<and> (k \<noteq> 0 \<and> trace_addr tr (i + k) \<noteq> None
\<longrightarrow> (nn, nn') \<in> trancl (reachable_step' f \<inter> {(x, y). Some x \<in> trace_addr tr ` {i ..< j}})))"
apply (induct k)
apply (clarsimp simp: trace_bottom_addr_def dest!: trace_addr_SomeD)
apply clarsimp
apply (frule(3) exec_trace_step_reachable_induct2)
apply clarsimp
apply (subgoal_tac "i + k \<in> {i ..< j} \<and> Suc (i + k) \<in> {i .. Suc (i + k)}")
apply (rule conjI)
apply (elim disjE conjE)
apply simp
apply (blast intro: sym)
apply (rule conjI)
apply (elim disjE conjE)
apply simp
apply (erule rtrancl_into_rtrancl)
apply simp
apply clarify
apply (elim disjE conjE)
apply simp
apply (erule rtrancl_into_trancl1)
apply simp
apply simp
done
lemma reachable_trace:
"tr \<in> exec_trace Gamma fname \<Longrightarrow> Gamma fname = Some f
\<Longrightarrow> wf_graph_function f ilen olen
\<Longrightarrow> trace_addr tr i = Some nn
\<Longrightarrow> trace_addr tr j = Some nn'
\<Longrightarrow> i < j
\<Longrightarrow> (nn, nn') \<in> trancl (reachable_step' f \<inter> {(x, y). Some x \<in> trace_addr tr ` {i ..< j}})"
apply (frule(3) reachable_trace_induct[where k="j - i" and j=j])
apply simp
apply (clarsimp simp: trace_bottom_addr_def dest!: trace_addr_SomeD)
done
lemma reachable_trace_eq:
"tr \<in> exec_trace Gamma fname \<Longrightarrow> Gamma fname = Some f
\<Longrightarrow> wf_graph_function f ilen olen
\<Longrightarrow> trace_addr tr i = Some nn
\<Longrightarrow> trace_addr tr j = Some nn'
\<Longrightarrow> i \<le> j
\<Longrightarrow> (nn, nn') \<in> rtrancl (reachable_step' f \<inter> {(x, y). Some x \<in> trace_addr tr ` {i ..< j}})"
apply (cases "i = j", simp_all)
apply (rule trancl_into_rtrancl)
apply (fastforce elim: reachable_trace)
done
lemma restrs_single_visit_neq:
"tr \<in> exec_trace Gamma fname \<Longrightarrow> Gamma fname = Some f
\<Longrightarrow> wf_graph_function f ilen olen
\<Longrightarrow> trace_addr tr i = Some nn
\<Longrightarrow> restrs_condition tr restrs i
\<Longrightarrow> trace_addr tr j = Some nn
\<Longrightarrow> restrs_condition tr restrs j
\<Longrightarrow> \<forall>x. NextNode x \<in> set cuts \<longrightarrow> (\<exists>y. restrs x \<subseteq> {y})
\<Longrightarrow> (nn, nn) \<notin> trancl (reachable_step' f \<inter> {(x, y). x \<notin> set cuts})
\<Longrightarrow> i < j \<Longrightarrow> False"
apply (frule(5) reachable_trace)
apply (erule notE, erule trancl_mono)
apply clarsimp
apply (case_tac a, simp_all)
apply (drule spec, drule(1) mp, clarsimp)
apply (clarsimp simp: restrs_condition_def)
apply (drule spec, drule(1) subsetD[rotated])+
apply clarsimp
apply (frule card_seteq[rotated -1, OF eq_imp_le], simp, clarsimp)[1]
apply (blast dest: linorder_not_less[THEN iffD2])
apply (simp add: reachable_step_def)+
done
lemma restrs_single_visit:
"tr \<in> exec_trace Gamma fname \<Longrightarrow> Gamma fname = Some f
\<Longrightarrow> wf_graph_function f ilen olen
\<Longrightarrow> trace_addr tr i = Some nn
\<Longrightarrow> restrs_condition tr restrs i
\<Longrightarrow> trace_addr tr j = Some nn
\<Longrightarrow> restrs_condition tr restrs j
\<Longrightarrow> \<forall>x. NextNode x \<in> set cuts \<longrightarrow> (\<exists>y. restrs x \<subseteq> {y})
\<Longrightarrow> (nn, nn) \<notin> trancl (reachable_step' f \<inter> {(x, y). x \<notin> set cuts})
\<Longrightarrow> i = j"
by (metis restrs_single_visit_neq linorder_neq_iff)
lemma restrs_single_visit2:
"trace_addr tr i = Some nn
\<Longrightarrow> trace_addr tr j = Some nn
\<Longrightarrow> tr \<in> exec_trace Gamma fname
\<Longrightarrow> Gamma fname = Some f
\<Longrightarrow> wf_graph_function f ilen olen
\<Longrightarrow> restrs_condition tr restrs i
\<Longrightarrow> restrs_condition tr restrs j
\<Longrightarrow> \<forall>x. NextNode x \<in> set cuts \<longrightarrow> (\<exists>y. restrs x \<subseteq> {y})
\<Longrightarrow> (nn, nn) \<notin> trancl (reachable_step' f \<inter> {(x, y). x \<notin> set cuts})
\<Longrightarrow> i = j"
by (metis restrs_single_visit)
lemma not_trancl_converse_step:
"converse stp `` {y} \<subseteq> S
\<Longrightarrow> \<forall>z \<in> S. (x, z) \<notin> rtrancl (stp \<inter> constraint)
\<Longrightarrow> (x, y) \<notin> trancl (stp \<inter> constraint)"
using tranclD2 by fastforce
lemma reachable_order_mono:
"(nn, nn') \<in> rtrancl R
\<Longrightarrow> (\<forall>(nn, nn') \<in> R. order nn \<le> order nn')
\<Longrightarrow> order nn \<le> (order nn' :: 'a :: linorder)"
apply (induct rule: rtrancl.induct)
apply simp
apply (blast intro: order_trans)
done
lemma not_reachable_by_order:
"(\<forall>(nn, nn') \<in> R. order nn \<le> order nn')
\<Longrightarrow> order nn > (order nn' :: 'a :: linorder)
\<Longrightarrow> (nn, nn') \<notin> rtrancl R"
by (metis reachable_order_mono linorder_not_less)
lemma Collect_less_nat:
"(n :: nat) - 1 = m
\<Longrightarrow> {i. i < n \<and> P i} = (if n = 0 then {} else {i. i < m \<and> P i} \<union> (if P m then {m} else {}))"
by (cases n, simp_all add: Collect_less_Suc)
lemma test_restrs_condition:
"\<lbrakk> graph = [ 1 \<mapsto> Call (NextNode 2) ''foo'' [] [], 2 \<mapsto> Basic (NextNode 3) [], 3 \<mapsto> Basic Ret [] ];
Gamma ''bar'' = Some (GraphFunction [] [] graph 1);
tr \<in> exec_trace Gamma ''bar'';
Gamma ''foo'' = Some (GraphFunction [] [] [ 1 \<mapsto> Basic Ret []] 1) \<rbrakk>
\<Longrightarrow> restrs_condition tr (restrs_list [(1, [1])]) 5"
apply (frule exec_trace_init, elim exE conjE)
apply (frule_tac i=0 in exec_trace_step_cases, clarsimp)
apply (clarsimp simp: all_exec_graph_step_cases)
apply (frule_tac i=1 in exec_trace_step_cases, clarsimp)
apply (clarsimp simp: all_exec_graph_step_cases upd_vars_def save_vals_def init_vars_def)
apply (frule_tac i=2 in exec_trace_step_cases, clarsimp)
apply (clarsimp simp: all_exec_graph_step_cases upd_vars_def save_vals_def init_vars_def TWO return_vars_def)
apply (frule_tac i=3 in exec_trace_step_cases, clarsimp)
apply (clarsimp simp: all_exec_graph_step_cases upd_vars_def save_vals_def init_vars_def)
apply (frule_tac i=4 in exec_trace_step_cases, clarsimp)
apply (clarsimp simp: all_exec_graph_step_cases upd_vars_def save_vals_def init_vars_def)
apply (frule_tac i=5 in exec_trace_step_cases, clarsimp)
apply (clarsimp simp: all_exec_graph_step_cases upd_vars_def save_vals_def init_vars_def)
apply (simp add: restrs_condition_def Collect_less_nat Collect_conv_if trace_addr_def
restrs_list_def)
done
primrec
first_reached_prop :: "visit_addr list
\<Rightarrow> (visit_addr \<Rightarrow> (trace \<times> trace) \<Rightarrow> bool)
\<Rightarrow> (trace \<times> trace) \<Rightarrow> bool"
where
"first_reached_prop [] p trs = False"
| "first_reached_prop (v # vs) p trs =
(if double_pc v trs then p v trs
else first_reached_prop vs p trs)"
definition
get_function_call :: "(graph_function \<times> graph_function)
\<Rightarrow> visit_addr
\<Rightarrow> (next_node \<times> string \<times> (state \<Rightarrow> variable) list \<times> string list) option"
where
"get_function_call gfs x = (case x of (side, NextNode n, restrs)
\<Rightarrow> (case function_graph (tuple_switch gfs side) n of
Some (Call nn fname inputs outputs) \<Rightarrow> Some (nn, fname, inputs, outputs)
| _ \<Rightarrow> None) | _ \<Rightarrow> None)"
definition
func_inputs_match :: "(graph_function \<times> graph_function)
\<Rightarrow> ((bool \<Rightarrow> variable list) \<Rightarrow> bool)
\<Rightarrow> visit_addr \<Rightarrow> visit_addr
\<Rightarrow> (trace \<times> trace) \<Rightarrow> bool"
where
"func_inputs_match gfs rel vis1 vis2 trs = (case (get_function_call gfs vis1,
get_function_call gfs vis2) of (Some (_, _, inputs1, _), Some (_, _, inputs2, _))
\<Rightarrow> rel (\<lambda>x. map (\<lambda>f. f (the (double_visit (tuple_switch (vis1, vis2) x) trs)))
(tuple_switch (inputs1, inputs2) x))
| _ \<Rightarrow> False)"
definition
succ_visit :: "next_node \<Rightarrow> visit_addr \<Rightarrow> visit_addr"
where
"succ_visit nn vis = (case vis of (side, nn', restrs)
\<Rightarrow> (side, nn', succ_restrs nn restrs))"
lemma fst_succ_visit[simp]:
"fst (succ_visit nn v) = fst v"
by (simp add: succ_visit_def split_def split: next_node.split)
lemma wf_graph_function_init_extract:
"\<lbrakk> wf_graph_function f ilen olen; tr \<in> exec_trace Gamma fn; Gamma fn = Some f;
the (tr 0) = [init_state fn f xs];
\<forall>n. 0 \<in> restrs n; length xs = ilen \<rbrakk> \<Longrightarrow>
acc_vars (function_inputs f) (the (visit tr (NextNode (entry_point f)) restrs)) = xs"
apply (cases f)
apply (frule wf_graph_function_length_function_inputs)
apply (frule exec_trace_init)
apply (clarsimp simp: visit_known restrs_condition_def init_state_def
wf_graph_function_acc_vars_save_vals)
done
lemma exec_trace_reachable_step:
"tr \<in> exec_trace Gamma fn \<Longrightarrow> Gamma fn = Some f
\<Longrightarrow> trace_addr tr (Suc i) = Some nn'
\<Longrightarrow> nn' \<noteq> Err
\<Longrightarrow> \<exists>nn. (nn, nn') \<in> reachable_step' f"
apply (clarsimp dest!: trace_addr_SomeD)
apply (frule(1) exec_trace_invariant)
apply (subgoal_tac "\<forall>stack. tr i = Some stack \<longrightarrow> exec_graph_invariant Gamma fn stack")
prefer 2
apply (metis exec_trace_invariant)
apply (frule_tac i=i in exec_trace_step_cases)
apply (rule_tac x="the (trace_bottom_addr tr i)" in exI)
apply (clarsimp simp: all_exec_graph_step_cases)
apply (safe, simp_all add: reachable_step_def)[1]
apply (auto simp: exec_graph_invariant_def trace_addr_def
get_state_function_call_def trace_bottom_addr_def
split: graph_function.split_asm next_node.split option.splits node.splits
next_node.split)
done
lemma init_pc:
"tr \<in> exec_trace Gamma fn \<Longrightarrow> Gamma fn = Some f
\<Longrightarrow> tr 0 = Some [init_state fn f xs]
\<Longrightarrow> \<forall>n. 0 \<in> restrs n
\<Longrightarrow> wf_graph_function f ilen olen
\<Longrightarrow> pc' (entry_point f) restrs tr"
apply (clarsimp simp: pc_def visit_def)
apply (rule_tac x=0 in exI)
apply (simp add: trace_addr_SomeI init_state_def restrs_condition_def
split: graph_function.split)
done
lemma eqs_eq_conj_len_assert:
"(eqs xs lens vs) = ((eqs xs lens $ vs)
\<and> (\<forall>addr. lens addr \<noteq> None \<longrightarrow> length (vs addr) = the (lens addr)))"
by (auto simp: eqs_def)
lemma length_acc_vars[simp]:
"length (acc_vars vs st) = length vs"
by (simp add: acc_vars_def)
lemma restrs_eventually_init:
"tr \<in> exec_trace Gamma fn
\<Longrightarrow> Gamma fn = Some f
\<Longrightarrow> converse (reachable_step' f) `` {NextNode (entry_point f)} \<subseteq> set []
\<Longrightarrow> rs = (restrs_list [(entry_point f, [1])])
\<Longrightarrow> restrs_eventually_condition tr rs"
apply (clarsimp simp: restrs_eventually_condition_def)
apply (rule_tac x=1 in exI)
apply (rule conjI)
apply (clarsimp simp: exec_trace_def nat_trace_rel_def)
apply (clarsimp simp: restrs_condition_def restrs_list_def)
apply (subst arg_cong[where f=card and y="{0}"], simp_all)
apply (safe, simp_all)
apply (case_tac x, simp_all)
apply (drule(2) exec_trace_reachable_step, simp)
apply auto[1]
apply (clarsimp simp: exec_trace_def trace_addr_def)
done
lemma graph_function_restr_trace_refine:
assumes wfs: "wf_graph_function f1 il1 ol1" "wf_graph_function f2 il2 ol2"
and inps: "converse (reachable_step' f1) `` {NextNode (entry_point f1)} \<subseteq> set []"
"converse (reachable_step' f2) `` {NextNode (entry_point f2)} \<subseteq> set []"
and gfs: "Gamma1 fn1 = Some f1" "Gamma2 fn2 = Some f2"
notes wf_facts = wf_graph_function_init_extract
wf_graph_function_init_extract
wf_graph_function_length_function_inputs
shows "graph_function_refine prec Gamma1 fn1 Gamma2 fn2
(eqs ieqs (Some o ils)) orel
= (\<forall>tr tr'. related_pair (False, NextNode (entry_point f1), restrs_list [])
(acc_vars (function_inputs f1))
(True, NextNode (entry_point f2), restrs_list [])
(acc_vars (function_inputs f2))
(eqs ieqs (Some o ils) o tuple_switch) (tr, tr')
\<and> exec_double_trace Gamma1 fn1 Gamma2 fn2 tr tr'
\<longrightarrow> restr_trace_refine prec Gamma1 fn1 Gamma2 fn2
(restrs_list [(entry_point f1, [1])]) (restrs_list [(entry_point f2, [1])])
orel tr tr')"
using wfs
apply (simp add: graph_function_refine_trace2 gfs)
apply (simp add: restr_trace_refine_def gfs
restrs_eventually_init
restrs_list_def split_def related_pair_def double_visit_def
wf_graph_function_init_extract)
apply (rule arg_cong[where f=All] ext)+
apply (auto simp: wf_graph_function_init_extract
wf_graph_function_length_function_inputs gfs
restrs_list_def
cong: if_cong
elim!: restrs_eventually_init[where Gamma=Gamma1, OF _ _ inps(1)]
restrs_eventually_init[where Gamma=Gamma2, OF _ _ inps(2)];
metis)
done
lemma restrs_list_Int:
"restrs_list rs = (\<lambda>j. \<Inter>(n, xsf) \<in> set rs. (if j = n then set xsf else UNIV))"
apply (induct rs rule: rev_induct)
apply (simp add: restrs_list_def)
apply (clarsimp simp add: restrs_list_def fun_eq_iff)
done
lemma restrs_list_Int2:
"restrs_list rs = (\<lambda>j. \<Inter>(_, xsf) \<in> {x \<in> set rs. fst x = j}. set xsf)"
by (auto simp add: restrs_list_Int fun_eq_iff set_eq_iff)
lemma restrs_list_Cons:
"restrs_list ((n, xs) # rs) = (\<lambda>j.
if j = n then set xs \<inter> restrs_list rs j else restrs_list rs j)"
by (simp add: restrs_list_Int fun_eq_iff)
lemma restrs_eventually_Cons:
"restrs_eventually_condition tr ((K UNIV) (n := set xs))
\<Longrightarrow> restrs_eventually_condition tr (restrs_list rs)
\<Longrightarrow> restrs_eventually_condition tr (restrs_list ((n, xs) # rs))"
apply (clarsimp simp: restrs_eventually_condition_def)
apply (rule_tac x="max na naa" in exI)
apply (simp add: restrs_condition_def restrs_list_Cons split: if_split_asm)
apply (simp add: max_def)
done
lemma ex_greatest_nat_le:
"P j \<Longrightarrow> j \<le> n \<Longrightarrow> \<exists>k \<le> n. (\<forall>i \<in> {Suc k .. n}. \<not> P i) \<and> P k"
apply (cases "P n")
apply (rule_tac x=n in exI, simp)
apply (cut_tac P="\<lambda>j. P (n - j)" and n = "n - j" in ex_least_nat_le, simp+)
apply (clarify, rule_tac x="n - k" in exI)
apply clarsimp
apply (drule_tac x="n - i" in spec, simp)
done
lemma neq_counting_le:
assumes neq: "\<forall>i. f i \<noteq> bound"
assumes mono: "\<forall>i. f (Suc i) \<le> Suc (f i)"
shows "f 0 < bound \<Longrightarrow> f i < bound"
proof (induct i)
case 0
show ?case using 0 by simp
next
case (Suc n)
show ?case using neq[rule_format, of n] neq[rule_format, of "Suc n"]
mono[rule_format, of n] Suc
by simp
qed
lemma restrs_eventually_upt:
"\<exists>m. card {i. i < m \<and> trace_addr tr i = Some (NextNode n)} = i \<and> tr m \<noteq> None
\<Longrightarrow> \<forall>m. card {i. i < m \<and> trace_addr tr i = Some (NextNode n)} \<noteq> j
\<Longrightarrow> restrs_eventually_condition tr ((\<lambda>_. UNIV) (n := {i ..< j}))"
apply (clarsimp simp: restrs_eventually_condition_def restrs_condition_def)
apply (rule exI, rule conjI, erule exI)
apply (clarsimp simp: Suc_le_eq)
apply (intro conjI impI allI)
apply (rule card_mono)
apply simp
apply clarsimp
apply (frule neq_counting_le, simp_all)
apply (clarsimp simp: Collect_less_Suc)
apply (drule_tac x=0 in spec, simp)
done
lemma set_fill_in_below:
"set (fill_in_below xs) = {k. \<exists>l. l : set xs \<and> k \<le> l}"
apply (induct xs rule: rev_induct, simp_all add: fill_in_below_def)
apply (simp add: set_eq_iff less_max_iff_disj less_Suc_eq_le)
apply auto
done
lemma restrs_visit_abstract:
assumes dist: "distinct (map fst rs)"
shows "restrs_list (restrs_visit rs nn gf)
= (\<lambda>j. if (nn, NextNode j) \<in> rtrancl (reachable_step' gf)
then {k. \<exists>l. l \<in> restrs_list rs j \<and> k \<le> l}
else restrs_list rs j)"
proof -
have P: "\<And>j. {x \<in> set (restrs_visit rs nn gf). fst x = j}
= (if (nn, NextNode j) \<in> rtrancl (reachable_step' gf)
then (\<lambda>(k, xsf). (k, fill_in_below xsf))
` {x \<in> set rs. fst x = j}
else {x \<in> set rs. fst x = j})"
by (force simp: restrs_visit_def elim: image_eqI[rotated])
show ?thesis
apply (rule ext, clarsimp simp: restrs_list_Int2 split_def P)
apply (cut_tac dist[THEN set_map_of_compr])
apply (case_tac "\<exists>b. (j, b) \<in> set rs")
apply (auto simp: set_fill_in_below)
done
qed
lemma last_upd_stack:
"List.last (upd_stack nn stf xs)
= (if length xs = 1 then (nn, stf (fst (snd (hd xs))), snd (snd (hd xs)))
else List.last xs)"
by (cases xs, simp_all)
lemma not_reachable_visits_same:
"tr \<in> exec_trace Gamma fn \<Longrightarrow> Gamma fn = Some gf
\<Longrightarrow> trace_addr tr i = Some n
\<Longrightarrow> (n, m) \<notin> rtrancl (reachable_step' gf)
\<Longrightarrow> wf_graph_function gf ilen olen
\<Longrightarrow> j > i
\<Longrightarrow> {k. k < j \<and> trace_addr tr k = Some m} = {k. k < i \<and> trace_addr tr k = Some m}"
using reachable_trace_eq[THEN subsetD[OF rtrancl_mono, OF Int_lower1]]
apply (safe, simp_all)
apply (rule ccontr, clarsimp simp: linorder_not_less)
done
lemma not_reachable_visits_same_symm:
"tr \<in> exec_trace Gamma fn \<Longrightarrow> Gamma fn = Some gf
\<Longrightarrow> trace_addr tr i = Some n
\<Longrightarrow> trace_addr tr j = Some n
\<Longrightarrow> (n, m) \<notin> rtrancl (reachable_step' gf)
\<Longrightarrow> wf_graph_function gf ilen olen
\<Longrightarrow> {k. k < j \<and> trace_addr tr k = Some m} = {k. k < i \<and> trace_addr tr k = Some m}"
using linorder_neq_iff[where x=i and y=j] not_reachable_visits_same
by blast
lemma restrs_eventually_at_visit:
"restrs_eventually_condition tr (restrs_list rs)
\<Longrightarrow> trace_addr tr i = Some nn
\<Longrightarrow> tr \<in> exec_trace Gamma fn \<Longrightarrow> Gamma fn = Some gf
\<Longrightarrow> distinct (map fst rs)
\<Longrightarrow> wf_graph_function gf ilen olen
\<Longrightarrow> restrs_condition tr (restrs_list (restrs_visit rs nn gf)) i"
apply (clarsimp simp: restrs_eventually_condition_def
restrs_condition_def restrs_visit_abstract)
apply (case_tac "i \<ge> n")
apply (drule spec, drule(1) mp)
apply auto[1]
apply (simp add: linorder_not_le)
apply (drule spec, drule mp, rule order_refl)
apply safe
apply (fastforce intro: card_mono)
apply (simp add: not_reachable_visits_same[symmetric])
done
lemma fold_double_trace_imp:
"fold double_trace_imp hyps hyp trs
= ((\<forall>h \<in> set hyps. h trs) \<longrightarrow> hyp trs)"
apply (induct hyps arbitrary: hyp, simp_all)
apply (auto simp add: double_trace_imp_def)
done
lemma exec_trace_addr_Suc:
"tr \<in> exec_trace Gamma f \<Longrightarrow> trace_addr tr n = Some (NextNode m) \<Longrightarrow> tr (Suc n) \<noteq> None"
apply (drule_tac i=n in exec_trace_step_cases)
apply (auto dest!: trace_addr_SomeD)
done
lemma num_visits_equals_j_first:
"card {i. i < m \<and> trace_addr tr i = Some n} = j
\<Longrightarrow> j \<noteq> 0
\<Longrightarrow> \<exists>m'. trace_addr tr m' = Some n \<and> card {i. i < m' \<and> trace_addr tr i = Some n} = j - 1"
apply (frule_tac P="\<lambda>m. card {i. i < m \<and> trace_addr tr i = Some n} = j" in ex_least_nat_le)
apply simp
apply clarify
apply (rule_tac x="k - 1" in exI)
apply (case_tac k)
apply (simp del: Collect_empty_eq)
apply (simp add: Collect_less_Suc split: if_split_asm)
done
lemma ex_least_nat:
"\<exists>n. P n \<Longrightarrow> \<exists>n :: nat. P n \<and> (\<forall>i < n. \<not> P i)"
apply clarsimp
apply (case_tac "n = 0")
apply fastforce
apply (cut_tac P="\<lambda>i. i \<noteq> 0 \<and> P i" in ex_least_nat_le, auto)
done
lemma visit_eqs:
"(visit trace nn restrs = None)
= (\<forall>i. trace_addr trace i = Some nn \<longrightarrow> \<not> restrs_condition trace restrs i)"
"(visit trace nn restrs = Some st)
= (\<exists>i. restrs_condition trace restrs i \<and> trace_addr trace i = Some nn
\<and> st = fst (snd (hd (the (trace i))))
\<and> (\<forall>j < i. restrs_condition trace restrs j \<longrightarrow> trace_addr trace j \<noteq> Some nn))"
apply (simp_all add: visit_def)
apply auto[1]
apply (safe elim!: exE[OF ex_least_nat] del: exE, simp_all)
apply (rule exI, rule conjI, assumption, simp)
apply (rule arg_cong[OF Least_equality], simp)
apply (blast intro: linorder_not_le[THEN iffD1])
apply (rule arg_cong[OF Least_equality], simp)
apply (blast intro: linorder_not_le[THEN iffD1])
done
lemmas visit_eqs' = visit_eqs(1) arg_cong[where f=Not, OF visit_eqs(1), simplified]
theorem restr_trace_refine_Restr1:
"j \<noteq> 0
\<Longrightarrow> distinct (map fst rs1)
\<Longrightarrow> wf_graph_function f1 ilen olen \<Longrightarrow> Gamma1 fn1 = Some f1
\<Longrightarrow> i \<noteq> 0 \<longrightarrow>
pc' n (restrs_list ((n, [i - 1]) # (restrs_visit rs1 (NextNode n) f1))) tr
\<Longrightarrow> \<not> pc' n (restrs_list ((n, [j - 1]) # (restrs_visit rs1 (NextNode n) f1))) tr
\<Longrightarrow> restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 (restrs_list ((n, [i ..< j]) # rs1)) rs2 orel tr tr'
\<Longrightarrow> restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 (restrs_list rs1) rs2 orel tr tr'"
apply (clarsimp simp: restr_trace_refine_def)
apply (subst(asm) restrs_eventually_Cons, simp_all add: K_def)
prefer 2
apply auto[1]
apply (rule restrs_eventually_upt)
apply (clarsimp simp: fold_double_trace_imp double_trace_imp_def pc_def)
apply (case_tac "i = 0")
apply (rule_tac x=0 in exI)
apply (clarsimp simp: dest!: exec_trace_init)
apply simp
apply (clarsimp simp: visit_eqs)
apply (simp(no_asm_use) add: restrs_list_Cons restrs_condition_def)
apply (drule_tac x=n in spec)+
apply simp
apply (rule_tac x="Suc ia" in exI, simp add: Collect_less_Suc)
apply (rule exec_trace_addr_Suc[simplified], assumption)
apply simp
apply (clarsimp simp: fold_double_trace_imp double_trace_imp_def pc_def)
apply (thin_tac "0 < i \<longrightarrow> q" for q)
apply (clarsimp simp: visit_eqs)
apply (frule num_visits_equals_j_first[OF refl, simplified neq0_conv])
apply clarsimp
apply (drule(5) restrs_eventually_at_visit)
apply (drule_tac x=m' in spec, clarsimp)
apply (clarsimp simp: restrs_list_Cons restrs_condition_def split: if_split_asm)
done
theorem restr_trace_refine_Restr2:
"j \<noteq> 0
\<Longrightarrow> distinct (map fst rs2)
\<Longrightarrow> wf_graph_function f2 ilen olen \<Longrightarrow> Gamma2 fn2 = Some f2
\<Longrightarrow> i \<noteq> 0 \<longrightarrow> pc' n (restrs_list ((n, [i - 1]) # (restrs_visit rs2 (NextNode n) f2))) tr'
\<Longrightarrow> \<not> pc' n (restrs_list ((n, [j - 1]) # (restrs_visit rs2 (NextNode n) f2))) tr'
\<Longrightarrow> restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 (restrs_list ((n, [i ..< j]) # rs2)) orel tr tr'
\<Longrightarrow> restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 (restrs_list rs2) orel tr tr'"
apply (clarsimp simp: restr_trace_refine_def)
apply (subst(asm) restrs_eventually_Cons, simp_all add: K_def)
prefer 2
apply auto[1]
apply (rule restrs_eventually_upt)
apply (clarsimp simp: fold_double_trace_imp double_trace_imp_def pc_def)
apply (case_tac "i = 0")
apply (rule_tac x=0 in exI)
apply (clarsimp dest!: exec_trace_init)
apply simp
apply (clarsimp simp: visit_eqs restrs_list_Cons restrs_condition_def)
apply (drule_tac x=n in spec)+
apply simp
apply (rule_tac x="Suc ia" in exI, simp add: Collect_less_Suc)
apply (rule exec_trace_addr_Suc[simplified], assumption)
apply simp
apply (clarsimp simp: fold_double_trace_imp double_trace_imp_def pc_def)
apply (thin_tac "0 < i \<longrightarrow> q" for q)
apply (clarsimp simp: visit_eqs)
apply (frule num_visits_equals_j_first[OF refl, simplified neq0_conv])
apply clarsimp
apply (drule(5) restrs_eventually_at_visit)
apply (drule_tac x=m' in spec, clarsimp)
apply (clarsimp simp: restrs_list_Cons restrs_condition_def split: if_split_asm)
done