| author | chunhan | 
| Tue, 07 Jan 2014 22:04:06 +0800 | |
| changeset 90 | 003cac7b8bf5 | 
| parent 74 | 271e9818b6f6 | 
| permissions | -rwxr-xr-x | 
| 74 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 1 | theory Static | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 2 | imports Static_type OS_type_def Flask_type Flask | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 3 | begin | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 4 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 5 | locale tainting_s = tainting | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 6 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 7 | begin | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 8 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 9 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 10 | definition init_sectxt_of_file:: "t_file \<Rightarrow> security_context_t option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 11 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 12 | "init_sectxt_of_file f \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 13 | if (is_init_file f) then init_sectxt_of_obj (O_file f) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 14 | else if (is_init_dir f) then init_sectxt_of_obj (O_dir f) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 15 | else None" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 16 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 17 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 18 | definition sroot :: "t_sfile" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 19 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 20 |   "sroot \<equiv> (Init [], sec_of_root, None, {})"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 21 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 22 | definition init_cf2sfile :: "t_file \<Rightarrow> t_sfile option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 23 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 24 | "init_cf2sfile f \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 25 | case (parent f) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 26 | None \<Rightarrow> Some sroot | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 27 | | Some pf \<Rightarrow> if (is_init_file f) then | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 28 | (case (init_sectxt_of_obj (O_file f), init_sectxt_of_obj (O_dir pf), get_parentfs_ctxts [] pf) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 29 | (Some sec, Some psec, Some aseclist) \<Rightarrow> Some (Init f, sec, Some psec, set aseclist) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 30 | | _ \<Rightarrow> None) else | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 31 | (case (init_sectxt_of_obj (O_dir f), init_sectxt_of_obj (O_dir pf), get_parentfs_ctxts [] pf) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 32 | (Some sec, Some psec, Some aseclist) \<Rightarrow> Some (Init f, sec, Some psec, set aseclist) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 33 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 34 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 35 | definition init_cf2sfiles :: "t_file \<Rightarrow> t_sfile set" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 36 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 37 |   "init_cf2sfiles f \<equiv> {sf. \<exists>f' \<in> init_same_inode_files f. init_cf2sfile f' = Some sf}"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 38 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 39 | definition init_cfd2sfd :: "t_process \<Rightarrow> t_fd \<Rightarrow> (security_context_t \<times> t_open_flags \<times> t_sfile) option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 40 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 41 | "init_cfd2sfd p fd = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 42 | (case (init_file_of_proc_fd p fd, init_oflags_of_proc_fd p fd, init_sectxt_of_obj (O_fd p fd)) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 43 | (Some f, Some flags, Some sec) \<Rightarrow> (case (init_cf2sfile f) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 44 | Some sf \<Rightarrow> Some (sec, flags, sf) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 45 | | _ \<Rightarrow> None) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 46 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 47 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 48 | definition init_cfds2sfds :: "t_process \<Rightarrow> (security_context_t \<times> t_open_flags \<times> t_sfile) set" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 49 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 50 |   "init_cfds2sfds p \<equiv> { sfd. \<exists> fd \<in> init_proc_file_fds p. init_cfd2sfd p fd = Some sfd}"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 51 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 52 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 53 | definition init_ch2sshm :: "t_shm \<Rightarrow> t_sshm option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 54 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 55 | "init_ch2sshm h \<equiv> (case (init_sectxt_of_obj (O_shm h)) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 56 | Some sec \<Rightarrow> Some (Init h, sec) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 57 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 58 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 59 | definition init_cph2spshs | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 60 | :: "t_process \<Rightarrow> (t_sshm \<times> t_shm_attach_flag) set" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 61 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 62 |   "init_cph2spshs p \<equiv> {(sh, flag)| sh flag h. (p, flag) \<in> init_procs_of_shm h \<and> 
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 63 | init_ch2sshm h = Some sh}" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 64 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 65 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 66 | definition init_cp2sproc :: "t_process \<Rightarrow> t_sproc option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 67 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 68 | "init_cp2sproc p \<equiv> (case (init_sectxt_of_obj (O_proc p)) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 69 | Some sec \<Rightarrow> Some (Init p, sec, (init_cfds2sfds p)) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 70 | | None \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 71 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 72 | (* acient files of init-file | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 73 | definition init_o2s_afs :: "t_file \<Rightarrow> security_context_t set" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 74 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 75 |   "init_o2s_afs f \<equiv> {sec. \<exists> f'. f' \<preceq> f \<and> init_sectxt_of_obj (O_dir f') = Some sec}" *)
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 76 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 77 | definition init_cm2smsg :: "t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsg option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 78 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 79 | "init_cm2smsg q m \<equiv> (case (init_sectxt_of_obj (O_msg q m)) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 80 | Some sec \<Rightarrow> Some (Init m, sec, (O_msg q m) \<in> seeds) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 81 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 82 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 83 | fun init_cqm2sms :: "t_msgq \<Rightarrow> t_msg list \<Rightarrow> (t_smsg list) option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 84 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 85 | "init_cqm2sms q [] = Some []" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 86 | | "init_cqm2sms q (m#ms) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 87 | (case (init_cqm2sms q ms, init_cm2smsg q m) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 88 | (Some sms, Some sm) \<Rightarrow> Some (sm # sms) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 89 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 90 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 91 | definition init_cq2smsgq :: "t_msgq \<Rightarrow> t_smsgq option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 92 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 93 | "init_cq2smsgq q \<equiv> (case (init_sectxt_of_obj (O_msgq q), init_cqm2sms q (init_msgs_of_queue q)) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 94 | (Some sec, Some sms) \<Rightarrow> Some (Init q, sec, sms) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 95 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 96 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 97 | fun init_obj2sobj :: "t_object \<Rightarrow> t_sobject option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 98 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 99 | "init_obj2sobj (O_proc p) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 100 | (case (init_cp2sproc p) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 101 | Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> seeds)) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 102 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 103 | | "init_obj2sobj (O_file f) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 104 | (case (init_cf2sfile f) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 105 | Some sf \<Rightarrow> Some (S_file sf (O_file f \<in> seeds)) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 106 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 107 | | "init_obj2sobj (O_dir f) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 108 | (case (init_cf2sfile f) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 109 | Some sf \<Rightarrow> Some (S_dir sf) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 110 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 111 | | "init_obj2sobj (O_msgq q) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 112 | (case (init_cq2smsgq q) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 113 | Some sq \<Rightarrow> Some (S_msgq sq) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 114 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 115 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 116 | | "init_obj2sobj (O_shm h) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 117 | (case (init_ch2sshm h) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 118 | Some sh \<Rightarrow> Some (S_shm sh) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 119 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 120 | | "init_obj2sobj (O_msg q m) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 121 | (case (init_cq2smsgq q, init_cm2smsg q m) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 122 | (Some sq, Some sm) \<Rightarrow> Some (S_msg sq sm) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 123 | | _ \<Rightarrow> None)" *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 124 | | "init_obj2sobj _ = None" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 125 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 126 | definition | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 127 |   "init_static_state \<equiv> {sobj. \<exists> obj. init_alive obj \<and> init_obj2sobj obj = Some sobj}"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 128 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 129 | (**************** translation from dynamic to static *******************) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 130 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 131 | definition cf2sfile :: "t_state \<Rightarrow> t_file \<Rightarrow> t_sfile option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 132 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 133 | "cf2sfile s f \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 134 | case (parent f) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 135 | None \<Rightarrow> Some sroot | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 136 | | Some pf \<Rightarrow> if (is_file s f) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 137 | then (case (sectxt_of_obj s (O_file f), sectxt_of_obj s (O_dir pf), get_parentfs_ctxts s pf) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 138 | (Some sec, Some psec, Some asecs) \<Rightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 139 | Some (if (\<not> deleted (O_file f) s \<and> is_init_file f) then Init f else Created, sec, Some psec, set asecs) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 140 | | _ \<Rightarrow> None) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 141 | else (case (sectxt_of_obj s (O_dir f), sectxt_of_obj s (O_dir pf), get_parentfs_ctxts s pf) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 142 | (Some sec, Some psec, Some asecs) \<Rightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 143 | Some (if (\<not> deleted (O_dir f) s \<and> is_init_dir f) then Init f else Created, sec, Some psec, set asecs) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 144 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 145 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 146 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 147 | definition cf2sfiles :: "t_state \<Rightarrow> t_file \<Rightarrow> t_sfile set" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 148 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 149 |   "cf2sfiles s f \<equiv> {sf. \<exists> f' \<in> (same_inode_files s f). cf2sfile s f' = Some sf}"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 150 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 151 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 152 | (* here cf2sfile is passed with True, because, process' fds are only for files not dirs *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 153 | definition cfd2sfd :: "t_state \<Rightarrow> t_process \<Rightarrow> t_fd \<Rightarrow> t_sfd option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 154 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 155 | "cfd2sfd s p fd \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 156 | (case (file_of_proc_fd s p fd, flags_of_proc_fd s p fd, sectxt_of_obj s (O_fd p fd)) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 157 | (Some f, Some flags, Some sec) \<Rightarrow> (case (cf2sfile s f) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 158 | Some sf \<Rightarrow> Some (sec, flags, sf) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 159 | | _ \<Rightarrow> None) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 160 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 161 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 162 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 163 | definition cpfd2sfds :: "t_state \<Rightarrow> t_process \<Rightarrow> t_sfd set" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 164 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 165 |   "cpfd2sfds s p \<equiv> {sfd. \<exists> fd \<in> proc_file_fds s p. cfd2sfd s p fd = Some sfd}"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 166 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 167 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 168 | definition ch2sshm :: "t_state \<Rightarrow> t_shm \<Rightarrow> t_sshm option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 169 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 170 | "ch2sshm s h \<equiv> (case (sectxt_of_obj s (O_shm h)) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 171 | Some sec \<Rightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 172 | Some (if (\<not> deleted (O_shm h) s \<and> h \<in> init_shms) then Init h else Created, sec) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 173 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 174 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 175 | definition cph2spshs :: "t_state \<Rightarrow> t_process \<Rightarrow> t_sproc_sshm set" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 176 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 177 |   "cph2spshs s p \<equiv> {(sh, flag)| sh flag h. (p, flag) \<in> procs_of_shm s h \<and> ch2sshm s h = Some sh}"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 178 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 179 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 180 | definition cp2sproc :: "t_state \<Rightarrow> t_process \<Rightarrow> t_sproc option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 181 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 182 | "cp2sproc s p \<equiv> (case (sectxt_of_obj s (O_proc p)) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 183 | Some sec \<Rightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 184 | Some (if (\<not> deleted (O_proc p) s \<and> p \<in> init_procs) then Init p else Created, sec, | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 185 | cpfd2sfds s p) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 186 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 187 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 188 | definition cm2smsg :: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsg option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 189 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 190 | "cm2smsg s q m \<equiv> (case (sectxt_of_obj s (O_msg q m)) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 191 | Some sec \<Rightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 192 | Some (if (\<not> deleted (O_msg q m) s \<and> m \<in> set (init_msgs_of_queue q)) then Init m else Created, | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 193 | sec, O_msg q m \<in> tainted s) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 194 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 195 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 196 | fun cqm2sms:: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_msg list \<Rightarrow> (t_smsg list) option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 197 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 198 | "cqm2sms s q [] = Some []" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 199 | | "cqm2sms s q (m#ms) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 200 | (case (cqm2sms s q ms, cm2smsg s q m) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 201 | (Some sms, Some sm) \<Rightarrow> Some (sm#sms) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 202 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 203 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 204 | definition cq2smsgq :: "t_state \<Rightarrow> t_msgq \<Rightarrow> t_smsgq option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 205 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 206 | "cq2smsgq s q \<equiv> (case (sectxt_of_obj s (O_msgq q), cqm2sms s q (msgs_of_queue s q)) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 207 | (Some sec, Some sms) \<Rightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 208 | Some (if (\<not> deleted (O_msgq q) s \<and> q \<in> init_msgqs) then Init q else Created, | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 209 | sec, sms) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 210 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 211 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 212 | fun co2sobj :: "t_state \<Rightarrow> t_object \<Rightarrow> t_sobject option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 213 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 214 | "co2sobj s (O_proc p) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 215 | (case (cp2sproc s p) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 216 | Some sp \<Rightarrow> Some (S_proc sp (O_proc p \<in> tainted s)) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 217 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 218 | | "co2sobj s (O_file f) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 219 | (case (cf2sfile s f) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 220 | Some sf \<Rightarrow> Some (S_file sf (O_file f \<in> tainted s)) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 221 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 222 | | "co2sobj s (O_dir f) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 223 | (case (cf2sfile s f) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 224 | Some sf \<Rightarrow> Some (S_dir sf) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 225 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 226 | | "co2sobj s (O_msgq q) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 227 | (case (cq2smsgq s q) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 228 | Some sq \<Rightarrow> Some (S_msgq sq) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 229 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 230 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 231 | | "co2sobj s (O_shm h) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 232 | (case (ch2sshm s h) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 233 | Some sh \<Rightarrow> Some (S_shm sh) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 234 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 235 | | "co2sobj s (O_msg q m) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 236 | (case (cq2smsgq s q, cm2smsg s q m) of | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 237 | (Some sq, Some sm) \<Rightarrow> Some (S_msg sq sm) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 238 | | _ \<Rightarrow> None)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 239 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 240 | | "co2sobj s _ = None" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 241 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 242 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 243 | definition s2ss :: "t_state \<Rightarrow> t_static_state" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 244 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 245 |   "s2ss s \<equiv> {sobj. \<exists> obj. alive s obj \<and> co2sobj s obj = Some sobj}"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 246 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 247 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 248 | (* ******************************************************** *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 249 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 250 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 251 | fun is_init_sfile :: "t_sfile \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 252 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 253 | "is_init_sfile (Init _, sec, psec,asec) = True" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 254 | | "is_init_sfile _ = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 255 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 256 | fun is_many_sfile :: "t_sfile \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 257 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 258 | "is_many_sfile (Created, sec, psec, asec) = True" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 259 | | "is_many_sfile _ = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 260 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 261 | fun is_init_sproc :: "t_sproc \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 262 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 263 | "is_init_sproc (Init p, sec, fds) = True" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 264 | | "is_init_sproc _ = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 265 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 266 | fun is_many_sproc :: "t_sproc \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 267 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 268 | "is_many_sproc (Created, sec,fds) = True" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 269 | | "is_many_sproc _ = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 270 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 271 | fun is_many_smsg :: "t_smsg \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 272 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 273 | "is_many_smsg (Created,sec,tag) = True" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 274 | | "is_many_smsg _ = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 275 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 276 | (* wrong def | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 277 | fun is_many_smsgq :: "t_smsgq \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 278 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 279 | "is_many_smsgq (Created,sec,sms) = (True \<and> (\<forall> sm \<in> (set sms). is_many_smsg sm))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 280 | | "is_many_smsgq _ = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 281 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 282 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 283 | fun is_many_smsgq :: "t_smsgq \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 284 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 285 | "is_many_smsgq (Created,sec,sms) = True" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 286 | | "is_many_smsgq _ = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 287 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 288 | fun is_many :: "t_sobject \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 289 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 290 | "is_many (S_proc sp tag) = is_many_sproc sp" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 291 | | "is_many (S_file sf tag) = is_many_sfile sf" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 292 | | "is_many (S_dir sf ) = is_many_sfile sf" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 293 | | "is_many (S_msgq sq ) = is_many_smsgq sq" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 294 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 295 | fun is_init_smsgq :: "t_smsgq \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 296 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 297 | "is_init_smsgq (Init q,sec,sms) = True" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 298 | | "is_init_smsgq _ = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 299 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 300 | fun is_init_sobj :: "t_sobject \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 301 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 302 | "is_init_sobj (S_proc sp tag ) = is_init_sproc sp" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 303 | | "is_init_sobj (S_file sf tag ) = is_init_sfile sf" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 304 | | "is_init_sobj (S_dir sf ) = is_init_sfile sf" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 305 | | "is_init_sobj (S_msgq sq ) = is_init_smsgq sq" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 306 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 307 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 308 | fun update_ss_sp:: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_sobject \<Rightarrow> t_static_state" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 309 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 310 | "update_ss_sp ss (S_proc sp tag) (S_proc sp' tag') = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 311 |      (if (is_many_sproc sp) then ss \<union> {S_proc sp' tag'} 
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 312 |       else (ss - {S_proc sp tag}) \<union> {S_proc sp' tag'})"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 313 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 314 | fun update_ss_sd:: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_sobject \<Rightarrow> t_static_state" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 315 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 316 | "update_ss_sd ss (S_dir sf tag) (S_dir sf' tag') = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 317 |      (if (is_many_sfile sf) then ss \<union> {S_dir sf' tag'} 
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 318 |       else (ss - {S_dir sf tag}) \<union> {S_dir sf' tag'})"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 319 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 320 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 321 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 322 | fun sparent :: "t_sfile \<Rightarrow> t_sfile option" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 323 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 324 | "sparent (Sroot si sec) = None" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 325 | | "sparent (Sfile si sec spf) = Some spf" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 326 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 327 | inductive is_ancesf :: "t_sfile \<Rightarrow> t_sfile \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 328 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 329 | "is_ancesf sf sf" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 330 | | "sparent sf = Some spf \<Longrightarrow> is_ancesf spf sf" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 331 | | "\<lbrakk>sparent sf = Some spf; is_ancesf saf spf\<rbrakk> \<Longrightarrow> is_ancesf saf sf" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 332 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 333 | definition sfile_reparent :: "t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 334 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 335 | "sfile_reparent (Sroot)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 336 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 337 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 338 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 339 | (* sfds rename aux definitions *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 340 | definition sfds_rename_notrelated | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 341 | :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 342 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 343 | "sfds_rename_notrelated sfds from to sfds' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 344 | (\<forall> sec flag sf. (sec,flag,sf) \<in> sfds \<and> (\<not> from \<preceq> sf) \<longrightarrow> (sec,flag,sf) \<in> sfds')" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 345 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 346 | definition sfds_rename_renamed | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 347 | :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 348 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 349 | "sfds_rename_renamed sfds sf from to sfds' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 350 | (\<forall> sec flag sf'. (sec,flag,sf) \<in> sfds \<and> (sf \<preceq> sf') \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 351 | (sec, flag, file_after_rename sf' from to) \<in> sfds' \<and> (sec,flag,sf) \<notin> sfds')" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 352 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 353 | definition sfds_rename_remain | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 354 | :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 355 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 356 | "sfds_rename_remain sfds sf from to sfds' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 357 | (\<forall> sec flag sf'. (sec,flag,sf') \<in> sfds \<and> (sf \<preceq> sf') \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 358 | (sec, flag, sf') \<in> sfds' \<and> (sec,flag, file_after_rename sf' from to) \<notin> sfds')" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 359 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 360 | (* for not many, choose on renamed or not *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 361 | definition sfds_rename_choices | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 362 | :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 363 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 364 | "sfds_rename_choices sfds sf from to sfds' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 365 | sfds_rename_remain sfds sf from to sfds' \<or> sfds_rename_renamed sfds sf from to sfds'" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 366 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 367 | (* for many, merge renamed with not renamed *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 368 | definition sfds_rename_both | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 369 | :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 370 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 371 | "sfds_rename_both sfds sf from to sfds' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 372 | (\<forall> sec flag sf'. (sec,flag,sf') \<in> sfds \<and> (sf \<preceq> sf') \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 373 | (sec, flag, sf') \<in> sfds' \<or> (sec,flag, file_after_rename sf' from to) \<in> sfds')" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 374 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 375 | (* added to the new sfds, are those only under the new sfile *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 376 | definition sfds_rename_added | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 377 | :: "t_sfd set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfd set \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 378 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 379 | "sfds_rename_added sfds from to sfds' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 380 | (\<forall> sec' flag' sf' sec flag. (sec',flag',sf') \<in> sfds' \<and> (sec,flag,sf') \<notin> sfds \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 381 | (\<exists> sf. (sec,flag,sf) \<in> sfds \<and> sf' = file_after_rename from to sf))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 382 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 383 | definition sproc_sfds_renamed | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 384 | :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 385 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 386 | "sproc_sfds_renamed ss sf from to ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 387 | (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 388 | (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_renamed sfds sf from to sfds'))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 389 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 390 | definition sproc_sfds_remain | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 391 | :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 392 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 393 | "sproc_sfds_remain ss sf from to ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 394 | (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 395 | (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_remain sfds sf from to sfds'))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 396 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 397 | (* for not many, choose on renamed or not *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 398 | definition sproc_sfds_choices | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 399 | :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 400 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 401 | "sproc_sfds_choices ss sf from to ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 402 | (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 403 | (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_choices sfds sf from to sfds'))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 404 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 405 | (* for many, merge renamed with not renamed *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 406 | definition sproc_sfds_both | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 407 | :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 408 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 409 | "sproc_sfds_both ss sf from to ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 410 | (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 411 | (\<exists> sfds'. S_proc (pi,sec,sfds',sshms) tagp \<in> ss \<and> sfds_rename_both sfds sf from to sfds'))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 412 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 413 | (* remove (\<forall> sp tag. S_proc sp tag \<in> ss \<longrightarrow> S_proc sp tag \<in> ss'), | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 414 | * cause sfds contains sfs informations *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 415 | definition ss_rename_notrelated | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 416 | :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 417 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 418 | "ss_rename_notrelated ss sf sf' ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 419 | (\<forall> sq. S_msgq sq \<in> ss \<longrightarrow> S_msgq sq \<in> ss') \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 420 | (\<forall> pi sec sfds sshms tagp. S_proc (pi,sec,sfds,sshms) tagp \<in> ss \<longrightarrow> (\<exists> sfds'. | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 421 | S_proc (pi,sec,sfds',sshms) tagp \<in> ss'\<and> sfds_rename_notrelated sfds sf sf' sfds')) \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 422 | (\<forall> sfs sf'' tag. S_file sfs tag \<in> ss \<and> sf'' \<in> sfs \<and> (\<not> sf \<preceq> sf'') \<longrightarrow> (\<exists> sfs'. | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 423 | S_file sfs tag \<in> ss' \<and> sf'' \<in> sfs')) \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 424 | (\<forall> sf'' tag. S_dir sf'' tag \<in> ss \<and> (\<not> sf \<preceq> sf'') \<longrightarrow> S_dir sf'' tag \<in> ss')" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 425 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 426 | (* rename from to, from should definited renamed if not many *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 427 | definition all_descendant_sf_renamed | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 428 | :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 429 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 430 | "all_descendant_sf_renamed ss sf from to ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 431 | (\<forall> sfs sf' tagf. sf \<preceq> sf' \<and> S_file sfs tagf \<in> ss \<and> sf' \<in> sfs \<longrightarrow> (\<exists> sfs'. | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 432 | S_file sfs' tagf \<in> ss' \<and> file_after_rename from to sf' \<in> sfs' \<and> sf' \<notin> sfs')) \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 433 | (\<forall> sf' tagf. sf \<preceq> sf' \<and> S_dir sf' tagf \<in> ss \<longrightarrow> S_dir (file_after_rename from to sf') tagf \<in> ss' \<and> S_dir sf' tagf \<notin> ss') \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 434 | sproc_sfds_renamed ss sf from to ss'" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 435 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 436 | (* not renamed *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 437 | definition all_descendant_sf_remain | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 438 | :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 439 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 440 | "all_descendant_sf_remain ss sf from to ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 441 | (\<forall> sfs sf' tag'. sf \<preceq> sf' \<and> S_file sfs tag' \<in> ss \<and> sf' \<in> sfs \<longrightarrow> (\<exists> sfs'. | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 442 | S_file sfs' tag' \<in> ss \<and> file_after_rename from to sf' \<notin> sfs' \<and> sf' \<in> sfs')) \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 443 | (\<forall> sf' tag'. sf \<preceq> sf' \<and> S_dir sf' tag' \<in> ss \<longrightarrow> S_dir (file_after_rename from to sf') tag' \<notin> ss' \<and> S_dir sf' tag' \<in> ss') \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 444 | sproc_sfds_remain ss sf from to ss'" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 445 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 446 | definition all_descendant_sf_choices | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 447 | :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 448 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 449 | "all_descendant_sf_choices ss sf from to ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 450 | all_descendant_sf_renamed ss sf from to ss' \<or> all_descendant_sf_remain ss sf from to ss'" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 451 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 452 | definition all_descendant_sf_both | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 453 | :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 454 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 455 | "all_descendant_sf_both ss sf from to ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 456 | (\<forall> sfs sf' tag'. sf \<preceq> sf' \<and> S_file sfs tag' \<in> ss \<and> sf' \<in> sfs \<longrightarrow> (\<exists> sfs'. | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 457 | S_file sfs' tag' \<in> ss \<and> file_after_rename from to sf' \<in> sfs' \<or> sf' \<in> sfs')) \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 458 | (\<forall> sf' tag'. sf \<preceq> sf' \<and> S_dir sf' tag' \<in> ss \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 459 | S_dir (file_after_rename from to sf') tag' \<in> ss' \<or> S_dir sf' tag' \<in> ss') \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 460 | sproc_sfds_both ss sf from to ss'" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 461 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 462 | definition ss_renamed_file | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 463 | :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 464 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 465 | "ss_renamed_file ss sf sf' ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 466 | (if (is_many_sfile sf) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 467 | then all_descendant_sf_choices ss sf sf sf' ss' | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 468 | else all_descendant_sf_renamed ss sf sf sf' ss')" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 469 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 470 | (* added to the new sfs, are those only under the new sfile *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 471 | definition sfs_rename_added | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 472 | :: "t_sfile set \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_sfile set \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 473 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 474 | "sfs_rename_added sfs from to sfs' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 475 | (\<forall> sf'. sf' \<in> sfs' \<and> sf' \<notin> sfs \<longrightarrow> (\<exists> sf. sf \<in> sfs \<and> sf' = file_after_rename from to sf))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 476 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 477 | (* added to the new sfile, are those only under the new sfile *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 478 | definition ss_rename_added | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 479 | :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 480 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 481 | "ss_rename_added ss from to ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 482 | (\<forall> pi sec fds fds' shms tagp. S_proc (pi, sec, fds,shms) tagp \<in> ss \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 483 | S_proc (pi,sec,fds',shms) tagp \<in> ss' \<longrightarrow> sfds_rename_added fds from to fds') \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 484 | (\<forall> sq. S_msgq sq \<in> ss' \<longrightarrow> S_msgq sq \<in> ss) \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 485 | (\<forall> sfs sfs' tagf. S_file sfs' tagf \<in> ss' \<and> S_file sfs tagf \<in> ss \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 486 | sfs_rename_added sfs from to sfs') \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 487 | (\<forall> sf' tagf. S_dir sf' tagf \<in> ss' \<and> S_dir sf' tagf \<notin> ss \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 488 | (\<exists> sf. S_dir sf tagf \<in> ss \<and> sf' = file_after_rename from to sf))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 489 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 490 | definition sfile_alive :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 491 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 492 | "sfile_alive ss sf \<equiv> (\<exists> sfs tagf. sf \<in> sfs \<and> S_file sfs tagf \<in> ss)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 493 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 494 | definition sf_alive :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 495 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 496 | "sf_alive ss sf \<equiv> (\<exists> tagd. S_dir sf tagd \<in> ss) \<or> sfile_alive ss sf" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 497 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 498 | (* constrains that the new static state should satisfy *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 499 | definition ss_rename:: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 500 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 501 | "ss_rename ss sf sf' ss' \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 502 | ss_rename_notrelated ss sf sf' ss' \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 503 | ss_renamed_file ss sf sf' ss' \<and> ss_rename_added ss sf sf' ss' \<and> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 504 | (\<forall> sf''. sf_alive ss sf'' \<and> sf \<prec> sf'' \<longrightarrow> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 505 | (if (is_many_sfile sf'') | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 506 | then all_descendant_sf_choices ss sf'' sf sf' ss' | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 507 | else all_descendant_sf_both ss sf'' sf sf' ss'))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 508 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 509 | (* two sfile, the last fname should not be equal *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 510 | fun sfile_same_fname:: "t_sfile \<Rightarrow> t_sfile \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 511 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 512 | "sfile_same_fname ((Init n, sec)#spf) ((Init n', sec')#spf') = (n = n')" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 513 | | "sfile_same_fname _ _ = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 514 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 515 | (* no same init sfile/only sfile in the target-to spf, should be in static_state addmissble check *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 516 | definition ss_rename_no_same_fname:: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 517 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 518 | "ss_rename_no_same_fname ss from spf \<equiv> | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 519 | \<not> (\<exists> to. sfile_same_fname from to \<and> parent to = Some spf \<and> sf_alive ss to)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 520 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 521 | (* is not a function, is a relation (one 2 many) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 522 | definition update_ss_rename :: "t_static_state \<Rightarrow> t_sfile \<Rightarrow> t_sfile \<Rightarrow> t_static_state" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 523 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 524 | "update_ss_rename ss sf sf' \<equiv> if (is_many_sfile sf) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 525 |      then (ss \<union> {S_file (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_file sf'' tag \<in> ss}
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 526 |               \<union> {S_dir  (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_dir sf'' tag \<in> ss})
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 527 |      else (ss - {S_file sf'' tag | sf'' tag. sf \<preceq> sf'' \<and> S_file sf'' tag \<in> ss}
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 528 |               - {S_dir  sf'' tag | sf'' tag. sf \<preceq> sf'' \<and> S_dir sf'' tag \<in> ss})
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 529 |               \<union> {S_file (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_file sf'' tag \<in> ss}
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 530 |               \<union> {S_dir  (file_after_rename sf sf' sf'') tag | sf'' tag. sf \<preceq> sf'' \<and> S_dir sf'' tag \<in> ss}" 
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 531 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 532 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 533 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 534 | fun sectxt_of_sproc :: "t_sproc \<Rightarrow> security_context_t" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 535 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 536 | "sectxt_of_sproc (pi,sec,fds) = sec" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 537 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 538 | fun sectxt_of_sfile :: "t_sfile \<Rightarrow> security_context_t" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 539 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 540 | "sectxt_of_sfile (fi,sec,psec,asecs) = sec" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 541 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 542 | fun asecs_of_sfile :: "t_sfile \<Rightarrow> security_context_t set" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 543 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 544 | "asecs_of_sfile (fi,sec,psec,asecs) = asecs" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 545 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 546 | definition search_check_s :: "security_context_t \<Rightarrow> t_sfile \<Rightarrow> bool \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 547 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 548 | "search_check_s pctxt sf if_file = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 549 | (if if_file | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 550 | then search_check_ctxt pctxt (sectxt_of_sfile sf) (asecs_of_sfile sf) True | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 551 | else search_check_ctxt pctxt (sectxt_of_sfile sf) (asecs_of_sfile sf) False)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 552 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 553 | definition sectxts_of_sfds :: "t_sfd set \<Rightarrow> security_context_t set" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 554 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 555 |   "sectxts_of_sfds sfds \<equiv> {ctxt. \<exists> flag sf. (ctxt, flag, sf) \<in> sfds}"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 556 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 557 | definition inherit_fds_check_s :: "security_context_t \<Rightarrow> t_sfd set \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 558 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 559 | "inherit_fds_check_s pctxt sfds \<equiv> inherit_fds_check_ctxt pctxt (sectxts_of_sfds sfds)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 560 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 561 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 562 | fun smsg_related :: "t_msg \<Rightarrow> t_smsg list \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 563 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 564 | "smsg_related m [] = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 565 | | "smsg_related m ((mi, sec, tag)#sms) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 566 | (if (mi = Init m) then True else smsg_related m sms)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 567 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 568 | fun smsgq_smsg_related :: "t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsgq \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 569 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 570 | "smsgq_smsg_related q m (qi, sec, smsgslist) = ((qi = Init q) \<and> (smsg_related m smsgslist))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 571 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 572 | fun smsg_relatainted :: "t_msg \<Rightarrow> t_smsg list \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 573 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 574 | "smsg_relatainted m [] = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 575 | | "smsg_relatainted m ((mi, sec, tag)#sms) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 576 | (if (mi = Init m \<and> tag = True) then True else smsg_relatainted m sms)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 577 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 578 | fun smsgq_smsg_relatainted :: "t_msgq \<Rightarrow> t_msg \<Rightarrow> t_smsgq \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 579 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 580 | "smsgq_smsg_relatainted q m (qi, sec, smsgslist) = | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 581 | ((qi = Init q) \<and> (smsg_relatainted m smsgslist))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 582 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 583 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 584 | fun sfile_related :: "t_sfile \<Rightarrow> t_file \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 585 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 586 | "sfile_related (fi,sec,psec,asecs) f = (fi = Init f)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 587 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 588 | fun sproc_related :: "t_process \<Rightarrow> t_sproc \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 589 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 590 | "sproc_related p (pi, sec, fds, shms) = (pi = Init p)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 591 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 592 | fun init_obj_related :: "t_sobject \<Rightarrow> t_object \<Rightarrow> bool" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 593 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 594 | "init_obj_related (S_proc (pi, sec, fds) tag) (O_proc p') = (pi = Init p')" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 595 | | "init_obj_related (S_file sf tag) (O_file f) = (sfile_related sf f)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 596 | | "init_obj_related (S_dir sf) (O_dir f) = (sfile_related sf f)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 597 | | "init_obj_related (S_msgq (qi, sec, sms)) (O_msgq q') = (qi = Init q')" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 598 | (* | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 599 | | "init_obj_related (S_msg (qi, sec, sms) (mi, secm, tagm)) (O_msg q' m') = (qi = Init q' \<and> mi = Init m')" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 600 | *) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 601 | | "init_obj_related _ _ = False" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 602 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 603 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 604 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 605 | (***************** for backward proof when Instancing static objects ******************) | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 606 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 607 | definition next_nat :: "nat set \<Rightarrow> nat" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 608 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 609 | "next_nat nset = (Max nset) + 1" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 610 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 611 | definition new_proc :: "t_state \<Rightarrow> t_process" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 612 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 613 | "new_proc \<tau> = next_nat (current_procs \<tau>)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 614 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 615 | definition new_inode_num :: "t_state \<Rightarrow> nat" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 616 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 617 | "new_inode_num \<tau> = next_nat (current_inode_nums \<tau>)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 618 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 619 | definition new_msgq :: "t_state \<Rightarrow> t_msgq" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 620 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 621 | "new_msgq s = next_nat (current_msgqs s)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 622 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 623 | definition new_msg :: "t_state \<Rightarrow> t_msgq \<Rightarrow> nat" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 624 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 625 | "new_msg s q = next_nat (set (msgs_of_queue s q))" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 626 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 627 | definition new_proc_fd :: "t_state \<Rightarrow> t_process \<Rightarrow> t_fd" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 628 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 629 | "new_proc_fd \<tau> p = next_nat (current_proc_fds \<tau> p)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 630 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 631 | definition all_fname_under_dir:: "t_file \<Rightarrow> t_state \<Rightarrow> t_fname set" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 632 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 633 |   "all_fname_under_dir d \<tau> = {fn. \<exists> f. fn # d = f \<and> f \<in> current_files \<tau>}"
 | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 634 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 635 | fun fname_all_a:: "nat \<Rightarrow> t_fname" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 636 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 637 | "fname_all_a 0 = []" | | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 638 | "fname_all_a (Suc n) = ''a''@(fname_all_a n)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 639 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 640 | definition fname_length_set :: "t_fname set \<Rightarrow> nat set" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 641 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 642 | "fname_length_set fns = length`fns" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 643 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 644 | definition next_fname:: "t_file \<Rightarrow> t_state \<Rightarrow> t_fname" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 645 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 646 | "next_fname pf \<tau> = fname_all_a ((Max (fname_length_set (all_fname_under_dir pf \<tau>))) + 1)" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 647 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 648 | definition new_childf:: "t_file \<Rightarrow> t_state \<Rightarrow> t_file" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 649 | where | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 650 | "new_childf pf \<tau> = next_fname pf \<tau> # pf" | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 651 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 652 | end | 
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 653 | |
| 
271e9818b6f6
remove shm and linkhard, make a simplified version of selinux
 chunhan parents: diff
changeset | 654 | end |