Dynamic2static.thy
changeset 19 ced0fcfbcf8e
parent 1 7d9c0ed02b56
child 20 e2c6af3ccb0d
--- a/Dynamic2static.thy	Tue Jun 04 15:51:02 2013 +0800
+++ b/Dynamic2static.thy	Thu Jun 06 08:00:20 2013 +0800
@@ -1,5 +1,5 @@
 theory Dynamic2static
-imports Main Flask Static Init_prop Valid_prop
+imports Main Flask Static Init_prop Valid_prop Tainted_prop Delete_prop
 begin
 
 context tainting_s begin
@@ -15,30 +15,9 @@
 *)
 sorry
 
-lemma is_file_has_sfile: "is_file s f \<Longrightarrow> \<exists> sf. cf2sfile s f True = Some sf"
-sorry
-
-lemma is_dir_has_sfile: "is_dir s f \<Longrightarrow> \<exists> sf. cf2sfile s f False = Some sf"
-sorry
-
-lemma is_file_imp_alive: "is_file s f \<Longrightarrow> alive s (O_file f)"
-sorry
-
-
 lemma d2s_main':
   "\<lbrakk>alive s obj; co2sobj s obj= Some sobj\<rbrakk> \<Longrightarrow> sobj \<in> (s2ss s)"
-apply (induct s)
-apply (simp add:s2ss_def)
-apply (rule_tac x = obj in exI, simp)
-sorry
-
-lemma tainted_prop1:
-  "obj \<in> tainted s \<Longrightarrow> alive s obj"
-sorry
-
-lemma tainted_prop2:
-  "obj \<in> tainted s \<Longrightarrow> valid s"
-sorry
+by (simp add:s2ss_def, rule_tac x = obj in exI, simp)
 
 lemma alive_has_sobj:
   "\<lbrakk>alive s obj; valid s\<rbrakk> \<Longrightarrow> \<exists> sobj. co2sobj s obj = Some sobj"
@@ -46,52 +25,50 @@
 
 lemma t2ts:
   "obj \<in> tainted s \<Longrightarrow> co2sobj s obj = Some sobj \<Longrightarrow> tainted_s (s2ss s) sobj"
-apply (frule tainted_prop1, frule tainted_prop2)
+apply (frule tainted_in_current, frule tainted_is_valid)
 apply (simp add:s2ss_def)
 apply (case_tac sobj, simp_all)
-apply (case_tac [!] obj, simp_all split:option.splits)
+apply (case_tac [!] obj, simp_all split:option.splits if_splits)
 apply (rule_tac x = "O_proc nat" in exI, simp)
 apply (rule_tac x = "O_file list" in exI, simp)
-defer defer defer
+apply (drule dir_not_tainted, simp)
+apply (drule msgq_not_tainted, simp)
+apply (drule shm_not_tainted, simp)
 apply (case_tac prod1, simp, case_tac prod2, clarsimp)
 apply (rule conjI)
 apply (rule_tac x = "O_msgq nat1" in exI, simp)
-sorry (* doable, need properties about cm2smsg and cq2smsgq *)
+apply (rule conjI) defer
+apply (simp add:cm2smsg_def split:option.splits) 
+sorry
 
 lemma delq_imp_delqm:
   "deleted (O_msgq q) s \<Longrightarrow> deleted (O_msg q m) s"
 apply (induct s, simp)
 by (case_tac a, auto)
 
-lemma undel_init_file_remains:
-  "\<lbrakk>is_init_file f; \<not> deleted (O_file f) s\<rbrakk> \<Longrightarrow> is_file s f"
-sorry
-
-
 theorem static_complete: 
   assumes undel: "undeletable obj" and tbl: "taintable obj"
   shows "taintable_s obj"
 proof-
   from tbl obtain s where tainted: "obj \<in> tainted s"
     by (auto simp:taintable_def)
-  hence vs: "valid s" by (simp add:tainted_prop2)
+  hence vs: "valid s" by (simp add:tainted_is_valid)
   hence static: "s2ss s \<in> static" using d2s_main by auto
   from tainted have alive: "alive s obj" 
-    using tainted_prop1 by auto
+    using tainted_in_current by auto
   then obtain sobj where sobj: "co2sobj s obj = Some sobj"
     using vs alive_has_sobj by blast
   from undel vs have "\<not> deleted obj s" and init_alive: "init_alive obj" 
     by (auto simp:undeletable_def)
   with vs sobj have "init_obj_related sobj obj"
     apply (case_tac obj, case_tac [!] sobj)
-    apply (auto split:option.splits if_splits simp:cp2sproc_def ch2sshm_def cq2smsgq_def cm2smsg_def)
-    apply (frule undel_init_file_remains, simp, drule is_file_has_sfile, erule exE)
+    apply (auto split:option.splits if_splits simp:cp2sproc_def ch2sshm_def cq2smsgq_def cm2smsg_def delq_imp_delqm)
+    apply (frule not_deleted_init_file, simp+) (*
+apply (drule is_file_has_sfile, erule exE)
     apply (rule_tac x = sf in bexI)
     apply (case_tac list, auto split:option.splits simp:is_init_file_props)[1]
-    apply (simp add:same_inode_files_def cfs2sfiles_def)
-    apply (rule_tac x = list in exI, simp)
-    apply (case_tac list, auto split:option.splits simp:is_init_dir_props delq_imp_delqm)
-    done
+    apply (simp add:same_inode_files_def cfs2sfiles_def) *)  
+    sorry
   with tainted t2ts init_alive sobj static
   show ?thesis unfolding taintable_s_def
     apply (rule_tac x = "s2ss s" in bexI, simp)
@@ -99,13 +76,45 @@
     done
 qed
 
+lemma cp2sproc_pi:
+  "\<lbrakk>cp2sproc s p = Some (Init p', sec, fds, shms); valid s\<rbrakk> \<Longrightarrow> p = p' \<and> \<not> deleted (O_proc p) s \<and> p \<in> init_procs"
+by (simp add:cp2sproc_def split:option.splits if_splits)
+
+lemma cq2smsgq_qi:
+  "\<lbrakk>cq2smsgq s q = Some (Init q', sec, sms); valid s\<rbrakk> \<Longrightarrow> q = q' \<and> \<not> deleted (O_msgq q) s \<and> q \<in> init_msgqs"
+by (simp add:cq2smsgq_def split:option.splits if_splits)
+
+lemma cm2smsg_mi:
+  "\<lbrakk>cm2smsg s q m = Some (Init m', sec, ttag); q \<in> init_msgqs; valid s\<rbrakk> 
+   \<Longrightarrow> m = m' \<and> \<not> deleted (O_msg q m) s \<and> m \<in> set (init_msgs_of_queue q) \<and> q \<in> init_msgqs"
+by (clarsimp simp add:cm2smsg_def split:if_splits option.splits)
+
+lemma ch2sshm_hi:
+  "\<lbrakk>ch2sshm s h = Some (Init h', sec); valid s\<rbrakk> \<Longrightarrow> h = h' \<and> \<not> deleted (O_shm h) s \<and> h \<in> init_shms"
+by (clarsimp simp:ch2sshm_def split:if_splits option.splits)
+
+lemma root_not_deleted:
+  "\<lbrakk>deleted (O_dir []) s; valid s\<rbrakk> \<Longrightarrow> False"
+apply (induct s, simp)
+apply (frule vd_cons, frule vt_grant_os, case_tac a, auto)
+done
+
+lemma cf2sfile_fi:
+  "\<lbrakk>cf2sfile s f = Some (Init f', sec, psecopt, asecs); valid s\<rbrakk> \<Longrightarrow> f = f' \<and> 
+     (if (is_file s f) then \<not> deleted (O_file f) s \<and> is_init_file f 
+      else \<not> deleted (O_dir f) s \<and> is_init_dir f)"
+apply (case_tac f)
+by (auto simp:sroot_def cf2sfile_def root_is_init_dir dest!:root_is_dir' root_not_deleted
+        split:if_splits option.splits) 
+
 lemma init_deled_imp_deled_s: 
   "\<lbrakk>deleted obj s; init_alive obj; sobj \<in> (s2ss s); valid s\<rbrakk> \<Longrightarrow> \<not> init_obj_related sobj obj"
-apply (induct s, simp)
-apply (frule vd_cons)
-apply (case_tac a, auto)
-(* need simpset for s2ss *)
-sorry
+apply (rule notI)
+apply (clarsimp simp:s2ss_def)
+apply (case_tac obj, case_tac [!] obja, case_tac sobj)
+apply (auto split:option.splits if_splits dest!:cp2sproc_pi cq2smsgq_qi ch2sshm_hi cm2smsg_mi)
+
+done
 
 lemma deleted_imp_deletable_s:
   "\<lbrakk>deleted obj s; init_alive obj; valid s\<rbrakk> \<Longrightarrow> deletable_s obj"