Dynamic2static.thy
changeset 63 051b0ee98852
parent 62 9fc384154e84
--- 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 \<Rightarrow> t_object set \<Rightarrow> t_state \<Rightarrow> bool"
+where
+  "enrich s objs s' \<equiv> \<forall> obj \<in> objs. \<exists> obj'. obj' \<notin> objs \<and> alive s' obj \<and> co2sobj s' obj' = co2sobj s' obj"
+
+definition reserve:: "t_state \<Rightarrow> t_object set \<Rightarrow> t_state \<Rightarrow> bool"
+where
+  "reserve s objs s' \<equiv> \<forall> obj. alive s obj \<longrightarrow> alive s' obj \<and> co2sobj s' obj = co2sobj s obj"
+
+definition enrichable :: "t_state \<Rightarrow> t_object set \<Rightarrow> bool"
+where
+  "enrichable s objs \<equiv> \<exists> s'. valid s' \<and> s2ss s' = s2ss s \<and> enrich s objs s' \<and> reserve s objs s'"
+
+definition is_created :: "t_state \<Rightarrow> t_object \<Rightarrow> bool"
+where
+  "is_created s obj \<equiv> init_alive obj \<longrightarrow> deleted obj s"
+
+definition is_inited :: "t_state \<Rightarrow> t_object \<Rightarrow> bool"
+where
+  "is_inited s obj \<equiv> init_alive obj \<and> \<not> deleted obj s"
+
+lemma is_inited_eq_not_created:
+  "is_inited s obj = (\<not> is_created s obj)"
+by (auto simp:is_created_def is_inited_def)
+
+(* recorded in our static world *)
+fun recorded :: "t_object \<Rightarrow> 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: 
+  "\<lbrakk>valid s; \<forall> obj \<in> objs. alive s obj \<and> is_created s obj \<and> recorded obj\<rbrakk>
+   \<Longrightarrow> 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: "\<And> objs. \<forall> obj \<in> objs. alive s obj \<and> is_created s obj \<and> recorded obj \<Longrightarrow> enrichable s objs"
+    and p2: "valid (e # s)" and p3: "\<forall>obj\<in>objs. alive (e # s) obj \<and> is_created (e # s) obj \<and> 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) \<Longrightarrow> (O_proc p) \<notin> 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 \<Rightarrow> bool"
+where
+  "potential_trace s \<equiv> \<forall> obj. alive s obj \<and> is_created s obj \<longrightarrow> 
+      (\<exists> s' obj'. valid s' \<and> s2ss s' = ss \<and> obj' \<noteq> obj \<and> co2sobj s' obj = co2sobj s' obj)
+     "
+
+lemma s2d_main_general:
+  "ss \<in> static \<Longrightarrow> \<exists> s. valid s \<and> s2ss s = ss \<and> (\<forall> obj \<in> objs. alive s obj \<and> is_created s obj \<longrightarrow> (\<exists> s'. valid s' \<and> s2ss s' = ss \<and> (\<exists> obj'. obj' \<noteq> obj \<and> 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 \<in> static \<Longrightarrow> \<exists> s. valid s \<and> 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 \<in> tainted s \<Longrightarrow> co2sobj s obj = Some sobj \<Longrightarrow> tainted_s (s2ss s) sobj"
@@ -219,14 +325,6 @@
 
 (************** static \<rightarrow> dynamic ***************)
 
-lemma s2d_main:
-  "ss \<in> static \<Longrightarrow> \<exists> s. valid s \<and> 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:
   "\<lbrakk>x \<in> S; {x. P x} = S\<rbrakk> \<Longrightarrow> P x"