diff -r 9b42765ce554 -r ced0fcfbcf8e Current_prop.thy --- a/Current_prop.thy Tue Jun 04 15:51:02 2013 +0800 +++ b/Current_prop.thy Thu Jun 06 08:00:20 2013 +0800 @@ -1,7 +1,8 @@ +(*<*) theory Current_prop imports Main Flask_type Flask My_list_prefix Init_prop Valid_prop Delete_prop begin -(*<*) +(*>*) context flask begin @@ -46,6 +47,38 @@ "\p \ current_procs s; valid s\ \ \ (\ deleted (O_proc p) s \ p \ init_procs)" using not_deleted_init_proc by auto +lemma has_same_inode_in_current: + "\has_same_inode s f f'; valid s\ \ f \ current_files s \ f' \ current_files s" +by (auto simp add:has_same_inode_def current_files_def) + +lemma has_same_inode_prop1: + "\has_same_inode s f f'; is_file s f; valid s\ \ is_file s f'" +by (auto simp:has_same_inode_def is_file_def) + +lemma has_same_inode_prop1': + "\has_same_inode s f f'; is_file s f'; valid s\ \ is_file s f" +by (auto simp:has_same_inode_def is_file_def) + +lemma has_same_inode_prop2: + "\has_same_inode s f f'; file_of_proc_fd s p fd = Some f; valid s\ \ is_file s f'" +apply (drule has_same_inode_prop1) +apply (simp add:file_of_pfd_is_file, simp+) +done + +lemma has_same_inode_prop2': + "\has_same_inode s f f'; file_of_proc_fd s p fd = Some f'; valid s\ \ is_file s f" +apply (drule has_same_inode_prop1') +apply (simp add:file_of_pfd_is_file, simp+) +done + +lemma tobj_in_init_alive: + "tobj_in_init obj \ init_alive obj" +by (case_tac obj, auto) + +lemma tobj_in_alive: + "tobj_in_init obj \ alive [] obj" +by (case_tac obj, auto simp:is_file_nil) + end end \ No newline at end of file