Current_prop.thy
author chunhan
Mon, 05 Aug 2013 12:30:26 +0800
changeset 30 01274a64aece
parent 29 622516c0fe34
child 34 e7f850d1e08e
permissions -rw-r--r--
move info_flow/path_by/shm into a new thy
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
29
622516c0fe34 path_by_shm reconstrain path without duplicated process AND shm
chunhan
parents: 28
diff changeset
     7
ML {*quick_and_dirty := true*}
622516c0fe34 path_by_shm reconstrain path without duplicated process AND shm
chunhan
parents: 28
diff changeset
     8
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     9
context flask begin
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    10
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    11
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
    12
apply (induct s arbitrary:p_flag)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    13
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
    14
apply (frule vd_cons, frule vt_grant_os)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    15
apply (case_tac a, auto split:if_splits option.splits)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    16
done
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    17
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    18
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
    19
apply (induct s arbitrary:p flag)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    20
apply (simp, drule init_procs_has_shm, simp)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    21
apply (frule vd_cons, frule vt_grant_os)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    22
apply (case_tac a, auto split:if_splits option.splits)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    23
done
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    24
26
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    25
lemma procs_of_shm_prop2':
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    26
  "\<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
    27
by (auto dest:procs_of_shm_prop2)
b6333712cb02 finished info_flow_shm(simple def) simpset
chunhan
parents: 25
diff changeset
    28
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    29
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
    30
  \<Longrightarrow> flag = flag'"
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    31
apply (induct s arbitrary:p flag flag')
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    32
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
    33
apply (frule vd_cons, frule vt_grant_os)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    34
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
    35
done
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    36
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    37
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
    38
apply (induct s arbitrary:p flag)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    39
apply (simp, drule init_procs_has_shm, simp)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    40
apply (frule vd_cons, frule vt_grant_os)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    41
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
    42
done
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    43
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    44
lemma procs_of_shm_prop4':
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    45
  "\<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
    46
by (auto dest:procs_of_shm_prop4)
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    47
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    48
lemma not_init_intro_proc:
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    49
  "\<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
    50
using not_deleted_init_proc by auto
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    51
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    52
lemma not_init_intro_proc':
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    53
  "\<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
    54
using not_deleted_init_proc by auto
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
    55
23
25e55731ed01 locale of tainting for seeds when same shm/inode bugs
chunhan
parents: 19
diff changeset
    56
lemma info_shm_flow_in_procs:
25e55731ed01 locale of tainting for seeds when same shm/inode bugs
chunhan
parents: 19
diff changeset
    57
  "\<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
    58
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
    59
27
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 26
diff changeset
    60
lemma flag_of_proc_shm_prop1:
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 26
diff changeset
    61
  "\<lbrakk>flag_of_proc_shm s p h = Some flag; valid s\<rbrakk> \<Longrightarrow> (p, flag) \<in> procs_of_shm s h"
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 26
diff changeset
    62
apply (induct s arbitrary:p flag)
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 26
diff changeset
    63
apply (simp, drule init_shmflag_has_proc, simp)
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 26
diff changeset
    64
apply (frule vd_cons, frule vt_grant_os)
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 26
diff changeset
    65
apply (case_tac a, auto split:if_splits option.splits dest:procs_of_shm_prop2)
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 26
diff changeset
    66
done
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 26
diff changeset
    67
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 26
diff changeset
    68
23
25e55731ed01 locale of tainting for seeds when same shm/inode bugs
chunhan
parents: 19
diff changeset
    69
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    70
lemma has_same_inode_in_current:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    71
  "\<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
    72
by (auto simp add:has_same_inode_def current_files_def)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    73
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    74
lemma has_same_inode_prop1:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    75
  "\<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
    76
by (auto simp:has_same_inode_def is_file_def)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    77
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    78
lemma has_same_inode_prop1':
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    79
  "\<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
    80
by (auto simp:has_same_inode_def is_file_def)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    81
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    82
lemma has_same_inode_prop2:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    83
  "\<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
    84
apply (drule has_same_inode_prop1)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    85
apply (simp add:file_of_pfd_is_file, simp+)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    86
done
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    87
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    88
lemma has_same_inode_prop2':
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    89
  "\<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
    90
apply (drule has_same_inode_prop1')
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    91
apply (simp add:file_of_pfd_is_file, simp+)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    92
done
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    93
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    94
lemma tobj_in_init_alive:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    95
  "tobj_in_init obj \<Longrightarrow> init_alive obj"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    96
by (case_tac obj, auto)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    97
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    98
lemma tobj_in_alive:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
    99
  "tobj_in_init obj \<Longrightarrow> alive [] obj"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   100
by (case_tac obj, auto simp:is_file_nil)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   101
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   102
end
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   103
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   104
end