Current_prop.thy
author chunhan
Mon, 24 Jun 2013 15:22:37 +0800
changeset 26 b6333712cb02
parent 25 259a50be4381
child 27 fc749f19b894
permissions -rw-r--r--
finished info_flow_shm(simple def) simpset
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
     1
(*<*)
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     2
theory Current_prop
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     3
imports Main Flask_type Flask My_list_prefix Init_prop Valid_prop Delete_prop
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     4
begin
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
     5
(*>*)
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     6
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     7
context flask begin
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     8
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     9
lemma procs_of_shm_prop1: "\<lbrakk> p_flag \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> h \<in> current_shms s"
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    10
apply (induct s arbitrary:p_flag)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    11
apply (case_tac p_flag, simp, drule init_procs_has_shm, simp)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    12
apply (frule vd_cons, frule vt_grant_os)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    13
apply (case_tac a, auto split:if_splits option.splits)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    14
done
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    15
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    16
lemma procs_of_shm_prop2: "\<lbrakk>(p, flag) \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s"
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    17
apply (induct s arbitrary:p flag)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    18
apply (simp, drule init_procs_has_shm, simp)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    19
apply (frule vd_cons, frule vt_grant_os)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    20
apply (case_tac a, auto split:if_splits option.splits)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    21
done
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    22
26
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    23
lemma procs_of_shm_prop2':
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    24
  "\<lbrakk>p \<notin> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<forall> flag h. (p, flag) \<notin> procs_of_shm s h"
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    25
by (auto dest:procs_of_shm_prop2)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    26
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    27
lemma procs_of_shm_prop3: "\<lbrakk>(p, flag) \<in> procs_of_shm s h; (p, flag') \<in> procs_of_shm s h; valid s\<rbrakk>
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    28
  \<Longrightarrow> flag = flag'"
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    29
apply (induct s arbitrary:p flag flag')
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    30
apply (simp, drule_tac flag = flag in init_procs_has_shm, drule_tac flag = flag' in init_procs_has_shm, simp)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    31
apply (frule vd_cons, frule vt_grant_os)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    32
apply (case_tac a, auto split:if_splits option.splits dest:procs_of_shm_prop2)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    33
done
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    34
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    35
lemma procs_of_shm_prop4: "\<lbrakk>(p, flag) \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> flag_of_proc_shm s p h = Some flag"
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    36
apply (induct s arbitrary:p flag)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    37
apply (simp, drule init_procs_has_shm, simp)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    38
apply (frule vd_cons, frule vt_grant_os)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    39
apply (case_tac a, auto split:if_splits option.splits dest:procs_of_shm_prop2)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    40
done
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    41
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    42
lemma procs_of_shm_prop4':
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    43
  "\<lbrakk>flag_of_proc_shm s p h = None; valid s\<rbrakk> \<Longrightarrow> \<forall> flag. (p, flag) \<notin> procs_of_shm s h"
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    44
by (auto dest:procs_of_shm_prop4)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    45
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    46
lemma not_init_intro_proc:
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    47
  "\<lbrakk>p \<notin> current_procs s; valid s\<rbrakk> \<Longrightarrow> deleted (O_proc p) s \<or> p \<notin> init_procs"
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    48
using not_deleted_init_proc by auto
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    49
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    50
lemma not_init_intro_proc':
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    51
  "\<lbrakk>p \<notin> current_procs s; valid s\<rbrakk> \<Longrightarrow> \<not> (\<not> deleted (O_proc p) s \<and> p \<in> init_procs)"
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    52
using not_deleted_init_proc by auto
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    53
23
25e55731ed01 locale of tainting for seeds when same shm/inode bugs
chunhan
parents: 19
diff changeset
    54
lemma info_shm_flow_in_procs:
25e55731ed01 locale of tainting for seeds when same shm/inode bugs
chunhan
parents: 19
diff changeset
    55
  "\<lbrakk>info_flow_shm s p p'; valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s \<and> p' \<in> current_procs s"
26
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    56
by (auto intro:procs_of_shm_prop2 simp:info_flow_shm_def one_flow_shm_def)
25
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
    57
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
    58
(*********** simpset for info_flow_shm **************)
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
    59
26
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    60
lemma info_flow_shm_attach:
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    61
  "valid (Attach p h flag # s) \<Longrightarrow> info_flow_shm (Attach p h flag # s) = (\<lambda> pa pb. 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    62
     (pa = p \<and> flag = SHM_RDWR \<and> (\<exists> flagb. (pb, flagb) \<in> procs_of_shm s h)) \<or>
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    63
     (pb = p \<and> (pa, SHM_RDWR) \<in> procs_of_shm s h) \<or>
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    64
     (info_flow_shm s pa pb)                       )"
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    65
apply (rule ext, rule ext, frule vt_grant_os)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    66
by (auto simp add:info_flow_shm_def one_flow_shm_def)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    67
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    68
lemma info_flow_shm_detach:
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    69
  "valid (Detach p h # s) \<Longrightarrow> info_flow_shm (Detach p h # s) = (\<lambda> pa pb. 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    70
     self_shm s pa pb \<or> ((p = pa \<or> p = pb) \<and> (\<exists> h'. h' \<noteq> h \<and> one_flow_shm s h' pa pb)) \<or>
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    71
     (pa \<noteq> p \<and> pb \<noteq> p \<and> info_flow_shm s pa pb) )"
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    72
apply (rule ext, rule ext, frule vt_grant_os)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    73
by (auto simp:info_flow_shm_def one_flow_shm_def)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    74
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    75
lemma info_flow_shm_deleteshm:
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    76
  "valid (DeleteShM p h # s) \<Longrightarrow> info_flow_shm (DeleteShM p h # s) = (\<lambda> pa pb. 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    77
     self_shm s pa pb \<or> (\<exists> h'. h' \<noteq> h \<and> one_flow_shm s h' pa pb)     )"
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    78
apply (rule ext, rule ext, frule vt_grant_os)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    79
by (auto simp:info_flow_shm_def one_flow_shm_def)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    80
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    81
lemma info_flow_shm_clone:
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    82
  "valid (Clone p p' fds shms # s) \<Longrightarrow> info_flow_shm (Clone p p' fds shms # s) = (\<lambda> pa pb. 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    83
     (pa = p' \<and> pb = p') \<or> (pa = p' \<and> pb \<noteq> p' \<and> (\<exists> h \<in> shms. one_flow_shm s h p pb)) \<or> 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    84
     (pb = p' \<and> pa \<noteq> p' \<and> (\<exists> h \<in> shms. one_flow_shm s h pa p)) \<or> 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    85
     (pa \<noteq> p' \<and> pb \<noteq> p' \<and> info_flow_shm s pa pb))"
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    86
apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons, clarsimp)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    87
apply (frule_tac p = p' in procs_of_shm_prop2', simp)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    88
apply (auto simp:info_flow_shm_def one_flow_shm_def)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    89
done
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    90
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    91
lemma info_flow_shm_execve:
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    92
  "valid (Execve p f fds # s) \<Longrightarrow> info_flow_shm (Execve p f fds # s) = (\<lambda> pa pb. 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    93
     (pa = p \<and> pb = p) \<or> (pa \<noteq> p \<and> pb \<noteq> p \<and> info_flow_shm s pa pb)    )"
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    94
apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    95
by (auto simp:info_flow_shm_def one_flow_shm_def)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    96
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    97
lemma info_flow_shm_kill:
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    98
  "valid (Kill p p' # s) \<Longrightarrow> info_flow_shm (Kill p p' # s) = (\<lambda> pa pb. 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    99
     pa \<noteq> p' \<and> pb \<noteq> p' \<and> info_flow_shm s pa pb                 )"
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   100
apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   101
by (auto simp:info_flow_shm_def one_flow_shm_def)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   102
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   103
lemma info_flow_shm_exit:
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   104
  "valid (Exit p # s) \<Longrightarrow> info_flow_shm (Exit p # s) = (\<lambda> pa pb. 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   105
     pa \<noteq> p \<and> pb \<noteq> p \<and> info_flow_shm s pa pb                          )"
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   106
apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   107
by (auto simp:info_flow_shm_def one_flow_shm_def)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   108
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   109
lemma info_flow_shm_other:
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   110
  "\<lbrakk>valid (e # s); 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   111
    \<forall> p h flag. e \<noteq> Attach p h flag;
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   112
    \<forall> p h. e \<noteq> Detach p h;
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   113
    \<forall> p h. e \<noteq> DeleteShM p h;
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   114
    \<forall> p p' fds shms. e \<noteq> Clone p p' fds shms;
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   115
    \<forall> p f fds. e \<noteq> Execve p f fds;
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   116
    \<forall> p p'. e \<noteq> Kill p p';
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   117
    \<forall> p. e \<noteq> Exit p
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   118
   \<rbrakk> \<Longrightarrow> info_flow_shm (e # s) = info_flow_shm s"
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   119
apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   120
apply (case_tac e, auto simp:info_flow_shm_def one_flow_shm_def dest:procs_of_shm_prop2)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   121
apply (erule_tac x = h in allE, simp)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   122
apply (drule procs_of_shm_prop1, auto)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   123
done
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   124
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   125
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   126
(*
25
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   127
lemma info_flow_shm_prop1: 
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   128
  "\<lbrakk>info_flow_shm s p p'; p \<noteq> p'; valid s\<rbrakk> 
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   129
   \<Longrightarrow> \<exists> h h' flag. (p, SHM_RDWR) \<in> procs_of_shm s h \<and> (p', flag) \<in> procs_of_shm s h'"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   130
by (induct rule: info_flow_shm.induct, auto)
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   131
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   132
lemma info_flow_shm_cases:
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   133
  "\<lbrakk>info_flow_shm \<tau> pa pb; \<And>p s. \<lbrakk>s = \<tau> ; pa = p; pb = p; p \<in> current_procs s\<rbrakk> \<Longrightarrow> P;
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   134
  \<And>s p p' h p'' flag. \<lbrakk>s = \<tau>; pa = p; pb = p''; info_flow_shm s p p'; (p', SHM_RDWR) \<in> procs_of_shm s h;
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   135
                       (p'', flag) \<in> procs_of_shm s h\<rbrakk>\<Longrightarrow> P\<rbrakk>
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   136
  \<Longrightarrow> P"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   137
by (erule info_flow_shm.cases, auto)
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   138
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   139
definition one_flow_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   140
where
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   141
  "one_flow_shm s p p' \<equiv> p \<noteq> p' \<and> (\<exists> h flag. (p, SHM_RDWR) \<in> procs_of_shm s h \<and> (p', flag) \<in> procs_of_shm s h)"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   142
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   143
inductive flows_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   144
where
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   145
  "p \<in> current_procs s \<Longrightarrow> flows_shm s p p"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   146
| "\<lbrakk>flows_shm s p p'; one_flow_shm s p' p''\<rbrakk> \<Longrightarrow> flows_shm s p p''"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   147
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   148
definition attached_procs :: "t_state \<Rightarrow> t_shm \<Rightarrow> t_process set"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   149
where
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   150
  "attached_procs s h \<equiv> {p. \<exists> flag. (p, flag) \<in> procs_of_shm s h}"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   151
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   152
definition flowed_procs:: "t_state \<Rightarrow> t_shm \<Rightarrow> t_process set"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   153
where
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   154
  "flowed_procs s h \<equiv> {p'. \<exists> p \<in> attached_procs s h. flows_shm s p p'}"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   155
26
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   156
inductive flowed_shm:: "t_state \<Rightarrow> t_process \<Rightarrow> t_shm set"
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   157
25
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   158
fun Info_flow_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process set"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   159
where
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   160
  "Info_flow_shm [] = (\<lambda> p. {p'. flows_shm [] p p'})"
26
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   161
| "Info_flow_shm (Attach p h flag # s) = (\<lambda> p'. 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   162
     if (p' = p) then flowed_procs s h 
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   163
     else if ()
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   164
    "
25
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   165
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   166
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   167
lemma info_flow_shm_attach:
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   168
  "valid (Attach p h flag # s) \<Longrightarrow> info_flow_shm (Attach p h flag # s) = (\<lambda> pa pb. (info_flow_shm s pa pb) \<or> 
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   169
     (if (pa = p) 
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   170
      then (if (flag = SHM_RDWR) 
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   171
            then (\<exists> flag. (pb, flag) \<in> procs_of_shm s h)
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   172
            else (pb = p)) 
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   173
      else (if (pb = p) 
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   174
            then (pa, SHM_RDWR) \<in> procs_of_shm s h
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   175
            else info_flow_shm s pa pb)) )"
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   176
apply (frule vd_cons, frule vt_grant_os, rule ext, rule ext)
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   177
apply (case_tac "info_flow_shm s pa pb", simp)
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   178
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   179
thm info_flow_shm.cases
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   180
apply (auto split:if_splits intro:info_flow_shm.intros elim:info_flow_shm_cases)
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   181
apply (erule info_flow_shm_cases, simp, simp split:if_splits)
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   182
apply (rule_tac p = pa and p' = p' in info_flow_shm.intros(2), simp+)
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   183
apply (rule notI, erule info_flow_shm.cases, simp+)
259a50be4381 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt
chunhan
parents: 23
diff changeset
   184
pr 5
26
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   185
*)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   186
lemmas info_flow_shm_simps = info_flow_shm_other info_flow_shm_attach info_flow_shm_detach info_flow_shm_deleteshm
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
   187
  info_flow_shm_clone info_flow_shm_execve info_flow_shm_kill info_flow_shm_exit
23
25e55731ed01 locale of tainting for seeds when same shm/inode bugs
chunhan
parents: 19
diff changeset
   188
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   189
lemma has_same_inode_in_current:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   190
  "\<lbrakk>has_same_inode s f f'; valid s\<rbrakk> \<Longrightarrow> f \<in> current_files s \<and> f' \<in> current_files s"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   191
by (auto simp add:has_same_inode_def current_files_def)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   192
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   193
lemma has_same_inode_prop1:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   194
  "\<lbrakk>has_same_inode s f f'; is_file s f; valid s\<rbrakk> \<Longrightarrow> is_file s f'"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   195
by (auto simp:has_same_inode_def is_file_def)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   196
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   197
lemma has_same_inode_prop1':
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   198
  "\<lbrakk>has_same_inode s f f'; is_file s f'; valid s\<rbrakk> \<Longrightarrow> is_file s f"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   199
by (auto simp:has_same_inode_def is_file_def)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   200
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   201
lemma has_same_inode_prop2:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   202
  "\<lbrakk>has_same_inode s f f'; file_of_proc_fd s p fd = Some f; valid s\<rbrakk> \<Longrightarrow> is_file s f'"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   203
apply (drule has_same_inode_prop1)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   204
apply (simp add:file_of_pfd_is_file, simp+)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   205
done
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   206
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   207
lemma has_same_inode_prop2':
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   208
  "\<lbrakk>has_same_inode s f f'; file_of_proc_fd s p fd = Some f'; valid s\<rbrakk> \<Longrightarrow> is_file s f"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   209
apply (drule has_same_inode_prop1')
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   210
apply (simp add:file_of_pfd_is_file, simp+)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   211
done
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   212
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   213
lemma tobj_in_init_alive:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   214
  "tobj_in_init obj \<Longrightarrow> init_alive obj"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   215
by (case_tac obj, auto)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   216
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   217
lemma tobj_in_alive:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   218
  "tobj_in_init obj \<Longrightarrow> alive [] obj"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   219
by (case_tac obj, auto simp:is_file_nil)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   220
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   221
end
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   222
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   223
end