# HG changeset patch # User chunhan # Date 1384835516 -28800 # Node ID 811e3028d169e2c0dc4325ceb65c4f474166af04 # Parent 5f86fb3ddd446c868d5a59f5aacfd30b7d9de769 update grant_check diff -r 5f86fb3ddd44 -r 811e3028d169 Dynamic_static.thy --- a/Dynamic_static.thy Tue Nov 19 09:27:25 2013 +0800 +++ b/Dynamic_static.thy Tue Nov 19 12:31:56 2013 +0800 @@ -86,34 +86,78 @@ where "is_created_proc s p \ p \ init_procs \ deleted (O_proc p) s" +lemma enrich_proc_aux1: + assumes vs': "valid s'" + and os: "os_grant s e" and grant: "grant s e" + and alive: "\ obj. alive s obj \ alive s' obj" + and cp2sp: "\ p. p \ current_procs s \ cp2sproc s' p = cp2sproc s p" + and cf2sf: "\ f. f \ current_files s \ cf2sfile s' f = cf2sfile s f" + shows "valid (e # s')" +proof (cases e) + case (Execve p f fds) + have p_in: "p \ current_procs s'" using os alive + apply (erule_tac x = "O_proc p" in allE) + by (auto simp:Execve) + have f_in: "is_file s' f" using os alive + apply (erule_tac x = "O_file f" in allE) + by (auto simp:Execve) + have fd_in: "fds \ current_proc_fds s' p" using os alive + apply (auto simp:Execve) + by (erule_tac x = "O_fd p x" in allE, auto) + have "os_grant s' e" using p_in f_in fd_in by (simp add:Execve) + moreover have "grant s' e" + proof- + from grant obtain up rp tp uf rf tf + where p1: "sectxt_of_obj s (O_proc p) = Some (up, rp, tp)" + and p2: "sectxt_of_obj s (O_file f) = Some (uf, rf, tf)" + by (simp add:Execve split:option.splits, blast) + with grant obtain pu nr nt where p3: "npctxt_execve (up, rp, tp) (uf, rf, tf) = Some (pu, nr, nt)" + by (simp add:Execve split:option.splits del:npctxt_execve.simps, blast) + from p1 have p1': "sectxt_of_obj s' (O_proc p) = Some (up, rp, tp)" + using os cp2sp + apply (erule_tac x = p in allE) + by (auto simp:Execve co2sobj.simps cp2sproc_def split:option.splits) + from p2 have p2': "sectxt_of_obj s' (O_file f) = Some (uf, rf, tf)" + using os cf2sf + apply (erule_tac x = "f" in allE) + apply (auto simp:Execve dest:is_file_in_current) + thm cf2sfile_def + apply (auto simp:cf2sfile_def split:option.splits) + apply (auto simp:Execve co2sobj.simps cf2sfile_simps split:option.splits) + apply (simp add:cf2sfiles_def) + apply (auto)[1] + using os pre + show ?thesis + proof- + have + + + + lemma enrich_proc_prop: "\valid s; is_created_proc s p; p' \ all_procs s\ \ valid (enrich_proc s p p') \ - co2sobj (enrich_proc s p p') (O_proc p') = co2sobj (enrich_proc s p p') (O_proc p) \ - (alive s obj \ alive (enrich_proc s p p') obj \ co2sobj (enrich_proc s p p') obj = co2sobj s obj)" + (p \ current_procs s \ co2sobj (enrich_proc s p p') (O_proc p') = co2sobj (enrich_proc s p p') (O_proc p)) \ + (\ obj. alive s obj \ alive (enrich_proc s p p') obj \ co2sobj (enrich_proc s p p') obj = co2sobj s obj)" proof (induct s) case Nil thus ?case by (auto simp:is_created_proc_def) next case (Cons e s) - hence p1: "\p \ current_procs s; valid s; is_created_proc s p; p' \ current_procs s\ + hence p1: "\valid s; is_created_proc s p; p' \ all_procs s\ \ valid (enrich_proc s p p') \ - p' \ current_procs (enrich_proc s p p') \ - co2sobj (enrich_proc s p p') (O_proc p') = co2sobj (enrich_proc s p p') (O_proc p) \ + (p \ current_procs s \ co2sobj (enrich_proc s p p') (O_proc p') = co2sobj (enrich_proc s p p') (O_proc p)) \ (alive s obj \ alive (enrich_proc s p p') obj \ co2sobj (enrich_proc s p p') obj = co2sobj s obj)" - and p2: "p \ current_procs (e # s)" and p3: "valid (e # s)" - and p4: "is_created_proc (e # s) p" and p5: "p' \ current_procs (e # s)" + and p2: "valid (e # s)" and p3: "is_created_proc (e # s) p" and p4: "p' \ all_procs (e # s)" by auto - from p4 p2 have p4': "is_created_proc s p" - by (case_tac e, auto simp:is_created_proc_def) - from p3 have vd: "valid s" and os: "os_grant s e" and grant: "grant s e" + from p2 have vd: "valid s" and os: "os_grant s e" and grant: "grant s e" by (auto dest:vd_cons vt_grant vt_grant_os) - from p1 p4' have a1: "\p \ current_procs s; p' \ current_procs s\ - \ valid (enrich_proc s p p')" by (auto simp:vd) + from p4 have p4': "p' \ all_procs s" by (case_tac e, auto) + from p1 p4' have a1: "is_created_proc s p \ valid (enrich_proc s p p')" by (auto simp:vd) have c1: "valid (enrich_proc (e # s) p p')" apply (case_tac e) - using a1 os p5 - apply (auto) + using a1 os p3 + apply (auto simp:is_created_proc_def) sorry moreover have c2: "p' \ current_procs (enrich_proc (e # s) p p')" sorry diff -r 5f86fb3ddd44 -r 811e3028d169 Flask.thy --- a/Flask.thy Tue Nov 19 09:27:25 2013 +0800 +++ b/Flask.thy Tue Nov 19 12:31:56 2013 +0800 @@ -1115,19 +1115,27 @@ (Some ctxts, Some ctxt) \ Some (ctxt#ctxts) | _ \ None)" +definition search_check_ctxt:: + "security_context_t \ security_context_t \ security_context_t set \ bool \ bool" +where + "search_check_ctxt pctxt fctxt asecs if_file \ ( + if if_file + then search_check_file pctxt fctxt \ search_check_allp pctxt asecs + else search_check_dir pctxt fctxt \ search_check_allp pctxt asecs )" + fun search_check :: "t_state \ security_context_t \ t_file \ bool" where "search_check s pctxt [] = (case (sectxt_of_obj s (O_dir [])) of - Some fctxt \ search_check_dir pctxt fctxt + Some fctxt \ search_check_ctxt pctxt fctxt {} False | _ \ False)" | "search_check s pctxt (f#pf) = (if (is_file s (f#pf)) then (case (sectxt_of_obj s (O_file (f#pf)), get_parentfs_ctxts s pf) of - (Some fctxt, Some pfctxts) \ (search_check_file pctxt fctxt \ search_check_allp pctxt (set pfctxts)) + (Some fctxt, Some pfctxts) \ search_check_ctxt pctxt fctxt (set pfctxts) True | _ \ False) else (case (sectxt_of_obj s (O_dir (f#pf)), get_parentfs_ctxts s pf) of - (Some fctxt, Some pfctxts) \ (search_check_dir pctxt fctxt \ search_check_allp pctxt (set pfctxts)) + (Some fctxt, Some pfctxts) \ search_check_ctxt pctxt fctxt (set pfctxts) False | _ \ False))" (* this means we should prove every current fd has a security context! *) @@ -1135,10 +1143,13 @@ where "sectxts_of_fds s p fds \ {ctxt. \ fd \ fds. sectxt_of_obj s (O_fd p fd) = Some ctxt}" +definition inherit_fds_check_ctxt:: "security_context_t \ security_context_t set \ bool" +where + "inherit_fds_check_ctxt p fds \ (\ ctxt \ fds. permission_check p ctxt C_fd P_inherit)" + definition inherit_fds_check :: "t_state \ security_context_t \ t_process \ t_fd set \ bool" where - "inherit_fds_check s npsectxt p fds \ - (\ ctxt \ sectxts_of_fds s p fds. permission_check npsectxt ctxt C_fd P_inherit)" + "inherit_fds_check s npsectxt p fds \ inherit_fds_check_ctxt npsectxt (sectxts_of_fds s p fds)" fun npctxt_execve :: "security_context_t \ security_context_t \ security_context_t option" where @@ -1165,10 +1176,13 @@ where "sectxts_of_shms s shms \ {ctxt. \ h \ shms. sectxt_of_obj s (O_shm h) = Some ctxt}" +definition inherit_shms_check_ctxt:: "security_context_t \ security_context_t set \ bool" +where + "inherit_shms_check_ctxt p shms \ (\ ctxt \ shms. permission_check p ctxt C_shm P_inherit)" + definition inherit_shms_check :: "t_state \ security_context_t \ t_shm set \ bool" where - "inherit_shms_check s npsectxt shms \ - (\ ctxt \ sectxts_of_shms s shms. permission_check npsectxt ctxt C_shm P_inherit)" + "inherit_shms_check s npsectxt shms \ inherit_shms_check_ctxt npsectxt (sectxts_of_shms s shms)" fun perm_of_mflag :: "t_open_must_flag \ permission_t set" where diff -r 5f86fb3ddd44 -r 811e3028d169 Static.thy --- a/Static.thy Tue Nov 19 09:27:25 2013 +0800 +++ b/Static.thy Tue Nov 19 12:31:56 2013 +0800 @@ -549,8 +549,8 @@ where "search_check_s pctxt sf if_file = (if if_file - then search_check_file pctxt (sectxt_of_sfile sf) \ search_check_allp pctxt (asecs_of_sfile sf) - else search_check_dir pctxt (sectxt_of_sfile sf) \ search_check_allp pctxt (asecs_of_sfile sf))" + then search_check_ctxt pctxt (sectxt_of_sfile sf) (asecs_of_sfile sf) True + else search_check_ctxt pctxt (sectxt_of_sfile sf) (asecs_of_sfile sf) False)" definition sectxts_of_sfds :: "t_sfd set \ security_context_t set" where @@ -558,8 +558,7 @@ definition inherit_fds_check_s :: "security_context_t \ t_sfd set \ bool" where - "inherit_fds_check_s pctxt sfds \ - (\ ctxt \ sectxts_of_sfds sfds. permission_check pctxt ctxt C_fd P_inherit)" + "inherit_fds_check_s pctxt sfds \ inherit_fds_check_ctxt pctxt (sectxts_of_sfds sfds)" definition sectxts_of_sproc_sshms :: "t_sproc_sshm set \ security_context_t set" where @@ -567,8 +566,7 @@ definition inherit_shms_check_s :: "security_context_t \ t_sproc_sshm set \ bool" where - "inherit_shms_check_s pctxt sshms \ - (\ ctxt \ sectxts_of_sproc_sshms sshms. permission_check pctxt ctxt C_shm P_inherit)" + "inherit_shms_check_s pctxt sshms \ inherit_shms_check_ctxt pctxt (sectxts_of_sproc_sshms sshms)" (* fun info_flow_sshm :: "t_sproc \ t_sproc \ bool"