diff -r 34d01e9a772e -r 7d9c0ed02b56 Co2sobj_prop.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Co2sobj_prop.thy Fri May 03 08:20:21 2013 +0100 @@ -0,0 +1,494 @@ +(*<*) +theory Co2sobj_prop +imports Main Flask Flask_type Static Static_type Sectxt_prop Init_prop Current_files_prop Current_sockets_prop +begin +(*<*) + +context tainting_s begin + +(************ simpset for cf2sfile ***************) + +declare get_parentfs_ctxts.simps [simp del] + +lemma cf2sfile_open_none: + "valid (Open p f flag fd None # s) \ cf2sfile (Open p f flag fd None # s) f b = cf2sfile s f b" +apply (frule vd_cons, frule vt_grant_os) +apply (frule noroot_events) +by (auto split:if_splits option.splits dest!:current_has_sec' + simp:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps + get_parentfs_ctxts_simps) + +lemma cf2sfile_open_some1: + "\valid (Open p f flag fd (Some inum) # s); f' \ current_files s\ + \ cf2sfile (Open p f flag fd (Some inum) # s) f' = cf2sfile s f'" +apply (frule vd_cons, frule vt_grant_os, frule noroot_events, rule ext) +apply (induct f') +apply (auto split:if_splits option.splits dest!:current_has_sec' + simp:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps + get_parentfs_ctxts_simps) +done + +lemma cf2sfile_open_some2: + "\valid (Open p f flag fd (Some inum) # s); is_file s f'\ + \ cf2sfile (Open p f flag fd (Some inum) # s) f' True = cf2sfile s f' True" +apply (frule vd_cons, drule is_file_in_current) +by (simp add:cf2sfile_open_some1) + +lemma cf2sfile_open_some3: + "\valid (Open p f flag fd (Some inum) # s); is_dir s f'\ + \ cf2sfile (Open p f flag fd (Some inum) # s) f' False = cf2sfile s f' False" +apply (frule vd_cons, drule is_dir_in_current) +by (simp add:cf2sfile_open_some1) + +lemma cf2sfile_open_some: + "valid (Open p f flag fd (Some inum) # s) \ cf2sfile (Open p f flag fd (Some inum) # s) f True = ( + case (parent f) of + Some pf \ (case (sectxt_of_obj (Open p f flag fd (Some inum) # s) (O_file f), sectxt_of_obj s (O_dir pf), + get_parentfs_ctxts s pf) of + (Some sec, Some psec, Some asecs) \ Some (Created, sec, Some psec, set asecs) + | _ \ None) + | None \ None)" +apply (frule vd_cons, frule vt_grant_os, frule noroot_events) +apply (auto split:if_splits option.splits dest!:current_has_sec' + simp:cf2sfile_def is_file_simps is_dir_simps current_files_simps sectxt_of_obj_simps + get_parentfs_ctxts_simps) + + + +lemma cf2sfile_open: + "valid (Open p f flag fd opt # s) \ cf2sfile (Open p f flag fd opt # s) = (\ f' b. + if (opt = None) then cf2sfile s f' b + else if (f' = f \ b = True) + then (case (parent f) of + Some pf \ (case (sectxt_of_obj s (O_file f), sectxt_of_obj s (O_dir pf), + get_parentfs_ctxts s pf) of + (Some sec, Some psec, Some asecs) \ + Some (if (\ deleted (O_file f) s \ f \ init_files) + then Init f else Created, sec, Some psec, set asecs) + | _ \ None) + | None \ None) + else cf2sfile s f' b)" +apply (frule vd_cons, frule vt_grant_os, rule ext, rule ext) +apply (auto split:if_splits option.splits + simp:sectxt_of_obj_simps) + + + + + + + + + + + + +lemma cf2sfile_keep_path: "\f \ f'; \ \ vt rc_cs\ \ cf2sfile \ f \ cf2sfile \ f'" +apply (induct f', simp add:no_junior_def) +apply (case_tac "f = a # f'", simp add:cf2sfile.simps) +apply (drule no_junior_noteq, simp, simp add:cf2sfile.simps) +done + +lemma ckp'_aux: "\ f \ a # f' \ \ f \ f'" +by (auto simp:no_junior_def) + +lemma cf2sfile_keep_path': "\\ f \ f'; \ \ vt rc_cs; f \ current_files \; f' \ current_files \\ \ \ cf2sfile \ f \ cf2sfile \ f'" +apply (induct f', simp add:no_junior_def cf2sfile.simps, rule notI, drule cf2sfile_root_file, simp) +apply (case_tac "f = a # f'", simp add:cf2sfile.simps) +apply (rule notI) +apply (frule ckp'_aux, simp, frule parentf_in_current', simp+) +apply (case_tac "cf2sfile \ f = cf2sfile \ (a # f')", drule cf2sfile_inj, simp+) +apply (simp add:cf2sfile.simps) +by (drule no_junior_noteq, simp+) + +lemma cf2sfile_keep_path'': "\cf2sfile \ f \ cf2sfile \ f'; \ \ vt rc_cs; f \ current_files \; f' \ current_files \\ \ f \ f'" +using cf2sfile_keep_path' +by (auto) + +lemma cf2sfile_open_some': "\f \ current_files \; Open p f' flags fd (Some inum) # \ \ vt rc_cs\ \ cf2sfile (Open p f' flags fd (Some inum) # \) f = cf2sfile \ f" +apply (frule vt_cons') +apply (drule vt_grant_os) +apply (induct f) +apply (simp add:cf2sfile.simps)+ +apply (frule parentf_in_current', simp) +apply (auto simp:os_grant.simps index_of_file.simps split:option.splits) +done + +lemma cf2sfile_open_some: "\Open p f flags fd (Some inum) # \ \ vt rc_cs; parent f = Some pf\ + \ cf2sfile (Open p f flags fd (Some inum) # \) f = (SCrea (Suc (length \)) # (cf2sfile \ pf))" +apply (case_tac f, simp) +apply (frule vt_grant_os) +apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_open_some') +done + +lemma cf2sfile_open_none: "cf2sfile (Open p f flags fd None # \) f' = cf2sfile \ f'" +apply (induct f') +by (simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits)+ + +lemma cf2sfile_open: "\Open p f flags fd opt # \ \ vt rc_cs; f' \ current_files (Open p f flags fd opt # \)\ \ + cf2sfile (Open p f flags fd opt # \) f' = ( + if (opt = None) then cf2sfile \ f' + else if (f' = f) then (case (parent f) of + Some pf \ (SCrea (Suc (length \)) # (cf2sfile \ pf)) + | _ \ []) + else cf2sfile \ f' )" +apply (frule vt_grant_os) +by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_open_none cf2sfile_open_some cf2sfile_open_some' split:if_splits option.splits) + +lemma cf2sfile_mkdir_some': "\Mkdir p f' inum # \ \ vt rc_cs; f \ current_files \\ \ cf2sfile (Mkdir p f' inum # \) f = cf2sfile \ f" +apply (frule vt_cons', drule vt_grant_os) +apply (induct f, (simp add:cf2sfile.simps)+) +apply (frule parentf_in_current', simp) +apply (auto simp:os_grant.simps index_of_file.simps split:option.splits) +done + +lemma cf2sfile_mkdir_some: "\Mkdir p f inum # \ \ vt rc_cs; parent f = Some pf\ + \ cf2sfile (Mkdir p f inum # \) f = (SCrea (Suc (length \)) # (cf2sfile \ pf))" +apply (case_tac f, simp) +apply (frule vt_grant_os) +apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_mkdir_some') +done + +lemma cf2sfile_mkdir: "\Mkdir p f inum # \ \ vt rc_cs; f' \ current_files (Mkdir p f inum # \)\ \ + cf2sfile (Mkdir p f inum # \) f' = ( + if (f' = f) then (case (parent f) of + Some pf \ (SCrea (Suc (length \)) # (cf2sfile \ pf)) + | _ \ []) + else cf2sfile \ f' ) " +apply (frule vt_grant_os) +by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_mkdir_some cf2sfile_mkdir_some' split:if_splits option.splits) + +lemma cf2sfile_linkhard_none: "\LinkHard p f\<^isub>1 f\<^isub>2 # \ \ vt rc_cs; f \ current_files \\ \ cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \) f = cf2sfile \ f" +apply (frule vt_cons') +apply (drule vt_grant_os) +apply (induct f) +apply (simp add:cf2sfile.simps)+ +apply (frule parentf_in_current', simp) +apply (auto simp:os_grant.simps index_of_file.simps split:option.splits) +done + +lemma cf2sfile_linkhard_some: + "\LinkHard p f\<^isub>1 f\<^isub>2 # \ \ vt rc_cs; parent f\<^isub>2 = Some pf\<^isub>2\ \ cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \) f\<^isub>2 = (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>2))" +apply (case_tac f\<^isub>2, simp) +apply (frule vt_grant_os) +apply (simp add:cf2sfile.simps os_grant.simps current_files_simps index_of_file.simps cf2sfile_linkhard_none) +done + +lemma cf2sfile_linkhard: "\LinkHard p f\<^isub>1 f\<^isub>2 # \ \ vt rc_cs; f \ current_files (LinkHard p f\<^isub>1 f\<^isub>2 # \)\ \ + cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \) f = ( + if (f = f\<^isub>2) then (case (parent f\<^isub>2) of + Some pf\<^isub>2 \ SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>2) + | _ \ []) + else cf2sfile \ f )" +apply (frule vt_grant_os) +by (auto simp:os_grant.simps current_files_simps intro:cf2sfile_linkhard_none cf2sfile_linkhard_some split:if_splits option.splits) + + +lemma no_junior_aux: "\ f\<^isub>2 \ a # f \ \ f\<^isub>2 \ f" +by (auto simp:no_junior_def) + +lemma cf2sfile_rename_aux: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; f \ current_files \\ \ f \ f\<^isub>3 \ \ f\<^isub>3 \ f" +apply (frule vt_grant_os, simp add:os_grant.simps, (erule exE|erule conjE)+) +apply (rule conjI, rule notI, simp) +apply (rule notI, drule vt_cons', simp add:ancenf_in_current) +done + +lemma cf2sfile_rename'1: + "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; f \ current_files \; \ (f\<^isub>2 \ f)\ \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = cf2sfile \ f" +apply (frule vt_cons',frule vt_grant_os, induct f) +apply (simp add:cf2sfile.simps) +apply (frule parentf_in_current', simp, simp) +apply (frule no_junior_aux, simp) +apply (simp add:os_grant.simps, (erule exE|erule conjE)+) +apply (drule cf2sfile_rename_aux, simp, erule conjE) +apply (auto simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits) +done + +lemma cf2sfile_rename'2: + "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; \ (f\<^isub>3 \ f)\ \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = cf2sfile \ f" +apply (frule vt_cons', induct f) +by (auto simp add:index_of_file.simps cf2sfile.simps no_junior_def split:option.splits nat.splits) + +lemma cf2sfile_rename1: + "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; parent f\<^isub>3 = Some pf\<^isub>3\ \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f\<^isub>3 = SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)" +apply (case_tac f\<^isub>3, simp add:cf2sfile.simps) +apply (auto dest!:cf2sfile_rename'2 simp add:no_junior_def cf2sfile.simps index_of_file.simps split:option.splits) +done + +lemma index_of_file_rename1: "\f\<^isub>2 \ f; f \ f\<^isub>2\ \ index_of_file (Rename p f\<^isub>2 f\<^isub>3 # \) (file_after_rename f\<^isub>2 f\<^isub>3 f) = index_of_file \ f" +apply (clarsimp simp add:index_of_file.simps split:option.splits) +by (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop3, simp, erule conjE, simp add:file_renaming_prop5) + +lemma cf2sfile_rename2: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; (file_before_rename f\<^isub>2 f\<^isub>3 f) \ current_files \; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>3 \ f\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ (file_before_rename f\<^isub>2 f\<^isub>3 f))" +apply (induct f, simp add:no_junior_def) +apply (case_tac "a # f = f\<^isub>3", simp) +apply (drule cf2sfile_rename1, simp, simp add:file_renaming_prop0' file_renaming_prop0) + +apply (frule no_junior_noteq, simp, simp) +apply (frule_tac file_renaming_prop1') +apply (frule_tac f = f\<^isub>2 and \ = \ in cf2sfile_keep_path, simp add:vt_cons') +apply (frule_tac f\<^isub>2 = f\<^isub>2 and a = a in file_renaming_prop8') +apply (frule_tac f\<^isub>2 = f\<^isub>2 and a = a in file_renaming_prop6') +apply (frule_tac f = "file_before_rename f\<^isub>2 f\<^isub>3 (a # f)" and \ = \ and f\<^isub>2 = f\<^isub>2 and f\<^isub>3 = f\<^isub>3 and p = p in index_of_file_rename1, simp add:file_before_rename_def) +apply (drule_tac f\<^isub>3 = f\<^isub>3 and f\<^isub>1 = f and a = a and f\<^isub>2 = f\<^isub>2 in file_renaming_prop9', simp) +apply (drule parentf_in_current', simp add:vt_cons') +apply (simp add:cf2sfile.simps) +apply (erule file_renaming_prop6[THEN sym]) +done + +lemma cf2sfile_rename2': "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; f \ current_files \; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>2 \ f\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ f)" +apply (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop5) +apply (frule_tac f\<^isub>3 = f\<^isub>3 in file_renaming_prop1) +apply (drule_tac f = "file_after_rename f\<^isub>2 f\<^isub>3 f" in cf2sfile_rename2, simp+) +done + +lemma cf2sfile_rename3: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; f \ current_files \; parent f\<^isub>3 = Some pf\<^isub>3\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ f)" +apply (case_tac "f\<^isub>2 \ f") +apply (rule cf2sfile_rename2', simp+) +apply (frule vt_grant_os) +apply (frule_tac \ = \ in cf2sfile_keep_path', simp add:vt_cons', simp add:os_grant.simps, simp) +apply (simp add:file_after_rename_def) +by (rule cf2sfile_rename'1, simp+) + +lemma cf2sfile_rename: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; f \ current_files (Rename p f\<^isub>2 f\<^isub>3 # \)\ \ + cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = ( + if (f\<^isub>3 \ f) then (case (parent f\<^isub>3) of + Some pf\<^isub>3 \ file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ (file_before_rename f\<^isub>2 f\<^isub>3 f)) + | _ \ []) + else cf2sfile \ f )" +apply (frule vt_grant_os) +apply (case_tac "f = f\<^isub>3", clarsimp simp:os_grant.simps, drule cf2sfile_rename1, simp+, simp add:file_renaming_prop0' file_renaming_prop0) + +apply (auto simp:os_grant.simps current_files_simps intro:cf2sfile_rename'2 cf2sfile_rename1 split:if_splits option.splits) +apply (rule cf2sfile_rename2, simp, drule rename_renaming_decom, simp+) +apply (drule_tac f\<^isub>2 = f\<^isub>2 and f\<^isub>1 = f\<^isub>1 and f\<^isub>3 = f\<^isub>3 in file_renaming_prop5, simp+) +done + +lemma cf2sfile_other: "\ + \ p f flags fd opt. e \ Open p f flags fd opt; + \ p f im. e \ Mkdir p f im; + \ p f\<^isub>1 f\<^isub>2. e \ LinkHard p f\<^isub>1 f\<^isub>2; + \ p f\<^isub>2 f\<^isub>3. e \ Rename p f\<^isub>2 f\<^isub>3\ \ cf2sfile (e # \) f' = cf2sfile \ f'" +apply (induct f', simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits) +apply (case_tac e, auto simp add:cf2sfile.simps index_of_file.simps split:option.splits nat.splits) +done + +lemmas cf2sfile_nil = init_cf2sfile + +lemma cf2sfile_nil': "f \ current_files [] \ cf2sfile [] f = map SInit f" +by (simp add:cf2sfile_nil current_files_simps) + +lemmas cf2sfile_simps = cf2sfile_nil cf2sfile_nil' cf2sfile_open cf2sfile_mkdir cf2sfile_linkhard cf2sfile_rename cf2sfile_other + +lemmas cf2sfile_simpss = cf2sfile_nil cf2sfile_nil' cf2sfile_open_some' cf2sfile_open_some cf2sfile_open_none cf2sfile_open cf2sfile_mkdir_some' cf2sfile_mkdir_some + cf2sfile_mkdir cf2sfile_linkhard_none cf2sfile_linkhard_some cf2sfile_linkhard cf2sfile_rename'1 cf2sfile_rename'2 cf2sfile_rename1 + cf2sfile_rename2 cf2sfile_rename2' cf2sfile_rename3 cf2sfile_rename cf2sfile_other + +lemma cf2sfile_open_some'_inum: "\Open p f' flags fd (Some inum) # \ \ vt rc_cs; inum_of_file \ f = Some im\ \ cf2sfile (Open p f' flags fd (Some inum) # \) f = cf2sfile \ f" +by (simp add:cf2sfile_open_some' current_files_def) + +lemma cf2sfile_mkdir_some'_inum: "\Mkdir p f' inum # \ \ vt rc_cs; inum_of_file \ f = Some im\ \ cf2sfile (Mkdir p f' inum # \) f = cf2sfile \ f" +by (simp add:cf2sfile_mkdir_some' current_files_def) + +lemma cf2sfile_linkhard_none_inum: "\LinkHard p f\<^isub>1 f\<^isub>2 # \ \ vt rc_cs; inum_of_file \ f = Some im\ \ cf2sfile (LinkHard p f\<^isub>1 f\<^isub>2 # \) f = cf2sfile \ f" +by (simp add:cf2sfile_linkhard_none current_files_def) + +lemma cf2sfile_rename'1_inum: + "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; inum_of_file \ f = Some im ; \ (f\<^isub>2 \ f)\ \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = cf2sfile \ f" +by (simp add:cf2sfile_rename'1 current_files_def) + +lemma cf2sfile_rename2_inum: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; inum_of_file \ (file_before_rename f\<^isub>2 f\<^isub>3 f) = Some im; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>3 \ f\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) f = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ (file_before_rename f\<^isub>2 f\<^isub>3 f))" +by (simp add:cf2sfile_rename2 current_files_def) + +lemma cf2sfile_rename2'_inum: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; inum_of_file \ f = Some im; parent f\<^isub>3 = Some pf\<^isub>3; f\<^isub>2 \ f\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ f)" +by (simp add:cf2sfile_rename2' current_files_def) + +lemma cf2sfile_rename3_inum: "\Rename p f\<^isub>2 f\<^isub>3 # \ \ vt rc_cs; inum_of_file \ f = Some im; parent f\<^isub>3 = Some pf\<^isub>3\ + \ cf2sfile (Rename p f\<^isub>2 f\<^isub>3 # \) (file_after_rename f\<^isub>2 f\<^isub>3 f) = file_after_rename (cf2sfile \ f\<^isub>2) (SCrea (Suc (length \)) # (cf2sfile \ pf\<^isub>3)) (cf2sfile \ f)" +by (simp add:cf2sfile_rename3 current_files_def) + +lemma cf2sfile_nil'_inum: "inum_of_file [] f = Some im \ cf2sfile [] f = map SInit f" +by (simp add:cf2sfile_nil' current_files_def) + +lemmas cf2sfile_simps_inum = cf2sfile_nil cf2sfile_nil'_inum cf2sfile_open_some'_inum cf2sfile_open_some cf2sfile_open_none cf2sfile_open cf2sfile_mkdir_some'_inum cf2sfile_mkdir_some + cf2sfile_mkdir cf2sfile_linkhard_none_inum cf2sfile_linkhard_some cf2sfile_linkhard cf2sfile_rename'1_inum cf2sfile_rename'2 cf2sfile_rename1 + cf2sfile_rename2_inum cf2sfile_rename2'_inum cf2sfile_rename3_inum cf2sfile_rename cf2sfile_other + +(*************** cp2sproc simpset *********************) + +lemma cp2sproc_nil: "p \ init_processes \ cp2sproc [] p = SInit p" +by (simp add:cp2sproc_def index_of_proc.simps) + +lemma cp2sproc_nil': "p \ current_procs [] \ cp2sproc [] p = SInit p" +by (simp add:cp2sproc_nil current_procs.simps) + +lemma cp2sproc_clone: "cp2sproc (Clone p p' # \) p'' = ( + if (p'' = p') then SCrea (Suc (length \)) + else cp2sproc \ p'' )" +by (auto simp:cp2sproc_def index_of_proc.simps d2s_aux.simps) + +lemma cp2sproc_other: "\ p p'. e \ Clone p p' \ cp2sproc (e # \) p'' = cp2sproc \ p''" +apply (case_tac e) +by (auto simp:cp2sproc_def index_of_proc.simps d2s_aux.simps) + +lemmas cp2sproc_simps = cp2sproc_nil cp2sproc_nil' cp2sproc_clone cp2sproc_other + +(******************** ch2sshm simpset **************************) + +lemma ch2sshm_nil: "h \ init_shms \ ch2sshm [] h = SInit h" +by (simp add:ch2sshm_def index_of_shm.simps) + +lemma ch2sshm_nil': "h \ current_shms [] \ ch2sshm [] h = SInit h" +by (simp add:ch2sshm_nil current_shms.simps) + +lemma ch2sshm_createshm: "ch2sshm (CreateShM p h # \) h' = (if (h' = h) then SCrea (Suc (length \)) else ch2sshm \ h')" +by (simp add:ch2sshm_def index_of_shm.simps d2s_aux.simps) + +lemma ch2sshm_other: "\ p h. e \ CreateShM p h \ ch2sshm (e # \) h' = ch2sshm \ h'" +apply (case_tac e) +by (auto simp add:ch2sshm_def index_of_shm.simps d2s_aux.simps) + +lemmas ch2sshm_simps = ch2sshm_nil ch2sshm_nil' ch2sshm_createshm ch2sshm_other + +(********************* cm2smsg simpset ***********************) + +lemma cm2smsg_nil: "m \ init_msgs \ cm2smsg [] m = SInit m" +by (simp add:cm2smsg_def index_of_msg.simps) + +lemma cm2smsg_nil': "m \ current_msgs [] \ cm2smsg [] m = SInit m" +by (simp add:cm2smsg_nil current_msgs.simps) + +lemma cm2smsg_createmsg: "cm2smsg (CreateMsg p m # \) m' = (if (m' = m) then SCrea (Suc (length \)) else cm2smsg \ m')" +by (simp add:cm2smsg_def index_of_msg.simps d2s_aux.simps) + +lemma cm2smsg_other: "\ p m. e \ CreateMsg p m \ cm2smsg (e # \) m' = cm2smsg \ m'" +apply (case_tac e) +by (auto simp:cm2smsg_def index_of_msg.simps d2s_aux.simps) + +lemmas cm2smsg_simps = cm2smsg_nil cm2smsg_nil' cm2smsg_createmsg cm2smsg_other + +(********************** cfd2fd_s simpset ******************************) + +lemma cfd2fd_s_nil: "fd \ init_fds_of_proc p \ cfd2fd_s [] p fd = SInit fd" +by (simp add:cfd2fd_s_def index_of_fd.simps) + +lemma cfd2fd_s_nil': "fd \ current_proc_fds [] p \ cfd2fd_s [] p fd = SInit fd" +by (simp add:cfd2fd_s_nil current_proc_fds.simps) + +lemma cfd2fd_s_open: "cfd2fd_s (Open p f flags fd opt # \) p' fd' = ( + if (p = p') then (if (fd = fd') then SCrea (Suc (length \)) + else cfd2fd_s \ p' fd') + else cfd2fd_s \ p' fd' )" +by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps) + +lemma cfd2fd_s_createsock: "cfd2fd_s (CreateSock p af st fd im # \) p' fd' = ( + if (p = p') then (if (fd = fd') then SCrea (Suc (length \)) + else cfd2fd_s \ p' fd') + else cfd2fd_s \ p' fd' )" +by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps) + +lemma cfd2fd_s_accept: "cfd2fd_s (Accept p fd addr port fd' im # \) p' fd'' = ( + if (p' = p) then (if (fd'' = fd') then SCrea (Suc (length \)) + else cfd2fd_s \ p' fd'') + else cfd2fd_s \ p' fd'' )" +by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps) + +lemma cfd2fd_s_clone: "cfd2fd_s (Clone p p' # \) p'' fd = (if (p'' = p') then cfd2fd_s \ p fd else cfd2fd_s \ p'' fd)" +by (simp add:cfd2fd_s_def index_of_fd.simps d2s_aux.simps) + +lemma cfd2fd_s_other: "\\ p f flags fd opt. e \ Open p f flags fd opt; + \ p af st fd im. e \ CreateSock p af st fd im; + \ p fd addr port fd' im. e \ Accept p fd addr port fd' im; + \ p p'. e \ Clone p p'\ \ cfd2fd_s (e # \) p'' fd'' = cfd2fd_s \ p'' fd''" +by (case_tac e, auto simp:cfd2fd_s_def index_of_fd.simps d2s_aux.simps) + +lemmas cfd2fd_s_simps = cfd2fd_s_nil cfd2fd_s_nil' cfd2fd_s_open cfd2fd_s_createsock cfd2fd_s_accept cfd2fd_s_clone cfd2fd_s_other + +(************* cim2im_s simpset **************************) + +(* no such lemma +lemma cim2im_s_nil: "init_itag_of_inum im = Some tag \ cim2im_s [] im = SInit im" +by (simp add:cim2im_s_def) +*) + +lemma cim2im_s_open: "cim2im_s (Open p f flags fd (Some im) # \) im' = (if (im' = im) then SCrea (Suc (length \)) else cim2im_s \ im')" +by (simp add:cim2im_s_def) + +lemma cim2im_s_open': "cim2im_s (Open p f flags fd None # \) im = cim2im_s \ im" +by (simp add:cim2im_s_def) + +lemma cim2im_s_mkdir: "cim2im_s (Mkdir p f im # \) im' = (if (im' = im) then SCrea (Suc (length \)) else cim2im_s \ im')" +by (simp add:cim2im_s_def) + +lemma cim2im_s_createsock: "cim2im_s (CreateSock p sf st fd im # \) im' = (if (im' = im) then SCrea (Suc (length \)) else cim2im_s \ im')" +by (simp add:cim2im_s_def) + +lemma cim2im_s_accept: "cim2im_s (Accept p fd addr port fd' im # \) im' = (if (im' = im) then SCrea (Suc (length \)) else cim2im_s \ im')" +by (simp add:cim2im_s_def) + +lemma cim2im_s_other: "\\ p f flags fd opt. e \ Open p f flags fd opt; + \ p f im. e \ Mkdir p f im; + \ p sf st fd im. e \ CreateSock p sf st fd im; + \ p fd addr port fd' im. e \ Accept p fd addr port fd' im\ \ cim2im_s (e # \) im = cim2im_s \ im" +by (case_tac e, auto simp:cim2im_s_def) + +lemmas cim2im_s_simps = cim2im_s_open cim2im_s_open' cim2im_s_mkdir cim2im_s_createsock cim2im_s_accept cim2im_s_other + + +lemma cig2ig_s_simp: "cig2ig_s (e # \) tag = cig2ig_s \ tag" +apply (case_tac tag) +by auto + +(******************* cobj2sobj no Suc (length \) ***********************) + +lemma cf2sfile_le_len: "\cf2sfile \ f = SCrea (Suc (length \)) # spf; f \ current_files \; \ \ vt rc_cs\ \ False" +apply (case_tac f, (simp add:cf2sfile.simps d2s_aux.simps)+) +apply (case_tac "index_of_file \ (a # list)", (simp add:d2s_aux.simps)+) +by (drule index_of_file_le_length', simp+) + +lemma cf2sfile_le_len': "\SCrea (Suc (length \)) # spf \ cf2sfile \ f; f \ current_files \; \ \ vt rc_cs\ \ False" +apply (induct f) +apply (simp add:no_junior_def cf2sfile.simps d2s_aux.simps) +apply (case_tac "cf2sfile \ (a # f) = SCrea (Suc (length \)) # spf") +apply (drule_tac f = "a # f" in cf2sfile_le_len, simp+) +apply (simp only:cf2sfile.simps d2s_aux.simps) +apply (drule_tac no_junior_noteq, simp+) +apply (rule impI, erule impE, simp+) +apply (drule parentf_in_current', simp+) +done + +lemma cp2sproc_le_len: "cp2sproc \ p = SCrea (Suc (length \)) \ False" +apply (simp add:cp2sproc_def, case_tac "index_of_proc \ p") +apply (simp add:d2s_aux.simps)+ +by (drule index_of_proc_le_length', simp) + +lemma ch2sshm_le_len: "ch2sshm \ h = SCrea (Suc (length \)) \ False" +apply (simp add:ch2sshm_def, case_tac "index_of_shm \ h") +apply (simp add:d2s_aux.simps)+ +by (drule index_of_shm_le_length', simp) + +lemma cm2smsg_le_len: "cm2smsg \ m = SCrea (Suc (length \)) \ False" +apply (simp add:cm2smsg_def, case_tac "index_of_msg \ m") +apply (simp add:d2s_aux.simps)+ +by (drule index_of_msg_le_length', simp) + +lemma cim2im_s_le_len: "cim2im_s \ im = SCrea (Suc (length \)) \ False" +apply (simp add:cim2im_s_def, case_tac "inum2ind \ im") +apply (simp add:d2s_aux.simps)+ +by (drule inum2ind_le_length', simp) + +lemma cfd2fd_s_le_len: "cfd2fd_s \ p fd = SCrea (Suc (length \)) \ False" +apply (simp add:cfd2fd_s_def, case_tac "index_of_fd \ p fd") +apply (simp add:d2s_aux.simps)+ +by (drule index_of_fd_le_length', simp) + +end + +(*<*) +end +(*>*) \ No newline at end of file