-
Notifications
You must be signed in to change notification settings - Fork 1
/
Access.thy
3037 lines (2492 loc) · 135 KB
/
Access.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: GPL-2.0-only
*)
theory Access
imports Deterministic_AC
begin
context begin interpretation Arch . (*FIXME: arch_split*)
section \<open>Policy and policy refinement\<close>
subsection \<open>Policy definition\<close>
text\<open>
The goal is to place limits on what untrusted agents can do while the
trusted agents are not running. This supports a framing constraint in
the descriptions (Hoare triples) of syscalls. Roughly the existing
proofs show the effect of the syscall, and this proof summarises what
doesn't (or isn't allowed to) change.
The basic intuition is to map all object references to the agent they
belong to and show that changes to the object reference graph are
allowed by the policy specified by the user. The policy is a labelled
graph amongst agents independent of the current state, i.e. a static
external summary of what should be talking to what and how.
The interesting cases occur between trusted and untrusted components:
e.g. we assume that it is unsafe for untrusted components (UT)
to send capabilities to trusted components (T), and so T must ensure
that all endpoints it shares with UT that it receives on do not have
the grant bit set.
\<close>
type_synonym 'a agent_map = "obj_ref \<Rightarrow> 'a"
type_synonym 'a agent_asid_map = "asid \<Rightarrow> 'a"
type_synonym 'a agent_irq_map = "10 word \<Rightarrow> 'a"
type_synonym 'a agent_domain_map = "domain \<Rightarrow> 'a set"
text\<open>
What one agent can do to another. We allow multiple edges between
agents in the graph.
Control is special. It implies the ability to do pretty much anything,
including get access the other rights, create, remove, etc.
DeleteDerived allows you to delete a cap derived from a cap you own
\<close>
datatype auth = Control | Receive | SyncSend | Notify
| Reset | Grant | Call | Reply | Write | Read
| ASIDPoolMapsASID | DeleteDerived
text\<open>
The interesting case is for endpoints.
Consider an EP between T and UT (across a trust boundary). We want to
be able to say that all EPs and tcbs that T does not expose to UT do
not change when it is not running. If UT had a direct Send right to T
then integrity (see below) could not guarantee this, as it doesn't
know which EP can be changed by UT. Thus we set things up like this
(all distinct labels):
T -Receive-> EP <-Send- UT
Now UT can interfere with EP and all of T's tcbs blocked for receive on EP,
but not endpoints internal to T, or tcbs blocked on other (suitably
labelled) EPs, etc.
\<close>
type_synonym 'a auth_graph = "('a \<times> auth \<times> 'a) set"
text \<open>
Each global namespace will need a labeling function.
We map each scheduling domain to a single label; concretely, each tcb
in a scheduling domain has to have the same label. We will want to
weaken this in the future.
The booleans @{text pasMayActivate} and @{text pasMayEditReadyQueues}
are used to weaken the integrity property. When @{const True},
@{text pasMayActivate} causes the integrity property to permit
activation of newly-scheduled threads. Likewise, @{text pasMayEditReadyQueues}
has the integrity property permit the removal of threads from ready queues,
as occurs when scheduling a new domain for instance. By setting each of these
@{const False} we get a more constrained integrity property that is useful for
establishing some of the proof obligations for the infoflow proofs, particularly
those over @{const handle_event} that neither activates new threads nor schedules
new domains.
The @{text pasDomainAbs} relation describes which labels may run in
a given scheduler domain. This relation is not relevant to integrity
but will be used in the information flow theory (with additional
constraints on its structure).
\<close>
end
record 'a PAS =
pasObjectAbs :: "'a agent_map"
pasASIDAbs :: "'a agent_asid_map"
pasIRQAbs :: "'a agent_irq_map"
pasPolicy :: "'a auth_graph"
pasSubject :: "'a" \<comment> \<open>The active label\<close>
pasMayActivate :: "bool"
pasMayEditReadyQueues :: "bool"
pasMaySendIrqs :: "bool"
pasDomainAbs :: "'a agent_domain_map"
text\<open>
Very often we want to say that the agent currently running owns a
given pointer.
\<close>
abbreviation is_subject :: "'a PAS \<Rightarrow> word32 \<Rightarrow> bool"
where
"is_subject aag ptr \<equiv> pasObjectAbs aag ptr = pasSubject aag"
text\<open>
Also we often want to say the current agent can do something to a
pointer that he doesn't own but has some authority to.
\<close>
abbreviation(input) aag_has_auth_to :: "'a PAS \<Rightarrow> auth \<Rightarrow> word32 \<Rightarrow> bool"
where
"aag_has_auth_to aag auth ptr \<equiv> (pasSubject aag, auth, pasObjectAbs aag ptr) \<in> pasPolicy aag"
abbreviation aag_subjects_have_auth_to_label :: "'a set \<Rightarrow> 'a PAS \<Rightarrow> auth \<Rightarrow> 'a \<Rightarrow> bool"
where
"aag_subjects_have_auth_to_label subs aag auth label
\<equiv> \<exists>s \<in> subs. (s, auth, label) \<in> pasPolicy aag"
abbreviation aag_subjects_have_auth_to :: "'a set \<Rightarrow> 'a PAS \<Rightarrow> auth \<Rightarrow> obj_ref \<Rightarrow> bool"
where
"aag_subjects_have_auth_to subs aag auth oref
\<equiv> aag_subjects_have_auth_to_label subs aag auth (pasObjectAbs aag oref)"
context begin interpretation Arch . (*FIXME: arch_split*)
subsection \<open>Policy wellformedness\<close>
text\<open>
Wellformedness of the agent authority function with respect to a label
(the current thread):
\begin{itemize}
\item For (the current untrusted) @{term "agent"}, large enough
authority must be contained within the agent's boundaries.
\item @{term "agent"} has all the self authority it could want.
\item If an agent can grant caps through an endpoint, then it is
authority-equivalent to all agents that can receive on that
endpoint.
\item Similarly, if an agent can grant through a reply cap, then
it is authority-equivalent to the original caller.
\item Anyone can send on any IRQ notification.
\item Call implies a send ability.
\item If an agent could reply to a call, then the caller has the
authority to delete the derived reply cap. This can happen if
the caller thread is deleted before the reply takes place.
\item Reply caps can be transferred, so the DeleteDerived
authority propagates transitively.
\end{itemize}
\<close>
definition policy_wellformed
where
"policy_wellformed aag maySendIrqs irqs agent \<equiv>
(\<forall>agent'. (agent, Control, agent') \<in> aag \<longrightarrow> agent = agent')
\<and> (\<forall>a. (agent, a, agent) \<in> aag)
\<and> (\<forall>s r ep. (s, Grant, ep) \<in> aag \<and> (r, Receive, ep) \<in> aag
\<longrightarrow> (s, Control, r) \<in> aag \<and> (r, Control, s) \<in> aag)
\<and> (maySendIrqs \<longrightarrow> (\<forall>irq ntfn. irq \<in> irqs \<and> (irq, Notify, ntfn) \<in> aag
\<longrightarrow> (agent, Notify, ntfn) \<in> aag))
\<and> (\<forall>s ep. (s, Call, ep) \<in> aag \<longrightarrow> (s, SyncSend, ep) \<in> aag)
\<and> (\<forall>s r ep. (s, Call, ep) \<in> aag \<and> (r, Receive, ep) \<in> aag \<longrightarrow> (r, Reply, s) \<in> aag)
\<and> (\<forall>s r. (s, Reply, r) \<in> aag \<longrightarrow> (r, DeleteDerived, s) \<in> aag)
\<and> (\<forall>l1 l2 l3. (l1, DeleteDerived, l2) \<in> aag \<longrightarrow> (l2, DeleteDerived, l3) \<in> aag
\<longrightarrow> (l1, DeleteDerived, l3) \<in> aag)
\<and> (\<forall>s r ep. (s, Call, ep) \<in> aag \<and> (r, Receive, ep) \<in> aag \<and> (r, Grant, ep) \<in> aag
\<longrightarrow> (s, Control, r) \<in> aag \<and> (r, Control, s) \<in> aag)"
abbreviation pas_wellformed
where
"pas_wellformed aag \<equiv>
policy_wellformed (pasPolicy aag) (pasMaySendIrqs aag) (range (pasIRQAbs aag)) (pasSubject aag)"
lemma aag_wellformed_Control:
"\<lbrakk> (x, Control, y) \<in> aag; policy_wellformed aag mirqs irqs x \<rbrakk> \<Longrightarrow> x = y"
unfolding policy_wellformed_def by metis
lemma aag_wellformed_refl:
"\<lbrakk> policy_wellformed aag mirqs irqs x \<rbrakk> \<Longrightarrow> (x, a, x) \<in> aag"
unfolding policy_wellformed_def by blast
lemma aag_wellformed_grant_Control_to_recv:
"\<lbrakk> (s, Grant, ep) \<in> aag; (r, Receive, ep) \<in> aag; policy_wellformed aag mirqs irqs l\<rbrakk>
\<Longrightarrow> (s, Control, r) \<in> aag"
unfolding policy_wellformed_def by blast
lemma aag_wellformed_grant_Control_to_send:
"\<lbrakk> (s, Grant, ep) \<in> aag; (r, Receive, ep) \<in> aag; policy_wellformed aag mirqs irqs l\<rbrakk>
\<Longrightarrow> (r, Control, s) \<in> aag"
unfolding policy_wellformed_def by blast
lemma aag_wellformed_reply:
"\<lbrakk> (s, Call, ep) \<in> aag; (r, Receive, ep) \<in> aag; policy_wellformed aag mirqs irqs l\<rbrakk>
\<Longrightarrow> (r, Reply, s) \<in> aag"
unfolding policy_wellformed_def by blast
lemma aag_wellformed_delete_derived':
"\<lbrakk> (s, Call, ep) \<in> aag; (r, Receive, ep) \<in> aag; policy_wellformed aag mirqs irqs l\<rbrakk>
\<Longrightarrow> (s, DeleteDerived, r) \<in> aag"
unfolding policy_wellformed_def by blast
lemma aag_wellformed_delete_derived:
"\<lbrakk> (s, Reply, r) \<in> aag; policy_wellformed aag mirqs irqs l\<rbrakk>
\<Longrightarrow> (r, DeleteDerived, s) \<in> aag"
unfolding policy_wellformed_def by blast
lemma aag_wellformed_delete_derived_trans:
"\<lbrakk> (l1, DeleteDerived, l2) \<in> aag; (l2, DeleteDerived, l3) \<in> aag;
policy_wellformed aag mirqs irqs l\<rbrakk>
\<Longrightarrow> (l1, DeleteDerived, l3) \<in> aag"
unfolding policy_wellformed_def by blast
lemma aag_wellformed_call_to_syncsend:
"\<lbrakk> (s, Call, ep) \<in> aag; policy_wellformed aag mirqs irqs l\<rbrakk>
\<Longrightarrow> (s, SyncSend, ep) \<in> aag"
unfolding policy_wellformed_def by blast
lemma aag_wellformed_grant_Control_to_send_by_reply:
"\<lbrakk> (s, Call, ep) \<in> aag; (r, Receive, ep) \<in> aag; (r, Grant, ep) \<in> aag;
policy_wellformed aag mirqs irqs l\<rbrakk>
\<Longrightarrow> (r, Control, s) \<in> aag"
unfolding policy_wellformed_def by blast
lemma aag_wellformed_grant_Control_to_recv_by_reply:
"\<lbrakk> (s, Call, ep) \<in> aag; (r, Receive, ep) \<in> aag; (r, Grant, ep) \<in> aag;
policy_wellformed aag mirqs irqs l\<rbrakk>
\<Longrightarrow> (s, Control, r) \<in> aag"
unfolding policy_wellformed_def by blast
subsection \<open>auth_graph_map\<close>
text\<open>
Abstract a graph by relabelling the nodes (agents). Clearly this can
collapse (and not create) distinctions.
\<close>
definition
auth_graph_map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a auth_graph \<Rightarrow> 'b auth_graph"
where
"auth_graph_map f aag \<equiv> {(f x, auth, f y) | x auth y. (x, auth, y) \<in> aag}"
lemma auth_graph_map_mem:
"(x, auth, y) \<in> auth_graph_map f S
= (\<exists>x' y'. x = f x' \<and> y = f y' \<and> (x', auth, y') \<in> S)"
by (simp add: auth_graph_map_def)
lemmas auth_graph_map_memD = auth_graph_map_mem[THEN iffD1]
lemma auth_graph_map_memE:
assumes hyp:"(x, auth, y) \<in> auth_graph_map f S"
obtains x' and y' where "x = f x'" and "y = f y'" and "(x', auth, y') \<in> S"
using hyp[THEN auth_graph_map_mem[THEN iffD1]] by fastforce
lemma auth_graph_map_memI:
"\<lbrakk> (x', auth, y') \<in> S; x = f x'; y = f y' \<rbrakk>
\<Longrightarrow> (x, auth, y) \<in> auth_graph_map f S"
by (fastforce simp add: auth_graph_map_mem)
lemma auth_graph_map_mono:
"S \<subseteq> S' \<Longrightarrow> auth_graph_map G S \<subseteq> auth_graph_map G S'"
unfolding auth_graph_map_def by auto
subsection \<open>Transform caps and tcb states into authority\<close>
text\<open>
Abstract the state to an agent authority graph. This definition states
what authority is conferred by a particular capability to the obj_refs
in it.
\<close>
definition cap_rights_to_auth :: "cap_rights \<Rightarrow> bool \<Rightarrow> auth set"
where
"cap_rights_to_auth r sync \<equiv>
{Reset}
\<union> (if AllowRead \<in> r then {Receive} else {})
\<union> (if AllowWrite \<in> r then (if sync then {SyncSend} else {Notify}) else {})
\<union> (if AllowGrant \<in> r then UNIV else {})
\<union> (if AllowGrantReply \<in> r \<and> AllowWrite \<in> r then {Call} else {})"
definition vspace_cap_rights_to_auth :: "cap_rights \<Rightarrow> auth set"
where
"vspace_cap_rights_to_auth r \<equiv>
(if AllowWrite \<in> r then {Write} else {})
\<union> (if AllowRead \<in> r then {Read} else {})"
definition reply_cap_rights_to_auth :: "bool \<Rightarrow> cap_rights \<Rightarrow> auth set"
where
"reply_cap_rights_to_auth master r \<equiv> if AllowGrant \<in> r \<or> master then UNIV else {Reply}"
definition cap_auth_conferred :: "cap \<Rightarrow> auth set"
where
"cap_auth_conferred cap \<equiv>
case cap of
Structures_A.NullCap \<Rightarrow> {}
| Structures_A.UntypedCap isdev oref bits freeIndex \<Rightarrow> {Control}
| Structures_A.EndpointCap oref badge r \<Rightarrow>
cap_rights_to_auth r True
| Structures_A.NotificationCap oref badge r \<Rightarrow>
cap_rights_to_auth (r - {AllowGrant, AllowGrantReply}) False
| Structures_A.ReplyCap oref m r \<Rightarrow> reply_cap_rights_to_auth m r
| Structures_A.CNodeCap oref bits guard \<Rightarrow> {Control}
| Structures_A.ThreadCap obj_ref \<Rightarrow> {Control}
| Structures_A.DomainCap \<Rightarrow> {Control}
| Structures_A.IRQControlCap \<Rightarrow> {Control}
| Structures_A.IRQHandlerCap irq \<Rightarrow> {Control}
| Structures_A.Zombie ptr b n \<Rightarrow> {Control}
| Structures_A.ArchObjectCap arch_cap \<Rightarrow>
(if is_page_cap (the_arch_cap cap)
then vspace_cap_rights_to_auth (acap_rights arch_cap)
else {Control})"
fun tcb_st_to_auth :: "Structures_A.thread_state \<Rightarrow> (obj_ref \<times> auth) set"
where
"tcb_st_to_auth (Structures_A.BlockedOnNotification ntfn) = {(ntfn, Receive)}"
| "tcb_st_to_auth (Structures_A.BlockedOnSend ep payl)
= {(ep, SyncSend)} \<union> (if sender_can_grant payl then {(ep, Grant),(ep, Call)} else {})
\<union> (if sender_can_grant_reply payl then {(ep, Call)} else {})"
| "tcb_st_to_auth (Structures_A.BlockedOnReceive ep payl)
= {(ep, Receive)} \<union> (if receiver_can_grant payl then {(ep, Grant)} else {})"
| "tcb_st_to_auth _ = {}"
subsection \<open>FIXME MOVE\<close>
definition pte_ref
where
"pte_ref pte \<equiv> case pte of
SmallPagePTE addr atts rights
\<Rightarrow> Some (ptrFromPAddr addr, pageBitsForSize ARMSmallPage, vspace_cap_rights_to_auth rights)
| LargePagePTE addr atts rights
\<Rightarrow> Some (ptrFromPAddr addr, pageBitsForSize ARMLargePage, vspace_cap_rights_to_auth rights)
| _ \<Rightarrow> None"
definition pde_ref2
where
"pde_ref2 pde \<equiv> case pde of
SectionPDE addr atts domain rights
\<Rightarrow> Some (ptrFromPAddr addr, pageBitsForSize ARMSection, vspace_cap_rights_to_auth rights)
| SuperSectionPDE addr atts rights
\<Rightarrow> Some (ptrFromPAddr addr, pageBitsForSize ARMSuperSection, vspace_cap_rights_to_auth rights)
| PageTablePDE addr atts domain
\<comment> \<open>The 0 is a hack, saying that we own only addr, although 12 would also be OK\<close>
\<Rightarrow> Some (ptrFromPAddr addr, 0, {Control})
| _ \<Rightarrow> None"
definition ptr_range
where
"ptr_range p sz \<equiv> {p .. p + 2 ^ sz - 1}"
(* We exclude the global page tables from the authority graph. Alternatively,
we could include them and add a wellformedness constraint to policies that
requires that every label has the necessary authority to whichever label owns
the global page tables, so that when a new page directory is created and
references to the global page tables are added to it, no new authority is gained.
Note: excluding the references to global page tables in this way
brings in some ARM arch-specific VM knowledge here *)
definition vs_refs_no_global_pts :: "Structures_A.kernel_object \<Rightarrow> (obj_ref \<times> vs_ref \<times> auth) set"
where
"vs_refs_no_global_pts \<equiv> \<lambda>ko. case ko of
ArchObj (ASIDPool pool) \<Rightarrow>
(\<lambda>(r,p). (p, VSRef (ucast r) (Some AASIDPool), Control)) ` graph_of pool
| ArchObj (PageDirectory pd) \<Rightarrow>
\<Union>(r,(p, sz, auth)) \<in> graph_of (pde_ref2 o pd) - {(x,y). (ucast (kernel_base >> 20) \<le> x)}.
(\<lambda>(p, a). (p, VSRef (ucast r) (Some APageDirectory), a)) ` (ptr_range p sz \<times> auth)
| ArchObj (PageTable pt) \<Rightarrow>
\<Union>(r,(p, sz, auth)) \<in> graph_of (pte_ref o pt).
(\<lambda>(p, a). (p, VSRef (ucast r) (Some APageTable), a)) ` (ptr_range p sz \<times> auth)
| _ \<Rightarrow> {}"
subsection \<open>Transferability: Moving caps between agents\<close>
text \<open>
Tells if cap can move/be derived between agents without grant
due to the inner workings of the system (Calling and replying for now)
\<close>
(* FIXME is_transferable should garantee directly that a non-NullCap cap is owned by its CDT
parents without using directly the CDT so that we can use it in integrity *)
inductive is_transferable for opt_cap
where
it_None: "opt_cap = None \<Longrightarrow> is_transferable opt_cap" |
it_Null: "opt_cap = Some NullCap \<Longrightarrow> is_transferable opt_cap" |
it_Reply: "opt_cap = Some (ReplyCap t False R) \<Longrightarrow> is_transferable opt_cap"
abbreviation "is_transferable_cap cap \<equiv> is_transferable (Some cap)"
abbreviation "is_transferable_in slot s \<equiv> is_transferable (caps_of_state s slot)"
lemma is_transferable_capE:
assumes hyp:"is_transferable_cap cap"
obtains "cap = NullCap" | t R where "cap = ReplyCap t False R"
by (rule is_transferable.cases[OF hyp]) auto
lemma is_transferable_None[simp]: "is_transferable None"
using is_transferable.simps by simp
lemma is_transferable_Null[simp]: "is_transferable_cap NullCap"
using is_transferable.simps by simp
lemma is_transferable_Reply[simp]: "is_transferable_cap (ReplyCap t False R)"
using is_transferable.simps by simp
lemma is_transferable_NoneE:
assumes hyp: "is_transferable opt_cap"
obtains "opt_cap = None" | cap where "opt_cap = Some cap" and "is_transferable_cap cap"
by (rule is_transferable.cases[OF hyp];simp)
lemma is_transferable_Untyped[simp]: "\<not> is_transferable (Some (UntypedCap dev ptr sz freeIndex))"
by (blast elim: is_transferable.cases)
lemma is_transferable_IRQ[simp]: "\<not> is_transferable_cap IRQControlCap"
by (blast elim: is_transferable.cases)
lemma is_transferable_Zombie[simp]: "\<not> is_transferable (Some (Zombie ptr sz n))"
by (blast elim: is_transferable.cases)
lemma is_transferable_Ntfn[simp]: "\<not> is_transferable (Some (NotificationCap ntfn badge R))"
by (blast elim: is_transferable.cases)
lemma is_transferable_Endpoint[simp]: "\<not> is_transferable (Some (EndpointCap ntfn badge R))"
by (blast elim: is_transferable.cases)
(*FIXME MOVE *)
lemmas descendants_incD
= descendants_inc_def[THEN meta_eq_to_obj_eq, THEN iffD1, rule_format]
lemma acap_class_reply:
"(acap_class acap \<noteq> ReplyClass t)"
(* warning: arch split *)
by (cases acap; simp)
lemma cap_class_reply:
"(cap_class cap = ReplyClass t) = (\<exists>r m. cap = ReplyCap t m r)"
by (cases cap; simp add: acap_class_reply)
lemma reply_cap_no_children:
"\<lbrakk> valid_mdb s; caps_of_state s p = Some (ReplyCap t False r) \<rbrakk> \<Longrightarrow> cdt s p' \<noteq> Some p"
apply (clarsimp simp: valid_mdb_def)
apply (frule descendants_incD[where p=p' and p'=p])
apply (rule child_descendant)
apply (fastforce)
apply clarsimp
apply (subgoal_tac "caps_of_state s p' \<noteq> None")
apply (clarsimp simp: cap_class_reply)
apply (simp only: reply_mdb_def reply_caps_mdb_def reply_masters_mdb_def)
apply (elim conjE allE[where x=p'])
apply simp
apply (drule(1) mdb_cte_atD)
apply (fastforce simp add: cte_wp_at_def caps_of_state_def)
done
(*FIXME MOVE *)
lemmas all_childrenI = all_children_def[THEN meta_eq_to_obj_eq, THEN iffD2, rule_format]
lemmas all_childrenD = all_children_def[THEN meta_eq_to_obj_eq, THEN iffD1, rule_format]
(*FIXME MOVE *)
lemma valid_mdb_descendants_inc[elim!]:
"valid_mdb s \<Longrightarrow> descendants_inc (cdt s) (caps_of_state s)"
by (simp add:valid_mdb_def)
(*FIXME MOVE *)
lemma valid_mdb_mdb_cte_at [elim!]:
"valid_mdb s \<Longrightarrow> mdb_cte_at (swp (cte_wp_at ((\<noteq>) NullCap)) s) (cdt s)"
by (simp add:valid_mdb_def)
(*FIXME MOVE *)
lemma descendants_inc_cap_classD:
"\<lbrakk> descendants_inc m caps; p \<in> descendants_of p' m; caps p = Some cap ; caps p' = Some cap'\<rbrakk>
\<Longrightarrow> cap_class cap = cap_class cap'"
by (fastforce dest:descendants_incD)
lemma is_transferable_all_children:
"valid_mdb s \<Longrightarrow> all_children (\<lambda>slot . is_transferable (caps_of_state s slot)) (cdt s)"
apply (rule all_childrenI)
apply (erule is_transferable.cases)
apply (fastforce dest: mdb_cte_atD valid_mdb_mdb_cte_at
simp: mdb_cte_at_def cte_wp_at_def caps_of_state_def)
apply (fastforce dest: mdb_cte_atD valid_mdb_mdb_cte_at
simp: mdb_cte_at_def cte_wp_at_def caps_of_state_def)
apply (blast dest:reply_cap_no_children)
done
subsection \<open>Generating a policy from the current cap, ASID and IRQs distribution\<close>
(* TODO split that section between stba sata and sita and move a maximun
of accesor functions back to AInvs *)
text \<open>
sbta_caps/sbta_asid imply that a thread and it's vspace are labelled
the same -- caps_of_state (tcb, vspace_index) will be the PD cap.
Thus, a thread is completely managed or completely self-managed.
We can (possibly) weaken this by only talking about addressable caps
(i.e., only cspace in a tcb). This would also mean that we should use
the current cspace for the current label ... a bit strange, though.
The set of all objects affected by a capability. We cheat a bit and
say that a DomainCap contains references to everything, as it
intuitively grants its owners that sort of access. This allows us to
reuse sbta for DomainCaps.
The sbta definition is non-inductive. We use the "inductive"
construct for convenience, i.e. to get a nice set of intro rules,
cases, etc.
\<close>
primrec aobj_ref' :: "arch_cap \<Rightarrow> obj_ref set"
where
"aobj_ref' (arch_cap.ASIDPoolCap p as) = {p}"
| "aobj_ref' arch_cap.ASIDControlCap = {}"
| "aobj_ref' (arch_cap.PageCap isdev x rs sz as4) = ptr_range x (pageBitsForSize sz)"
| "aobj_ref' (arch_cap.PageDirectoryCap x as2) = {x}"
| "aobj_ref' (arch_cap.PageTableCap x as3) = {x}"
primrec obj_refs :: "cap \<Rightarrow> obj_ref set"
where
"obj_refs cap.NullCap = {}"
| "obj_refs (cap.ReplyCap r m cr) = {r}"
| "obj_refs cap.IRQControlCap = {}"
| "obj_refs (cap.IRQHandlerCap irq) = {}"
| "obj_refs (cap.UntypedCap d r s f) = {}"
| "obj_refs (cap.CNodeCap r bits guard) = {r}"
| "obj_refs (cap.EndpointCap r b cr) = {r}"
| "obj_refs (cap.NotificationCap r b cr) = {r}"
| "obj_refs (cap.ThreadCap r) = {r}"
| "obj_refs (cap.Zombie ptr b n) = {ptr}"
| "obj_refs (cap.ArchObjectCap x) = aobj_ref' x"
| "obj_refs cap.DomainCap = UNIV" (* hack, see above *)
inductive_set state_bits_to_policy for caps thread_sts thread_bas cdt vrefs
where
sbta_caps: "\<lbrakk> caps ptr = Some cap; oref \<in> obj_refs cap;
auth \<in> cap_auth_conferred cap \<rbrakk>
\<Longrightarrow> (fst ptr, auth, oref) \<in> state_bits_to_policy caps thread_sts thread_bas cdt vrefs"
| sbta_untyped: "\<lbrakk> caps ptr = Some cap; oref \<in> untyped_range cap \<rbrakk>
\<Longrightarrow> (fst ptr, Control, oref) \<in> state_bits_to_policy caps thread_sts thread_bas cdt vrefs"
| sbta_ts: "\<lbrakk> (oref', auth) \<in> thread_sts oref \<rbrakk>
\<Longrightarrow> (oref, auth, oref') \<in> state_bits_to_policy caps thread_sts thread_bas cdt vrefs"
| sbta_bounds: "\<lbrakk>thread_bas oref = Some oref'; auth \<in> {Receive, Reset} \<rbrakk>
\<Longrightarrow> (oref, auth, oref') \<in> state_bits_to_policy caps thread_sts thread_bas cdt vrefs"
| sbta_cdt: "\<lbrakk> cdt slot' = Some slot ; \<not>is_transferable (caps slot') \<rbrakk>
\<Longrightarrow> (fst slot, Control, fst slot')
\<in> state_bits_to_policy caps thread_sts thread_bas cdt vrefs"
| sbta_cdt_transferable: "\<lbrakk>cdt slot' = Some slot\<rbrakk>
\<Longrightarrow> (fst slot, DeleteDerived, fst slot')
\<in> state_bits_to_policy caps thread_sts thread_bas cdt vrefs"
| sbta_vref: "\<lbrakk> (ptr', ref, auth) \<in> vrefs ptr \<rbrakk>
\<Longrightarrow> (ptr, auth, ptr') \<in> state_bits_to_policy caps thread_sts thread_bas cdt vrefs"
fun cap_asid' :: "cap \<Rightarrow> asid set"
where
"cap_asid' (ArchObjectCap (PageCap _ _ _ _ mapping))
= fst ` set_option mapping"
| "cap_asid' (ArchObjectCap (PageTableCap _ mapping))
= fst ` set_option mapping"
| "cap_asid' (ArchObjectCap (PageDirectoryCap _ asid_opt))
= set_option asid_opt"
| "cap_asid' (ArchObjectCap (ASIDPoolCap _ asid))
= {x. asid_high_bits_of x = asid_high_bits_of asid \<and> x \<noteq> 0}"
| "cap_asid' (ArchObjectCap ASIDControlCap) = UNIV"
| "cap_asid' _ = {}"
inductive_set state_asids_to_policy_aux for aag caps asid_tab vrefs
where
sata_asid: "\<lbrakk> caps ptr = Some cap; asid \<in> cap_asid' cap \<rbrakk>
\<Longrightarrow> ( pasObjectAbs aag (fst ptr), Control, pasASIDAbs aag asid)
\<in> state_asids_to_policy_aux aag caps asid_tab vrefs"
| sata_asid_lookup: "\<lbrakk> asid_tab (asid_high_bits_of asid) = Some poolptr;
(pdptr, VSRef (asid && mask asid_low_bits)
(Some AASIDPool), a) \<in> vrefs poolptr \<rbrakk>
\<Longrightarrow> (pasASIDAbs aag asid, a, pasObjectAbs aag pdptr)
\<in> state_asids_to_policy_aux aag caps asid_tab vrefs"
| sata_asidpool: "\<lbrakk> asid_tab (asid_high_bits_of asid) = Some poolptr; asid \<noteq> 0 \<rbrakk> \<Longrightarrow>
(pasObjectAbs aag poolptr, ASIDPoolMapsASID, pasASIDAbs aag asid)
\<in> state_asids_to_policy_aux aag caps asid_tab vrefs"
fun cap_irqs_controlled :: "cap \<Rightarrow> irq set"
where
"cap_irqs_controlled cap.IRQControlCap = UNIV"
| "cap_irqs_controlled (cap.IRQHandlerCap irq) = {irq}"
| "cap_irqs_controlled _ = {}"
inductive_set state_irqs_to_policy_aux for aag caps
where
sita_controlled: "\<lbrakk> caps ptr = Some cap; irq \<in> cap_irqs_controlled cap \<rbrakk>
\<Longrightarrow> (pasObjectAbs aag (fst ptr), Control, pasIRQAbs aag irq) \<in> state_irqs_to_policy_aux aag caps"
abbreviation state_irqs_to_policy
where
"state_irqs_to_policy aag s \<equiv> state_irqs_to_policy_aux aag (caps_of_state s)"
definition irq_map_wellformed_aux
where
"irq_map_wellformed_aux aag irqs \<equiv> \<forall>irq. pasObjectAbs aag (irqs irq) = pasIRQAbs aag irq"
abbreviation irq_map_wellformed
where
"irq_map_wellformed aag s \<equiv> irq_map_wellformed_aux aag (interrupt_irq_node s)"
definition state_vrefs
where
"state_vrefs s = case_option {} vs_refs_no_global_pts o kheap s"
abbreviation state_asids_to_policy
where
"state_asids_to_policy aag s \<equiv> state_asids_to_policy_aux aag (caps_of_state s)
(arm_asid_table (arch_state s)) (state_vrefs s)"
definition tcb_states_of_state
where
"tcb_states_of_state s \<equiv> \<lambda>p. option_map tcb_state (get_tcb p s)"
(* FIXME: RENAME: should be named something thread_st_auth. thread_states is confusing as
it outputs autorities *)
definition thread_states
where
"thread_states s \<equiv> case_option {} tcb_st_to_auth \<circ> tcb_states_of_state s"
definition thread_bound_ntfns
where
"thread_bound_ntfns s \<equiv> \<lambda>p. case (get_tcb p s) of None \<Rightarrow> None
| Some tcb \<Rightarrow> tcb_bound_notification tcb"
definition state_objs_to_policy
where
"state_objs_to_policy s =
state_bits_to_policy (caps_of_state s) (thread_states s) (thread_bound_ntfns s)
(cdt s) (state_vrefs s)"
lemmas state_objs_to_policy_mem = eqset_imp_iff[OF state_objs_to_policy_def]
lemmas state_objs_to_policy_intros
= state_bits_to_policy.intros[THEN state_objs_to_policy_mem[THEN iffD2]]
(* FIXME this is non resilient at all to adding or removing rules to or from state_bits_to_policy *)
lemmas sta_caps = state_objs_to_policy_intros(1)
and sta_untyped = state_objs_to_policy_intros(2)
and sta_ts = state_objs_to_policy_intros(3)
and sta_bas = state_objs_to_policy_intros(4)
and sta_cdt = state_objs_to_policy_intros(5)
and sta_cdt_transferable = state_objs_to_policy_intros(6)
and sta_vref = state_objs_to_policy_intros(7)
lemmas state_objs_to_policy_cases
= state_bits_to_policy.cases[OF state_objs_to_policy_mem[THEN iffD1]]
end
context pspace_update_eq begin
interpretation Arch . (*FIXME: arch_split*)
lemma state_vrefs[iff]: "state_vrefs (f s) = state_vrefs s"
by (simp add: state_vrefs_def pspace)
lemma thread_states[iff]: "thread_states (f s) = thread_states s"
by (simp add: thread_states_def pspace get_tcb_def swp_def tcb_states_of_state_def)
lemma thread_bound_ntfns[iff]: "thread_bound_ntfns (f s) = thread_bound_ntfns s"
by (simp add: thread_bound_ntfns_def pspace get_tcb_def swp_def split: option.splits)
(*
lemma ipc_buffers_of_state[iff]: "ipc_buffers_of_state (f s) = ipc_buffers_of_state s"
by (simp add: ipc_buffers_of_state_def pspace get_tcb_def swp_def)
*)
end
context begin interpretation Arch . (*FIXME: arch_split*)
lemma tcb_states_of_state_preserved:
"\<lbrakk> get_tcb thread s = Some tcb; tcb_state tcb' = tcb_state tcb \<rbrakk>
\<Longrightarrow> tcb_states_of_state (s\<lparr>kheap := kheap s(thread \<mapsto> TCB tcb')\<rparr>) = tcb_states_of_state s"
apply rule
apply (auto split: option.splits simp: tcb_states_of_state_def get_tcb_def)
done
lemma thread_states_preserved:
"\<lbrakk> get_tcb thread s = Some tcb; tcb_state tcb' = tcb_state tcb \<rbrakk>
\<Longrightarrow> thread_states (s\<lparr>kheap := kheap s(thread \<mapsto> TCB tcb')\<rparr>) = thread_states s"
by (simp add: tcb_states_of_state_preserved thread_states_def)
lemma thread_bound_ntfns_preserved:
"\<lbrakk> get_tcb thread s = Some tcb; tcb_bound_notification tcb' = tcb_bound_notification tcb \<rbrakk>
\<Longrightarrow> thread_bound_ntfns (s\<lparr>kheap := kheap s(thread \<mapsto> TCB tcb')\<rparr>) = thread_bound_ntfns s"
by (auto simp: thread_bound_ntfns_def get_tcb_def split: option.splits)
(* FIXME: move all three *)
lemma null_filterI:
"\<lbrakk> cs p = Some cap; cap \<noteq> cap.NullCap \<rbrakk> \<Longrightarrow> null_filter cs p = Some cap"
unfolding null_filter_def by auto
lemma null_filterI2:
"\<lbrakk> cs p = None \<rbrakk> \<Longrightarrow> null_filter cs p = None"
unfolding null_filter_def by auto
lemma null_filterE2:
"\<lbrakk> null_filter cps x = None;
\<lbrakk> cps x = Some NullCap \<rbrakk> \<Longrightarrow> R ; \<lbrakk> cps x = None \<rbrakk> \<Longrightarrow> R \<rbrakk>
\<Longrightarrow> R"
by (simp add: null_filter_def split: if_split_asm)
lemma is_transferable_null_filter[simp]:
"is_transferable (null_filter caps ptr) = is_transferable (caps ptr)"
by (auto simp: is_transferable.simps null_filter_def)
lemma sbta_null_filter:
"state_bits_to_policy (null_filter cs) sr bar cd vr = state_bits_to_policy cs sr bar cd vr"
apply (rule subset_antisym)
apply clarsimp
apply (erule state_bits_to_policy.induct,
(fastforce intro: state_bits_to_policy.intros
elim: null_filterE null_filterE2 split: option.splits)+)[1]
apply clarsimp
apply (erule state_bits_to_policy.induct)
by (fastforce intro: state_bits_to_policy.intros null_filterI split:option.splits)+
lemma sata_null_filter:
"state_asids_to_policy_aux aag (null_filter cs) asid_tab vrefs
= state_asids_to_policy_aux aag cs asid_tab vrefs"
by (auto elim!: state_asids_to_policy_aux.induct null_filterE
intro: state_asids_to_policy_aux.intros sata_asid[OF null_filterI])
subsection \<open>Policy Refinement\<close>
text\<open>
We map scheduler domains to labels. This asserts that the labels on
tcbs are consistent with the labels on the domains they run in.
We need this to show that the ready queues are not reordered by
unauthorised subjects (see integrity_ready_queues).
\<close>
inductive_set domains_of_state_aux for ekheap
where
domtcbs: "\<lbrakk> ekheap ptr = Some etcb; d = tcb_domain etcb \<rbrakk>
\<Longrightarrow> (ptr, d) \<in> domains_of_state_aux ekheap"
abbreviation domains_of_state
where
"domains_of_state s \<equiv> domains_of_state_aux (ekheap s)"
definition tcb_domain_map_wellformed_aux
where
"tcb_domain_map_wellformed_aux aag etcbs_doms \<equiv>
\<forall>(ptr, d) \<in> etcbs_doms. pasObjectAbs aag ptr \<in> pasDomainAbs aag d"
abbreviation tcb_domain_map_wellformed
where
"tcb_domain_map_wellformed aag s \<equiv> tcb_domain_map_wellformed_aux aag (domains_of_state s)"
lemma tcb_domain_map_wellformed_mono:
"\<lbrakk> domains_of_state s' \<subseteq> domains_of_state s; tcb_domain_map_wellformed pas s \<rbrakk> \<Longrightarrow> tcb_domain_map_wellformed pas s'"
by (auto simp: tcb_domain_map_wellformed_aux_def get_etcb_def)
text\<open>
We sometimes need to know that our current subject may run in the current domain.
\<close>
abbreviation
"pas_cur_domain aag s \<equiv> pasSubject aag \<in> pasDomainAbs aag (cur_domain s)"
text\<open>
The relation we want to hold between the current state and
the policy:
\begin{itemize}
\item The policy should be well-formed.
\item The abstraction of the state should respect the policy.
\end{itemize}
\<close>
abbreviation state_objs_in_policy :: "'a PAS \<Rightarrow> det_ext state \<Rightarrow> bool"
where
"state_objs_in_policy aag s \<equiv>
auth_graph_map (pasObjectAbs aag) (state_objs_to_policy s) \<subseteq> pasPolicy aag"
definition pas_refined :: "'a PAS \<Rightarrow> det_ext state \<Rightarrow> bool"
where
"pas_refined aag \<equiv> \<lambda>s.
pas_wellformed aag
\<and> irq_map_wellformed aag s
\<and> tcb_domain_map_wellformed aag s
\<and> state_objs_in_policy aag s
\<and> state_asids_to_policy aag s \<subseteq> pasPolicy aag
\<and> state_irqs_to_policy aag s \<subseteq> pasPolicy aag"
lemma pas_refined_mem:
"\<lbrakk> (x, auth, y) \<in> state_objs_to_policy s; pas_refined aag s \<rbrakk>
\<Longrightarrow> (pasObjectAbs aag x, auth, pasObjectAbs aag y) \<in> pasPolicy aag"
by (auto simp: pas_refined_def intro: auth_graph_map_memI)
lemma pas_refined_wellformed[elim!]:
"pas_refined aag s \<Longrightarrow> pas_wellformed aag"
unfolding pas_refined_def by simp
lemmas pas_refined_Control
= aag_wellformed_Control[OF _ pas_refined_wellformed]
and pas_refined_refl = aag_wellformed_refl [OF pas_refined_wellformed]
lemma caps_of_state_pasObjectAbs_eq:
"\<lbrakk> caps_of_state s p = Some cap; Control \<in> cap_auth_conferred cap;
is_subject aag (fst p); pas_refined aag s;
x \<in> obj_refs cap \<rbrakk>
\<Longrightarrow> is_subject aag x"
apply (frule sta_caps, simp+)
apply (drule pas_refined_mem, simp+)
apply (drule pas_refined_Control, simp+)
done
lemma pas_refined_state_objs_to_policy_subset:
"\<lbrakk> pas_refined aag s;
state_objs_to_policy s' \<subseteq> state_objs_to_policy s;
state_asids_to_policy aag s' \<subseteq> state_asids_to_policy aag s;
state_irqs_to_policy aag s' \<subseteq> state_irqs_to_policy aag s;
domains_of_state s' \<subseteq> domains_of_state s;
interrupt_irq_node s' = interrupt_irq_node s
\<rbrakk>
\<Longrightarrow> pas_refined aag s'"
by (simp add: pas_refined_def) (blast dest: tcb_domain_map_wellformed_mono dest: auth_graph_map_mono[where G="pasObjectAbs aag"])
lemma aag_wellformed_all_auth_is_owns':
"\<lbrakk> Control \<in> S; pas_wellformed aag \<rbrakk> \<Longrightarrow> (\<forall>auth \<in> S. aag_has_auth_to aag auth x) = (is_subject aag x)"
apply (rule iffI)
apply (drule (1) bspec)
apply (frule (1) aag_wellformed_Control )
apply fastforce
apply (clarsimp simp: aag_wellformed_refl)
done
lemmas aag_wellformed_all_auth_is_owns
= aag_wellformed_all_auth_is_owns'[where S=UNIV, simplified]
aag_wellformed_all_auth_is_owns'[where S="{Control}", simplified]
lemmas aag_wellformed_control_is_owns
= aag_wellformed_all_auth_is_owns'[where S="{Control}", simplified]
lemmas pas_refined_all_auth_is_owns = aag_wellformed_all_auth_is_owns[OF pas_refined_wellformed]
lemma pas_refined_sita_mem:
"\<lbrakk> (x, auth, y) \<in> state_irqs_to_policy aag s; pas_refined aag s \<rbrakk>
\<Longrightarrow> (x, auth, y) \<in> pasPolicy aag"
by (auto simp: pas_refined_def)
section \<open>Integrity definition\<close>
subsection \<open>How kernel objects can change\<close>
fun blocked_on :: "obj_ref \<Rightarrow> Structures_A.thread_state \<Rightarrow> bool"
where
"blocked_on ref (Structures_A.BlockedOnReceive ref' _) = (ref = ref')"
| "blocked_on ref (Structures_A.BlockedOnSend ref' _) = (ref = ref')"
| "blocked_on ref (Structures_A.BlockedOnNotification ref') = (ref = ref')"
| "blocked_on _ _ = False"
lemma blocked_on_def2:
"blocked_on ref ts = (ref \<in> fst ` tcb_st_to_auth ts)"
by (cases ts, simp_all)
fun receive_blocked_on :: "obj_ref \<Rightarrow> Structures_A.thread_state \<Rightarrow> bool"
where
"receive_blocked_on ref (Structures_A.BlockedOnReceive ref' _) = (ref = ref')"
| "receive_blocked_on ref (Structures_A.BlockedOnNotification ref') = (ref = ref')"
| "receive_blocked_on _ _ = False"
lemma receive_blocked_on_def2:
"receive_blocked_on ref ts = ((ref, Receive) \<in> tcb_st_to_auth ts)"
by (cases ts, simp_all)
fun can_receive_ipc :: "Structures_A.thread_state \<Rightarrow> bool"
where
"can_receive_ipc (Structures_A.BlockedOnReceive _ _) = True"
| "can_receive_ipc (Structures_A.BlockedOnSend _ pl) =
(sender_is_call pl \<and> (sender_can_grant pl \<or> sender_can_grant_reply pl))"
| "can_receive_ipc (Structures_A.BlockedOnNotification _) = True"
| "can_receive_ipc Structures_A.BlockedOnReply = True"
| "can_receive_ipc _ = False"
lemma receive_blocked_on_can_receive_ipc[elim!,simp]:
"receive_blocked_on ref ts \<Longrightarrow>can_receive_ipc ts"
by (cases ts;simp)
lemma receive_blocked_on_can_receive_ipc'[elim!,simp]:
"case_option False (receive_blocked_on ref) tsopt \<Longrightarrow> case_option False can_receive_ipc tsopt"
by (cases tsopt;simp)
fun send_blocked_on :: "obj_ref \<Rightarrow> Structures_A.thread_state \<Rightarrow> bool"
where
"send_blocked_on ref (Structures_A.BlockedOnSend ref' _) = (ref = ref')"
| "send_blocked_on _ _ = False"
lemma send_blocked_on_def2:
"send_blocked_on ref ts = ((ref, SyncSend) \<in> tcb_st_to_auth ts)"
by (cases ts, simp_all)
lemma send_blocked_on_def3:
"send_blocked_on ref ts = (\<exists> pl. ts = BlockedOnSend ref pl)"
by (cases ts; force)
fun send_is_call :: "Structures_A.thread_state \<Rightarrow> bool"
where
"send_is_call (Structures_A.BlockedOnSend _ payl) = sender_is_call payl"
| "send_is_call _ = False"
definition tcb_bound_notification_reset_integrity
:: "obj_ref option \<Rightarrow> obj_ref option \<Rightarrow> 'a set \<Rightarrow> 'a PAS \<Rightarrow> bool"
where
"tcb_bound_notification_reset_integrity ntfn ntfn' subjects aag
= ((ntfn = ntfn') \<comment> \<open>no change to bound ntfn\<close>
\<or> (ntfn' = None \<and> aag_subjects_have_auth_to subjects aag Reset (the ntfn))
\<comment> \<open>ntfn is unbound\<close>)"
lemma tcb_bound_notification_reset_integrity_refl[simp]:
"tcb_bound_notification_reset_integrity ntfn ntfn subjects aag"
by (simp add: tcb_bound_notification_reset_integrity_def)
definition direct_send :: "'a set \<Rightarrow> 'a PAS \<Rightarrow> obj_ref \<Rightarrow> tcb \<Rightarrow> bool"
where
"direct_send subjects aag ep tcb \<equiv> receive_blocked_on ep (tcb_state tcb) \<and>
(aag_subjects_have_auth_to subjects aag SyncSend ep
\<or> aag_subjects_have_auth_to subjects aag Notify ep)"
abbreviation ep_recv_blocked :: "obj_ref \<Rightarrow> thread_state \<Rightarrow> bool"
where
"ep_recv_blocked ep ts \<equiv> case ts of BlockedOnReceive w _ \<Rightarrow> w = ep
| _ \<Rightarrow> False"
definition direct_call :: "'a set \<Rightarrow> 'a PAS \<Rightarrow> obj_ref \<Rightarrow> thread_state \<Rightarrow> bool"
where
"direct_call subjects aag ep tcbst \<equiv> ep_recv_blocked ep (tcbst) \<and>
aag_subjects_have_auth_to subjects aag Call ep"
definition indirect_send :: "'a set \<Rightarrow> 'a PAS \<Rightarrow> obj_ref \<Rightarrow> obj_ref \<Rightarrow> tcb \<Rightarrow> bool"
where
"indirect_send subjects aag ntfn recv_ep tcb \<equiv> ep_recv_blocked recv_ep (tcb_state tcb)
\<comment> \<open>tcb is blocked on sync ep\<close>
\<and> (tcb_bound_notification tcb = Some ntfn)
\<and> aag_subjects_have_auth_to subjects aag Notify ntfn"
definition call_blocked :: "obj_ref \<Rightarrow> thread_state \<Rightarrow> bool"
where
"call_blocked ep tst \<equiv> \<exists> pl. tst = BlockedOnSend ep pl \<and> sender_is_call pl"
definition allowed_call_blocked :: "obj_ref \<Rightarrow> thread_state \<Rightarrow> bool"
where
"allowed_call_blocked ep tst \<equiv> \<exists> pl. tst = BlockedOnSend ep pl \<and> sender_is_call pl \<and>