-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathExampleSpecIRQ_SI.thy
918 lines (793 loc) · 38.2 KB
/
ExampleSpecIRQ_SI.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
(*
* Copyright 2014, NICTA
*
* This software may be distributed and modified according to the terms of
* the GNU General Public License version 2. Note that NO WARRANTY is provided.
* See "LICENSE_GPLv2.txt" for details.
*
* @TAG(NICTA_GPL)
*)
(*
* This file contains an example spec, and shows it obeys the well_formed constraints
*
* This file contains the invariants of the user-space initialiser,
* the well-formed conditions of the input capDL specification,
* and the predicates describing the initial state.
*)
theory ExampleSpecIRQ_SI
imports WellFormed_SI
begin
(****************************************************
* Definitions of all the objects and capabilities. *
***************************************************)
definition small_section_size :: nat
where
"small_section_size = 20"
definition "guard = 0"
definition "guard_size = 1"
definition "cnode_a1_size = 7"
definition "cnode_a2_size = 7"
definition "cnode_b_size = 8"
definition "cnode_extra_size = 2"
lemmas constants [simp] = guard_def guard_size_def cnode_a1_size_def cnode_a2_size_def
cnode_b_size_def cnode_extra_size_def
definition "tcb_a_id = 9"
definition "tcb_b_id = 10"
definition "cnode_a1_id = 6"
definition "cnode_a2_id = 7"
definition "cnode_b_id = 5"
definition "cnode_extra_id = 11"
definition "ep_id = 12"
definition "aep_id = 13"
definition "pd_a_id = 1"
definition "pt_a_id = 8"
definition "pd_b_id = 2"
definition "frame_a1_id = 3"
definition "frame_a2_id = 4"
definition "frame_b_id = 0"
lemmas ids [simp] = tcb_a_id_def tcb_b_id_def cnode_a1_id_def
cnode_a2_id_def cnode_b_id_def cnode_extra_id_def
ep_id_def aep_id_def pd_a_id_def pt_a_id_def pd_b_id_def
frame_a1_id_def frame_a2_id_def frame_b_id_def
definition
"tcb_a =
\<lparr>cdl_tcb_caps = [tcb_cspace_slot \<mapsto> CNodeCap cnode_a1_id guard guard_size cnode_a1_size,
tcb_vspace_slot \<mapsto> PageDirectoryCap pd_a_id Real None,
tcb_replycap_slot \<mapsto> NullCap,
tcb_caller_slot \<mapsto> NullCap,
tcb_ipcbuffer_slot \<mapsto> FrameCap frame_a1_id {AllowRead, AllowWrite} small_frame_size Real None,
tcb_pending_op_slot \<mapsto> NullCap],
cdl_tcb_fault_endpoint = 0,
cdl_tcb_intent = \<lparr>cdl_intent_op = None, cdl_intent_error = False, cdl_intent_cap = 0,
cdl_intent_extras = [], cdl_intent_recv_slot = None\<rparr>,
cdl_tcb_has_fault = False,
cdl_tcb_domain = minBound\<rparr>"
definition
"tcb_b =
\<lparr>cdl_tcb_caps = [tcb_cspace_slot \<mapsto> CNodeCap cnode_b_id guard guard_size cnode_b_size,
tcb_vspace_slot \<mapsto> PageDirectoryCap pd_b_id Real None,
tcb_replycap_slot \<mapsto> NullCap,
tcb_caller_slot \<mapsto> NullCap,
tcb_ipcbuffer_slot \<mapsto> FrameCap frame_b_id {AllowRead, AllowWrite} small_section_size Real None,
tcb_pending_op_slot \<mapsto> NullCap],
cdl_tcb_fault_endpoint = 0,
cdl_tcb_intent = \<lparr>cdl_intent_op = None, cdl_intent_error = False, cdl_intent_cap = 0,
cdl_intent_extras = [], cdl_intent_recv_slot = None\<rparr>,
cdl_tcb_has_fault = False,
cdl_tcb_domain = minBound\<rparr>"
definition
"new_cap_map sz caps \<equiv> empty_cap_map sz ++ caps"
definition
"new_cnode sz caps \<equiv> \<lparr>cdl_cnode_caps = new_cap_map sz caps,
cdl_cnode_size_bits = sz\<rparr>"
definition
"cnode_a1 \<equiv>
new_cnode cnode_a1_size
[0 \<mapsto> TcbCap tcb_a_id,
1 \<mapsto> CNodeCap cnode_a2_id guard guard_size cnode_a2_size]"
definition
"cnode_a2 \<equiv>
new_cnode cnode_a2_size
[0 \<mapsto> EndpointCap ep_id 0 {Write},
2 \<mapsto> CNodeCap cnode_a1_id guard guard_size cnode_a1_size,
3 \<mapsto> PageDirectoryCap pd_a_id Real None,
4 \<mapsto> PageTableCap pt_a_id Real None,
8 \<mapsto> FrameCap frame_a1_id {AllowRead, AllowWrite} small_frame_size Real None,
10 \<mapsto> AsyncEndpointCap aep_id 0 {Read},
11 \<mapsto> FrameCap frame_a2_id {AllowRead, AllowWrite} small_frame_size Real None,
12 \<mapsto> IrqHandlerCap 4]"
definition
"cnode_b \<equiv>
new_cnode cnode_b_size
[0 \<mapsto> TcbCap tcb_b_id,
2 \<mapsto> CNodeCap cnode_b_id guard guard_size cnode_b_size,
4 \<mapsto> EndpointCap ep_id 0 {Read},
7 \<mapsto> PageDirectoryCap pd_b_id Real None,
8 \<mapsto> FrameCap frame_b_id {AllowRead, AllowWrite} small_section_size Real None,
254 \<mapsto> IrqHandlerCap 254]"
definition
"cnode_extra \<equiv>
new_cnode cnode_extra_size
[0 \<mapsto> CNodeCap cnode_extra_id guard guard_size cnode_extra_size,
1 \<mapsto> EndpointCap ep_id 0 UNIV,
2 \<mapsto> AsyncEndpointCap aep_id 0 {Read, Write}]"
definition
"pd_a \<equiv> \<lparr>cdl_page_directory_caps = new_cap_map pd_size [0 \<mapsto> PageTableCap pt_a_id Fake None]\<rparr>"
definition
"pd_b \<equiv> \<lparr>cdl_page_directory_caps = new_cap_map pd_size
[2 \<mapsto> FrameCap frame_b_id {AllowRead, AllowWrite} small_section_size Fake None]\<rparr>"
definition
"pt_a \<equiv> \<lparr>cdl_page_table_caps = new_cap_map pt_size
[0 \<mapsto> FrameCap frame_a1_id {AllowRead, AllowWrite} small_frame_size Fake None,
255 \<mapsto> FrameCap frame_a2_id {AllowRead, AllowWrite} small_frame_size Fake None]\<rparr>"
definition
"empty_frame \<equiv> \<lparr>cdl_frame_size_bits = small_frame_size\<rparr>"
definition
"empty_section \<equiv> \<lparr>cdl_frame_size_bits = small_section_size\<rparr>"
definition
example_irq_node :: "cdl_irq \<Rightarrow> cdl_object_id"
where
"example_irq_node = (\<lambda>irq. ucast irq + 0x100)"
definition
"new_irq_node obj_id \<equiv>
CNode \<lparr> cdl_cnode_caps = (\<lambda>slot. if slot = 0
then Some (AsyncEndpointCap obj_id 0 {Read, Write})
else None),
cdl_cnode_size_bits = 0 \<rparr>"
definition
irq_objects :: cdl_heap
where
"irq_objects \<equiv>
\<lambda>obj_id. if obj_id = 0x104 then Some (new_irq_node aep_id)
else if 0x100 \<le> obj_id \<and> obj_id < 0x200 then (Some empty_irq_node)
else None"
lemma
"irq_objects =
(\<lambda>obj_id. if 0x100 \<le> obj_id \<and> obj_id < 0x200
then (Some empty_irq_node)
else None) ++
(\<lambda>obj_id. if obj_id = 0x104 then Some (new_irq_node aep_id)
else None)"
apply (rule ext)
apply (clarsimp simp: irq_objects_def map_add_def)
done
definition
"example_spec =
\<lparr>cdl_arch = ARM11,
cdl_objects = [tcb_a_id \<mapsto> Tcb tcb_a,
tcb_b_id \<mapsto> Tcb tcb_b,
pd_a_id \<mapsto> PageDirectory pd_a,
pt_a_id \<mapsto> PageTable pt_a,
pd_b_id \<mapsto> PageDirectory pd_b,
cnode_a1_id \<mapsto> CNode cnode_a1,
cnode_a2_id \<mapsto> CNode cnode_a2,
cnode_b_id \<mapsto> CNode cnode_b,
cnode_extra_id \<mapsto> CNode cnode_extra,
frame_a1_id \<mapsto> Frame empty_frame,
frame_a2_id \<mapsto> Frame empty_frame,
frame_b_id \<mapsto> Frame empty_section,
ep_id \<mapsto> Endpoint,
aep_id \<mapsto> AsyncEndpoint,
0x104 \<mapsto> new_irq_node aep_id,
0x1FE \<mapsto> empty_irq_node],
cdl_cdt = [(cnode_a2_id, 0) \<mapsto> (cnode_extra_id, 1),
(cnode_b_id, 4) \<mapsto> (cnode_extra_id, 1),
(cnode_a2_id, 10) \<mapsto> (cnode_extra_id, 2)], (* All caps are orig caps,
except the endpoint and async endpoint. *)
cdl_current_thread = undefined,
cdl_irq_node = example_irq_node,
cdl_asid_table = undefined,
cdl_current_domain = minBound\<rparr>"
declare cap_object_simps [simp]
lemmas cnode_defs = cnode_a1_def cnode_a2_def cnode_b_def cnode_extra_def
lemmas obj_defs = cnode_defs tcb_a_def tcb_b_def pd_a_def pd_b_def pt_a_def
lemmas example_spec_def_expanded = example_spec_def
[unfolded obj_defs tcb_slot_defs tcb_pending_op_slot_def
new_cnode_def new_cap_map_def empty_cap_map_def]
(*************************
* Helper lemmas. *
*************************)
lemma cap_has_object_IrqHandlerCap [simp]:
"\<not>cap_has_object (IrqHandlerCap irq)"
by (clarsimp simp: cap_has_object_def)+
lemma badge_bits_2p [simp]:
"(0::word32) < 2 ^ badge_bits"
by (clarsimp simp: p2_gt_0 badge_bits_def)
lemma cdl_cnode_size_bits_new_cnode [simp]:
"cdl_cnode_size_bits (new_cnode sz caps) = sz"
by (clarsimp simp: new_cnode_def)
lemma cnode_cap_size_simps [simp]:
"cnode_cap_size (CNodeCap a b c sz) = sz"
by (clarsimp simp: cnode_cap_size_def)
lemma object_size_bits_new_cnode [simp]:
"object_size_bits (CNode (new_cnode sz caps)) = sz"
by (clarsimp simp: object_size_bits_def)
lemma object_slots_Endpoint [simp]:
"object_slots Endpoint = empty"
by (simp add: object_slots_def)
lemma cdl_frame_size_bits_empty_frame [simp]:
"cdl_frame_size_bits empty_frame = small_frame_size"
by (simp add: empty_frame_def)
lemma cdl_frame_size_bits_empty_section [simp]:
"cdl_frame_size_bits empty_section = small_section_size"
by (simp add: empty_section_def)
lemma object_slots_empty_objects [simp]:
"object_slots (Frame f) slot = None"
by (clarsimp simp: object_slots_def)+
lemma is_fake_pt_cap_simps:
"\<not> is_fake_pt_cap (PageTableCap obj_id Real asid)"
"is_fake_pt_cap (PageTableCap obj_id Fake asid)"
by (clarsimp simp: is_fake_pt_cap_def)+
lemma frame_cap_not_cnode:
"\<not>is_cnode_cap (FrameCap a b c d e)"
by (clarsimp simp: cap_type_def)
lemma empty_cap_map_NullCap [simp]:
"empty_cap_map sz slot = Some cap \<Longrightarrow> cap = NullCap"
by (clarsimp simp: empty_cap_map_def split: split_if_asm)
lemma new_cap_map_empty_NullCap [simp]:
"new_cap_map sz empty slot = Some cap \<Longrightarrow> cap = NullCap"
by (clarsimp simp: new_cap_map_def)
lemma new_cap_map_slot:
"\<lbrakk>new_cap_map sz caps slot = Some cap; cap \<noteq> NullCap\<rbrakk> \<Longrightarrow> caps slot = Some cap"
by (clarsimp simp: new_cap_map_def empty_cap_map_def split: option.splits split_if_asm)
lemma cdl_cnode_caps_new_cnode:
"\<lbrakk>cdl_cnode_caps (new_cnode sz caps) slot = Some cap; cap \<noteq> NullCap\<rbrakk> \<Longrightarrow> caps slot = Some cap"
by (clarsimp simp: new_cnode_def, erule (1) new_cap_map_slot)
lemma new_cap_map_caps_D:
"new_cap_map sz caps slot = Some cap \<Longrightarrow> caps slot = Some cap \<or> cap = NullCap"
by (clarsimp simp: new_cap_map_def)
lemma cdl_cnode_caps_new_cnode_D:
"\<lbrakk>cdl_cnode_caps (new_cnode sz caps) slot = Some cap\<rbrakk>
\<Longrightarrow> caps slot = Some cap \<or> cap = NullCap"
by (clarsimp simp: new_cnode_def, erule (1) new_cap_map_slot)
lemma object_slots_new_cnode_D:
"object_slots (CNode (new_cnode sz caps)) slot = Some cap
\<Longrightarrow> caps slot = Some cap \<or> cap = NullCap"
by (clarsimp simp: object_slots_def dest!: cdl_cnode_caps_new_cnode_D)
lemma object_slots_new_cnode_cap_has_object [dest!]:
"\<lbrakk>object_slots (CNode (new_cnode sz caps)) slot = Some cap; cap_has_object cap\<rbrakk>
\<Longrightarrow> caps slot = Some cap"
by (clarsimp simp: object_slots_def dest!: cdl_cnode_caps_new_cnode_D)
lemma cdl_cnode_caps_empty_cnode [dest!]:
"cdl_cnode_caps (empty_cnode sz) slot = Some cap \<Longrightarrow> cap = NullCap"
by (clarsimp simp: empty_cnode_def)
lemma cdl_cnode_caps_new_cnode_cnode_cap:
"\<lbrakk>cdl_cnode_caps (new_cnode sz caps) slot = Some cap; is_cnode_cap cap\<rbrakk>
\<Longrightarrow> caps slot = Some cap"
by (erule cdl_cnode_caps_new_cnode, clarsimp)
lemma object_slots_new_cnode_cnode_cap:
"\<lbrakk>object_slots (CNode (new_cnode sz caps)) slot = Some cap; is_cnode_cap cap\<rbrakk>
\<Longrightarrow> caps slot = Some cap"
by (clarsimp simp: object_slots_def, erule cdl_cnode_caps_new_cnode, clarsimp)
lemma object_slots_empty_irq_node [simp, dest!]:
"object_slots empty_irq_node slot = Some cap \<Longrightarrow> cap = NullCap"
by (clarsimp simp: empty_irq_node_def object_slots_empty_cnode)
lemma tcb_domain_simp [simp]:
"tcb_domain (Tcb \<lparr>cdl_tcb_caps = caps,
cdl_tcb_fault_endpoint = 0,
cdl_tcb_intent = intent,
cdl_tcb_has_fault = fault,
cdl_tcb_domain = domain\<rparr>) = domain"
by (simp add: tcb_domain_def)
lemma cdl_irq_node_example_spec [simp]:
"cdl_irq_node example_spec = example_irq_node"
by (clarsimp simp: example_spec_def)
lemma range_example_irq_node_helper:
"range example_irq_node = (\<lambda>irq. irq + 0x100) ` range (ucast :: 8 word \<Rightarrow> 32 word)"
by (auto simp: example_irq_node_def image_def)
lemma range_example_irq_node:
"range example_irq_node = {x. 0x100 \<le> x \<and> x < 0x200}"
apply (clarsimp simp: range_example_irq_node_helper ucast_range_less)
apply (clarsimp simp: image_def)
apply rule
apply (clarsimp simp: word_le_nat_alt word_less_nat_alt unat_plus_if')
apply clarsimp
apply (rule_tac x="x - 0x100" in exI)
apply unat_arith
done
lemma irq_cnodes_example_spec:
"irq_cnodes example_spec = {x. 0x100 \<le> x \<and> x < 0x200}"
by (simp add: irq_cnodes_def range_example_irq_node)
(*************************
* End of helper lemmas. *
*************************)
(************************
Helpers on the specific state.
***************************)
lemma onehundred_not_le_one:
"\<not>(0x100 \<le> (1::32 word))"
by unat_arith
lemma object_type_simps [simp]:
"object_type (Tcb t) = TcbType"
"object_type (CNode c) = CNodeType"
"object_type (Endpoint) = EndpointType"
"object_type (AsyncEndpoint) = AsyncEndpointType"
"object_type (PageDirectory pd) = PageDirectoryType"
"object_type (PageTable pt) = PageTableType"
"object_type (Frame f) = FrameType (cdl_frame_size_bits f)"
"object_type empty_irq_node = CNodeType"
by (clarsimp simp: object_type_def empty_irq_node_def)+
lemma cap_type'_simps [simp]:
"cap_type' (TcbCap obj_id) = Some TcbType"
"cap_type' (CNodeCap obj_id g gs sz) = Some CNodeType"
"cap_type' (EndpointCap obj_id b rights) = Some EndpointType"
"cap_type' (AsyncEndpointCap obj_id b rights) = Some AsyncEndpointType"
"cap_type' (PageDirectoryCap obj_id r asid) = Some PageDirectoryType"
"cap_type' (PageTableCap obj_id r asid) = Some PageTableType"
"cap_type' (FrameCap obj_id rights sz r asid) = Some (FrameType sz)"
"cap_type' (IrqHandlerCap irq) = Some CNodeType"
by (clarsimp simp: cap_type'_def cap_type_def)+
lemma is_irqhandler_cap_simps [simp]:
"\<not> is_irqhandler_cap (TcbCap obj_id)"
"\<not> is_irqhandler_cap (CNodeCap obj_id g gs sz)"
"\<not> is_irqhandler_cap (EndpointCap obj_id b rights)"
"\<not> is_irqhandler_cap (AsyncEndpointCap obj_id b rights)"
"\<not> is_irqhandler_cap (PageDirectoryCap obj_id r asid)"
"\<not> is_irqhandler_cap (PageTableCap obj_id r asid)"
"\<not> is_irqhandler_cap (FrameCap obj_id rights sz r asid)"
by (clarsimp simp: is_irqhandler_cap_def)+
lemma cap_irq_simp [simp]:
"cap_irq (IrqHandlerCap irq) = irq"
by (simp add: cap_irq_def)
lemma example_irq_node_simps [simp]:
"example_irq_node 4 = 0x104"
"example_irq_node 0xFE = 0x1FE"
by (simp add: example_irq_node_def)+
lemma irq_objects_simps [simp]:
"irq_objects 0 = None"
"irq_objects 1 = None"
"irq_objects 2 = None"
"irq_objects 3 = None"
"irq_objects 4 = None"
"irq_objects 5 = None"
"irq_objects 6 = None"
"irq_objects 7 = None"
"irq_objects 8 = None"
"irq_objects 9 = None"
"irq_objects 0xA = None"
"irq_objects 0xB = None"
"irq_objects 0xC = None"
"irq_objects 0xD = None"
by (clarsimp simp: irq_objects_def onehundred_not_le_one)+
lemma opt_cap_example_spec [simp]:
"opt_cap (4, slot) example_spec = object_slots (Frame empty_frame) slot"
"opt_cap (5, slot) example_spec = object_slots (CNode cnode_b) slot"
"opt_cap (6, slot) example_spec = object_slots (CNode cnode_a1) slot"
"opt_cap (7, slot) example_spec = object_slots (CNode cnode_a2) slot"
"opt_cap (0xB, slot) example_spec = object_slots (CNode cnode_extra) slot"
by (auto simp: example_spec_def opt_cap_def slots_of_def opt_object_def
map_add_def irq_objects_def
split: split_if_asm)
lemma irq_objects_some_object:
"irq_objects obj_id = Some obj \<Longrightarrow>
(obj_id = 0x104 \<and> obj = new_irq_node aep_id) \<or> obj = empty_irq_node"
by (clarsimp simp: irq_objects_def split: split_if_asm)
(*
lemma real_cnode_at_example_spec:
"real_cnode_at obj_id example_spec
\<Longrightarrow> obj_id = cnode_a1_id \<or> obj_id = cnode_a2_id \<or> obj_id = cnode_b_id \<or> obj_id = cnode_extra_id"
apply (clarsimp simp: real_cnode_at_def object_at_def is_cnode_def irq_cnodes_example_spec)
by (auto simp: example_spec_def object_at_def is_cnode_def irq_objects_def
split: split_if_asm cdl_object.splits)
*)
lemma real_cnode_at_example_spec:
"real_cnode_at obj_id example_spec =
(obj_id = cnode_a1_id \<or> obj_id = cnode_a2_id \<or> obj_id = cnode_b_id \<or> obj_id = cnode_extra_id)"
apply (clarsimp simp: real_cnode_at_def object_at_def is_cnode_def irq_cnodes_example_spec)
apply (auto simp: example_spec_def irq_objects_def map_add_def
split: split_if_asm cdl_object.splits)
done
lemma pt_at_example_spec:
"pt_at obj_id example_spec = (obj_id = pt_a_id)"
apply (clarsimp simp: real_cnode_at_def object_at_def is_cnode_def irq_cnodes_example_spec)
apply (auto simp: example_spec_def object_at_def is_pt_def irq_objects_def
new_irq_node_def empty_irq_node_def
split: split_if_asm cdl_object.splits)
done
lemma pd_at_example_spec:
"pd_at obj_id example_spec = (obj_id = pd_a_id \<or> obj_id = pd_b_id)"
apply (clarsimp simp: real_cnode_at_def object_at_def is_cnode_def irq_cnodes_example_spec)
apply (auto simp: example_spec_def object_at_def is_pd_def irq_objects_def
new_irq_node_def empty_irq_node_def
split: split_if_asm cdl_object.splits)
done
lemma slots_of_example_spec_obj_ids:
"\<lbrakk>slots_of obj_id example_spec 0 = Some cap; cap \<noteq> NullCap\<rbrakk>\<Longrightarrow>
((obj_id = tcb_a_id) \<or>
(obj_id = tcb_b_id) \<or>
(obj_id = cnode_a1_id) \<or>
(obj_id = cnode_a2_id) \<or>
(obj_id = cnode_b_id) \<or>
(obj_id = cnode_extra_id) \<or>
(obj_id = pd_a_id) \<or>
(obj_id = pt_a_id) \<or>
(obj_id = pd_b_id) \<or>
(obj_id = 0x104))"
by (clarsimp simp: example_spec_def slots_of_def opt_object_def object_slots_def
split: split_if_asm)
lemma irq_handler_cap_example_spec:
"\<lbrakk>is_irqhandler_cap cap; opt_cap (obj_id, slot) example_spec = Some cap\<rbrakk>
\<Longrightarrow> (obj_id = cnode_a2_id \<and> slot = 12) \<or>
(obj_id = cnode_b_id \<and> slot = 254)"
by (clarsimp simp: example_spec_def opt_cap_def slots_of_def opt_object_def
object_slots_def empty_irq_node_def new_irq_node_def new_cnode_def
obj_defs new_cap_map_def empty_cap_map_def
split: split_if_asm)
lemma irqhandler_cap_at_example_spec:
"irqhandler_cap_at (obj_id, slot) example_spec
= ((obj_id = cnode_a2_id \<and> slot = 12) \<or>
(obj_id = cnode_b_id \<and> slot = 254))"
apply (clarsimp simp: cap_at_def)
apply (rule iffI)
apply clarsimp
apply (drule (1) irq_handler_cap_example_spec)
apply clarsimp
apply (erule disjE)
apply (clarsimp simp: cnode_a2_def object_slots_def new_cnode_def new_cap_map_def)
apply (clarsimp simp: cnode_b_def object_slots_def new_cnode_def new_cap_map_def)
done
lemma cap_at_has_no_parents_in_cdt_example_spec:
"cap_at_has_no_parents_in_cdt (obj_id, slot) example_spec
= ((obj_id \<noteq> cnode_a2_id \<or> slot \<noteq> 0) \<and>
(obj_id \<noteq> cnode_a2_id \<or> slot \<noteq> 10) \<and>
(obj_id \<noteq> cnode_b_id \<or> slot \<noteq> 4))"
by (auto simp: cap_at_has_no_parents_in_cdt_def opt_parent_def example_spec_def)
lemma is_orig_cap_example_spec:
"original_cap_at (obj_id, slot) example_spec
= ((obj_id \<noteq> cnode_a2_id \<or> slot \<noteq> 0) \<and>
(obj_id \<noteq> cnode_a2_id \<or> slot \<noteq> 10) \<and>
(obj_id \<noteq> cnode_b_id \<or> slot \<noteq> 4))"
by (fastforce simp: original_cap_at_def cap_at_has_no_parents_in_cdt_example_spec irqhandler_cap_at_example_spec)
(****************************************
* Proof that the state is well formed. *
****************************************)
lemma well_formed_tcb_a:
"well_formed_tcb example_spec obj_id (Tcb tcb_a)"
by (auto simp: well_formed_tcb_def object_slots_def tcb_a_def tcb_slot_defs
tcb_pending_op_slot_def tcb_has_fault_def is_default_cap_def
default_cap_def cap_type_def irq_cnodes_example_spec)
lemma well_formed_tcb_b:
"well_formed_tcb example_spec obj_id (Tcb tcb_b)"
by (auto simp: well_formed_tcb_def object_slots_def tcb_b_def tcb_slot_defs
tcb_pending_op_slot_def tcb_has_fault_def is_default_cap_def
default_cap_def cap_type_def irq_cnodes_example_spec)
lemma well_formed_orig_caps_unique_example:
"well_formed_orig_caps_unique example_spec"
apply (clarsimp simp: well_formed_orig_caps_unique_def)
apply (clarsimp simp: real_cnode_at_example_spec is_orig_cap_example_spec)
apply (elim disjE, (clarsimp simp: cnode_defs split: split_if_asm)+)
done
lemma well_formed_fake_pt_caps_unique_example:
"well_formed_fake_pt_caps_unique example_spec"
apply (clarsimp simp: well_formed_fake_pt_caps_unique_def
pd_at_example_spec)
apply (fastforce simp: example_spec_def opt_cap_def slots_of_def opt_object_def
object_slots_def is_fake_pt_cap_simps
pd_a_def pd_b_def new_cap_map_def irq_objects_def
split: split_if_asm option.splits)
done
lemma well_formed_orig_cap_tcb [simp]:
"well_formed_orig_cap (TcbCap obj_id)"
by (clarsimp simp: well_formed_orig_cap_def default_cap_def cap_type_def
cap_rights_def ep_related_cap_def)
lemma well_formed_cap_ex [simp]:
"well_formed_cap (CNodeCap a 0 0 2)"
"well_formed_cap (TcbCap 0)"
by (clarsimp simp: well_formed_cap_def guard_bits_def)+
lemma well_formed_cap_example [simp]:
"\<lbrakk>cdl_objects example_spec obj_id = Some obj;
object_slots obj slot = Some cap; cap \<noteq> NullCap\<rbrakk>
\<Longrightarrow> well_formed_cap cap"
apply (clarsimp simp: well_formed_cap_def)
apply (clarsimp simp: well_formed_cap_def example_spec_def
obj_defs new_cap_map_def new_irq_node_def new_cnode_def
object_slots_def empty_cap_map_def guard_bits_def
tcb_slot_defs tcb_pending_op_slot_def vm_read_write_def
split: cdl_cap.splits split_if_asm)
done
lemma well_formed_cdt_example [simp]:
"\<lbrakk>cdl_objects example_spec obj_id = Some obj;
object_slots obj slot = Some cap; cap \<noteq> NullCap\<rbrakk>
\<Longrightarrow> well_formed_cdt example_spec (obj_id, slot) cap"
apply (clarsimp simp: well_formed_cdt_def)
apply (clarsimp simp: real_cnode_at_example_spec)
apply (case_tac "(obj_id = cnode_a2_id \<and> slot = 0) \<or>
(obj_id = cnode_b_id \<and> slot = 4)")
apply (rule_tac x=cnode_extra_id in exI, clarsimp, rule conjI)
apply (fastforce simp: example_spec_def cnode_defs opt_object_def
split: split_if_asm)
apply (rule_tac x=1 in exI)
apply (clarsimp simp: is_orig_cap_example_spec)
apply (clarsimp simp: example_spec_def opt_cap_def slots_of_def opt_object_def
cnode_defs object_slots_def new_cnode_def new_cap_map_def
irq_objects_def map_add_def empty_irq_node_def
split: split_if_asm)
apply (case_tac "(obj_id = cnode_a2_id \<and> slot = 10)")
apply (rule_tac x=cnode_extra_id in exI, clarsimp, rule conjI)
apply (fastforce simp: example_spec_def cnode_defs opt_object_def
split: split_if_asm)
apply (rule_tac x=2 in exI)
apply (clarsimp simp: is_orig_cap_example_spec)
apply (clarsimp simp: example_spec_def opt_cap_def slots_of_def opt_object_def
cnode_defs object_slots_def new_cnode_def new_cap_map_def
irq_objects_def map_add_def empty_irq_node_def
split: split_if_asm)
apply clarsimp
apply (rule_tac x=obj_id in exI, clarsimp, rule conjI)
apply (clarsimp simp: example_spec_def cnode_defs opt_object_def
dest!: object_slots_new_cnode_D
split: split_if_asm)
apply (fastforce simp: is_orig_cap_example_spec opt_cap_def slots_of_def opt_object_def)
done
lemma well_formed_orig_cap_example [simp]:
"\<lbrakk>cdl_objects example_spec obj_id = Some obj;
object_slots obj slot = Some cap; cap \<noteq> NullCap;
original_cap_at (obj_id, slot) example_spec \<rbrakk>
\<Longrightarrow> well_formed_orig_cap cap"
apply (clarsimp simp: is_orig_cap_example_spec well_formed_orig_cap_def)
apply (clarsimp simp: example_spec_def object_slots_def obj_defs new_cnode_def new_cap_map_def
new_irq_node_def ep_related_cap_def cap_type_def default_cap_def cap_rights_def
split: split_if_asm)
done
lemma well_formed_caps_example:
"cdl_objects example_spec obj_id = Some obj \<Longrightarrow>
well_formed_caps example_spec obj_id obj"
apply (clarsimp simp: well_formed_caps_def, rule conjI)
apply (clarsimp simp: is_orig_cap_example_spec)
apply (clarsimp simp: example_spec_def obj_defs object_type_def cap_type_def object_slots_def
is_copyable_cap_def is_irqhandler_cap_def
dest!: cdl_cnode_caps_new_cnode_D new_cap_map_caps_D
split: split_if_asm)
apply (rule conjI)
apply (clarsimp simp: well_formed_cap_to_real_object_def real_object_at_def irq_cnodes_example_spec)
apply (clarsimp simp: example_spec_def obj_defs object_slots_def onehundred_not_le_one
dest!: cdl_cnode_caps_new_cnode_D new_cap_map_caps_D
irq_objects_some_object
split: split_if_asm)
apply (fastforce simp: new_irq_node_def empty_irq_node_def split: split_if_asm)
apply (fastforce simp: new_irq_node_def empty_irq_node_def split: split_if_asm)
apply (rule conjI)
apply (clarsimp simp: well_formed_cap_types_match_def)
apply (rule conjI)
apply (clarsimp simp: example_spec_def object_slots_def obj_defs
irq_objects_def map_add_def new_irq_node_def
onehundred_not_le_one
dest!: cdl_cnode_caps_new_cnode_D new_cap_map_caps_D
split: split_if_asm)
apply (clarsimp simp: example_spec_def object_slots_def obj_defs
irq_objects_def map_add_def new_irq_node_def
onehundred_not_le_one
dest!: cdl_cnode_caps_new_cnode_D new_cap_map_caps_D
split: split_if_asm)
apply (clarsimp simp: example_spec_def obj_defs is_cnode_def is_tcb_def is_fake_vm_cap_def
object_slots_def object_type_def cap_type_def new_irq_node_def
empty_irq_node_def empty_cnode_def
dest!: cdl_cnode_caps_new_cnode_D irq_objects_some_object
split: split_if_asm cdl_object.splits)
done
lemma real_object_at_example_spec:
"real_object_at obj_id example_spec =
((obj_id = tcb_a_id) \<or>
(obj_id = tcb_b_id) \<or>
(obj_id = cnode_a1_id) \<or>
(obj_id = cnode_a2_id) \<or>
(obj_id = cnode_b_id) \<or>
(obj_id = cnode_extra_id) \<or>
(obj_id = ep_id) \<or>
(obj_id = aep_id) \<or>
(obj_id = pd_a_id) \<or>
(obj_id = pt_a_id) \<or>
(obj_id = pd_b_id) \<or>
(obj_id = frame_a1_id) \<or>
(obj_id = frame_a2_id) \<or>
(obj_id = frame_b_id))"
apply (clarsimp simp: real_object_at_def irq_cnodes_example_spec)
apply (clarsimp simp: example_spec_def irq_objects_def dom_def onehundred_not_le_one
split: split_if_asm)
done
lemma real_object_at_example_spec_simp [simp]:
"real_object_at 0 example_spec"
"real_object_at 1 example_spec = True"
"real_object_at 2 example_spec"
"real_object_at 3 example_spec"
"real_object_at 4 example_spec"
"real_object_at 5 example_spec"
"real_object_at 6 example_spec"
"real_object_at 7 example_spec"
"real_object_at 8 example_spec"
"real_object_at 9 example_spec"
"real_object_at 0xA example_spec"
"real_object_at 0xB example_spec"
"real_object_at 0xC example_spec"
"real_object_at 0xD example_spec"
"\<not>real_object_at 0x104 example_spec"
"\<not>real_object_at 0x1FE example_spec"
by (clarsimp simp: real_object_at_example_spec)+
lemma cdl_objects_example_spec_simps [simp]:
"cdl_objects example_spec 4 = Some (Frame empty_frame)"
"cdl_objects example_spec 0xD = Some AsyncEndpoint"
"cdl_objects example_spec 0x1FE = Some empty_irq_node"
by (clarsimp simp: example_spec_def map_add_def)+
lemma well_formed_cap_to_object_example:
"cdl_objects example_spec obj_id = Some obj
\<Longrightarrow> well_formed_cap_to_object example_spec obj_id obj"
apply (clarsimp simp: well_formed_cap_to_object_def is_orig_cap_example_spec)
apply (intro conjI)
apply (case_tac "obj_id = ep_id \<or>
obj_id = aep_id \<or>
obj_id = cnode_extra_id")
apply (rule_tac x=cnode_extra_id in exI)
apply (fastforce simp: real_cnode_at_example_spec cnode_extra_def
object_slots_def new_cnode_def new_cap_map_def)
apply (case_tac "obj_id = tcb_b_id \<or>
obj_id = cnode_b_id \<or>
obj_id = pd_b_id \<or>
obj_id = frame_b_id")
apply (rule_tac x=cnode_b_id in exI)
apply (fastforce simp: real_cnode_at_example_spec cnode_b_def
object_slots_def new_cnode_def new_cap_map_def)
apply (case_tac "obj_id = tcb_a_id \<or>
obj_id = cnode_a2_id")
apply (rule_tac x=cnode_a1_id in exI)
apply (fastforce simp: real_cnode_at_example_spec cnode_a1_def
object_slots_def new_cnode_def new_cap_map_def)
apply (case_tac "obj_id = cnode_a1_id \<or>
obj_id = pd_a_id \<or>
obj_id = pt_a_id \<or>
obj_id = frame_a1_id \<or>
obj_id = ep_id")
apply (rule_tac x=cnode_a2_id in exI)
apply (fastforce simp: real_cnode_at_example_spec cnode_a2_def
object_slots_def new_cnode_def new_cap_map_def)
apply (case_tac "obj_id = frame_a2_id")
apply (rule_tac x=cnode_a2_id in exI)
apply (rule_tac x=11 in exI) (* Not sure why fastforce gives up here. *)
apply (fastforce simp: real_cnode_at_example_spec cnode_a2_def
object_slots_def new_cnode_def new_cap_map_def)
apply (case_tac "obj_id = 0x104")
apply (rule_tac x=cnode_a2_id in exI)
apply (rule_tac x=12 in exI) (* Not sure why fastforce gives up here. *)
apply (fastforce simp: real_cnode_at_example_spec cnode_a2_def
object_slots_def new_cnode_def new_cap_map_def)
apply (case_tac "obj_id = 0x1FE")
apply (rule_tac x=cnode_b_id in exI)
apply (fastforce simp: real_cnode_at_example_spec cnode_b_def
object_slots_def new_cnode_def new_cap_map_def)
apply (clarsimp simp: example_spec_def)
apply clarsimp
apply (clarsimp simp: example_spec_def)
apply (clarsimp simp: opt_cap_def slots_of_def opt_object_def obj_defs
object_slots_def object_size_bits_def
new_cap_map_def empty_cap_map_def frame_cap_not_cnode
empty_irq_node_def new_irq_node_def
split: split_if_asm
| drule (1) cdl_cnode_caps_new_cnode_cnode_cap)+ (* Takes 20 seconds. *)
done
lemma well_formed_cap_to_non_empty_pt_example:
"cdl_objects example_spec obj_id = Some obj \<Longrightarrow>
well_formed_cap_to_non_empty_pt example_spec obj_id obj"
apply (clarsimp simp: well_formed_cap_to_non_empty_pt_def pt_at_example_spec)
apply (rule exI [where x=pd_a_id])
apply (clarsimp simp: well_formed_cap_to_non_empty_pt_def example_spec_def is_pt_def
object_at_def opt_cap_def slots_of_def object_slots_def opt_object_def
obj_defs new_cap_map_def is_pd_def empty_cap_map_def
split: split_if_asm)
done
lemma well_formed_vspace_example:
"cdl_objects example_spec obj_id = Some obj
\<Longrightarrow> well_formed_vspace example_spec obj_id obj"
apply (clarsimp simp: well_formed_vspace_def well_formed_cap_to_non_empty_pt_example)
apply (clarsimp simp: example_spec_def is_pt_def is_pd_def object_slots_def empty_cap_map_def
new_irq_node_def
split: split_if_asm)
apply (fastforce simp: cap_type_def is_fake_vm_cap_def obj_defs new_cap_map_def small_section_size_def
split: split_if_asm)
apply (clarsimp simp: obj_defs new_cap_map_def cap_type_def small_frame_size_def
is_fake_vm_cap_def is_fake_pt_cap_simps small_section_size_def
split: split_if_asm)
apply (clarsimp simp: obj_defs new_cap_map_def cap_type_def small_frame_size_def
is_fake_vm_cap_def is_fake_pt_cap_simps small_section_size_def
split: split_if_asm)
done
lemma well_formed_irqhandler_caps_unique_example_spec:
"well_formed_irqhandler_caps_unique example_spec"
apply (clarsimp simp: well_formed_irqhandler_caps_unique_def)
apply (drule (1) irq_handler_cap_example_spec)+
apply (clarsimp simp: example_spec_def opt_cap_def slots_of_def opt_object_def
object_slots_def is_irqhandler_cap_def obj_defs
new_cnode_def new_cap_map_def
split: split_if_asm)
done
lemma ucast_0xFE:
"(ucast :: 8 word \<Rightarrow> 32 word) irq = 0xFE \<Longrightarrow> irq = 0xFE"
by (rule ucast_up_inj, simp+)
lemma ucast_4:
"(ucast :: 8 word \<Rightarrow> 32 word) irq = 4 \<Longrightarrow> irq = 4"
by (rule ucast_up_inj, simp+)
lemma rangeD:
"\<lbrakk>range f = A; f x = y\<rbrakk> \<Longrightarrow> y \<in> A"
by (fastforce simp: image_def)
lemma slots_of_example_irq_node:
"\<lbrakk>slots_of (example_irq_node irq) example_spec 0 = Some cap;
cap \<noteq> NullCap\<rbrakk>
\<Longrightarrow> (irq = 4)"
apply (frule (1) slots_of_example_spec_obj_ids)
apply (insert range_example_irq_node)
apply (erule disjE, drule (1) rangeD, simp add: onehundred_not_le_one)+
apply (clarsimp simp: example_irq_node_def ucast_4)
done
lemma bound_irqs_example_spec [simp]:
"bound_irqs example_spec = {4}"
apply (clarsimp simp: bound_irqs_def)
apply rule
apply clarsimp
apply (erule (1) slots_of_example_irq_node)
apply (clarsimp simp: example_spec_def slots_of_def opt_object_def
object_slots_def new_irq_node_def)
done
lemma well_formed_irqhandler_caps_example_spec:
"well_formed_irqhandler_caps example_spec"
apply (clarsimp simp: well_formed_irqhandler_caps_def)
apply (rule exI [where x=cnode_a2_id])
apply (rule exI [where x=12])
apply (rule exI [where x="IrqHandlerCap 4"])
apply (clarsimp simp: object_slots_def cnode_a2_def new_cnode_def new_cap_map_def)
done
lemma well_formed_irq_table_example_spec:
"well_formed_irq_table example_spec"
apply (clarsimp simp: well_formed_irq_table_def)
apply (clarsimp simp: example_irq_node_def)
apply (insert ucast_8_32_inj)
apply (clarsimp simp: inj_on_def)
done
lemma well_formed_tcb_example_spec:
"cdl_objects example_spec obj_id = Some obj \<Longrightarrow>
well_formed_tcb example_spec obj_id obj"
apply (case_tac "obj_id = tcb_a_id")
apply (cut_tac obj_id = tcb_a_id in well_formed_tcb_a)
apply (clarsimp simp: example_spec_def split: split_if_asm)
apply (case_tac "obj_id = tcb_b_id")
apply (cut_tac obj_id = tcb_b_id in well_formed_tcb_b)
apply (clarsimp simp: example_spec_def split: split_if_asm)
apply (clarsimp simp: example_spec_def well_formed_tcb_def is_tcb_def
empty_irq_node_def new_irq_node_def
split: split_if_asm)
done
lemma well_formed_irq_node_example_spec:
"cdl_objects example_spec obj_id = Some obj \<Longrightarrow>
well_formed_irq_node example_spec obj_id obj"
apply (clarsimp simp: well_formed_irq_node_def irq_cnodes_example_spec)
apply (clarsimp simp: example_spec_def object_slots_def empty_irq_node_def new_irq_node_def
empty_cnode_def empty_cap_map_def dom_def
is_default_cap_def default_cap_def onehundred_not_le_one
split: split_if_asm)
done
lemma well_formed_example:
"well_formed example_spec"
apply (clarsimp simp: well_formed_def)
apply (intro conjI)
apply (rule well_formed_orig_caps_unique_example)
apply (rule well_formed_irqhandler_caps_unique_example_spec)
apply (rule well_formed_fake_pt_caps_unique_example)
apply (rule well_formed_irqhandler_caps_example_spec)
apply (rule well_formed_irq_table_example_spec)
apply (clarsimp split: option.splits, rename_tac obj)
apply (clarsimp simp: well_formed_caps_example well_formed_cap_to_object_example
well_formed_orig_caps_unique_example)
apply (rule conjI)
apply (fact well_formed_tcb_example_spec)
apply (rule conjI)
apply (fact well_formed_vspace_example)
apply (rule conjI)
apply (fact well_formed_irq_node_example_spec)
apply (clarsimp simp: real_cnode_at_example_spec)
apply (auto simp: example_spec_def object_size_bits_def object_default_state_def2
pd_size_def word_bits_def empty_cnode_def is_cnode_def
object_slots_def empty_cap_map_def tcb_slot_defs
default_tcb_def obj_defs tcb_pending_op_slot_def
small_frame_size_def small_section_size_def pt_size_def
new_cnode_def new_cap_map_def empty_irq_node_def
new_irq_node_def
split: split_if_asm)
done
end