diff -r 9fc384154e84 -r 051b0ee98852 Dynamic_static.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Dynamic_static.thy Thu Oct 24 09:41:33 2013 +0800 @@ -0,0 +1,153 @@ +theory Dynamic_static +imports Main Flask Static Init_prop Valid_prop Tainted_prop Delete_prop Co2sobj_prop S2ss_prop S2ss_prop2 +begin + +context tainting_s begin + +definition init_ss_eq:: "t_static_state \ t_static_state \ bool" (infix "\" 100) +where + "ss \ ss' \ ss \ ss' \ {sobj. is_init_sobj sobj \ sobj \ ss'} \ ss" + +lemma [simp]: "ss \ ss" +by (auto simp:init_ss_eq_def) + +definition init_ss_in:: "t_static_state \ t_static_state set \ bool" (infix "\" 101) +where + "ss \ sss \ \ ss' \ sss. ss \ ss'" + +lemma s2ss_included_sobj: + "\alive s obj; co2sobj s obj= Some sobj\ \ sobj \ (s2ss s)" +by (simp add:s2ss_def, rule_tac x = obj in exI, simp) + +lemma init_ss_in_prop: + "\s2ss s \ static; co2sobj s obj = Some sobj; alive s obj; init_obj_related sobj obj\ + \ \ ss \ static. sobj \ ss" +apply (simp add:init_ss_in_def init_ss_eq_def) +apply (erule bexE, erule conjE) +apply (rule_tac x = ss' in bexI, auto dest!:s2ss_included_sobj) +done + + + + + +definition enrich:: "t_state \ t_object set \ t_state \ bool" +where + "enrich s objs s' \ \ obj \ objs. \ obj'. obj' \ objs \ alive s' obj \ co2sobj s' obj' = co2sobj s' obj" + +definition reserve:: "t_state \ t_object set \ t_state \ bool" +where + "reserve s objs s' \ \ obj. alive s obj \ alive s' obj \ co2sobj s' obj = co2sobj s obj" + +definition enrichable :: "t_state \ t_object set \ bool" +where + "enrichable s objs \ \ s'. valid s' \ s2ss s' = s2ss s \ enrich s objs s' \ reserve s objs s'" + +definition is_created :: "t_state \ t_object \ bool" +where + "is_created s obj \ init_alive obj \ deleted obj s" + +definition is_inited :: "t_state \ t_object \ bool" +where + "is_inited s obj \ init_alive obj \ \ deleted obj s" + +lemma is_inited_eq_not_created: + "is_inited s obj = (\ is_created s obj)" +by (auto simp:is_created_def is_inited_def) + +(* recorded in our static world *) +fun recorded :: "t_object \ bool" +where + "recorded (O_proc p) = True" +| "recorded (O_file f) = True" +| "recorded (O_dir f) = True" +| "recorded (O_node n) = False" (* cause socket is temperary not considered *) +| "recorded (O_shm h) = True" +| "recorded (O_msgq q) = True" +| "recorded _ = False" + + + + + + +lemma d2s_main_execve: + "valid (Execve p f fds # s) \ s2ss (Execve p f fds # s) \ static" +apply (frule vd_cons, frule vt_grant_os, clarsimp simp:s2ss_execve) +sorry + +lemma d2s_main: + "valid s \ s2ss s \ static" +apply (induct s, simp add:s2ss_nil_prop init_ss_in_def) +apply (rule_tac x = "init_static_state" in bexI, simp, simp add:s_init) +apply (frule vd_cons, frule vt_grant_os, simp) +apply (case_tac a) +apply (clarsimp simp add:s2ss_execve) +apply (rule conjI, rule impI) + + + +sorry + + +lemma many_sq_imp_sms: + "\S_msgq (Create, sec, sms) \ ss; ss \ static\ \ \ sm \ (set sms). is_many_smsg sm" +sorry + + + +lemma enrichability: + "\valid s; \ obj \ objs. alive s obj \ is_created s obj \ recorded obj\ + \ enrichable s objs" +proof (induct s arbitrary:objs) + case Nil + hence "objs = {}" + apply (auto simp:is_created_def) + apply (erule_tac x = x in ballE) + apply (auto simp:init_alive_prop) + done + thus ?case using Nil unfolding enrichable_def enrich_def reserve_def + by (rule_tac x = "[]" in exI, auto) +next + case (Cons e s) + hence p1: "\ objs. \ obj \ objs. alive s obj \ is_created s obj \ recorded obj \ enrichable s objs" + and p2: "valid (e # s)" and p3: "\obj\objs. alive (e # s) obj \ is_created (e # s) obj \ recorded obj" + and os: "os_grant s e" and se: "grant s e" and vd: "valid s" + by (auto dest:vt_grant_os vd_cons vt_grant) + show ?case sorry (* + proof (cases e) + case (Execve p f fds) + hence p4: "e = Execve p f fds" by simp + from p3 have p5: "is_inited s (O_proc p) \ (O_proc p) \ objs" + by (auto simp:is_created_def is_inited_def p4 elim!:ballE[where x = "O_proc p"]) + show "enrichable (e # s) objs" + proof (case "is_inited s (O_proc p)") + apply (simp add:enrichable_def p4) + + + + apply auto + apply (auto simp:enrichable_def) +apply (induct s) + + + +done +*) +qed + +lemma s2d_main: + "ss \ static \ \ s. valid s \ s2ss s = ss" +apply (erule static.induct) +apply (rule_tac x = "[]" in exI, simp add:s2ss_nil_prop valid.intros) + +apply (erule exE|erule conjE)+ + +apply (simp add:update_ss_def) + +sorry + + +end + +end \ No newline at end of file