diff -r 9fc384154e84 -r 051b0ee98852 Dynamic2static.thy --- a/Dynamic2static.thy Tue Oct 22 10:08:27 2013 +0800 +++ b/Dynamic2static.thy Thu Oct 24 09:41:33 2013 +0800 @@ -54,6 +54,112 @@ sorry +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 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 + 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 + + +(* for the object set, there exists another trace which keeps this objects but also add new identical objects + * that have the same static-signature + *) + +definition potential_trace:: "t_state \ bool" +where + "potential_trace s \ \ obj. alive s obj \ is_created s obj \ + (\ s' obj'. valid s' \ s2ss s' = ss \ obj' \ obj \ co2sobj s' obj = co2sobj s' obj) + " + +lemma s2d_main_general: + "ss \ static \ \ s. valid s \ s2ss s = ss \ (\ obj \ objs. alive s obj \ is_created s obj \ (\ s'. valid s' \ s2ss s' = ss \ (\ obj'. obj' \ obj \ co2sobj s' obj = co2sobj s' obj')))" +apply (erule static.induct) +apply (rule_tac x = "[]" in exI, simp add:s2ss_nil_prop valid.intros) defer + +apply (erule exE|erule conjE)+ + +apply (simp add:update_ss_def) + +sorry + +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 + lemma t2ts: "obj \ tainted s \ co2sobj s obj = Some sobj \ tainted_s (s2ss s) sobj" @@ -219,14 +325,6 @@ (************** static \ dynamic ***************) -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)+ - -sorry lemma set_eq_D: "\x \ S; {x. P x} = S\ \ P x"