# HG changeset patch # User chunhan # Date 1369898937 -28800 # Node ID cc1e46225a8190df1339dfb0119f81302903c74b # Parent 7b5e9fbeaf932ee1609277aa93a92f5b6e3e6f6e shm attached at most once diff -r 7b5e9fbeaf93 -r cc1e46225a81 Co2sobj_prop.thy --- a/Co2sobj_prop.thy Thu May 30 09:15:19 2013 +0800 +++ b/Co2sobj_prop.thy Thu May 30 15:28:57 2013 +0800 @@ -859,10 +859,6 @@ apply (case_tac a, auto split:if_splits option.splits) done -definition - "fff x z = {xx. \ y. xx = (x,y) \ (x, y) \ z}" -thm fff_def - lemma procs_of_shm_prop2: "\(p, flag) \ procs_of_shm s h; valid s\ \ p \ current_procs s" apply (induct s arbitrary:p flag) apply (simp, drule init_procs_has_shm, simp) @@ -870,26 +866,61 @@ apply (case_tac a, auto split:if_splits option.splits) done +lemma procs_of_shm_prop3: "\(p, flag) \ procs_of_shm s h; (p, flag') \ procs_of_shm s h; valid s\ + \ flag = flag'" +apply (induct s arbitrary:p) +apply (simp) +apply( drule init_procs_has_shm, simp) +apply (frule vd_cons, frule vt_grant_os) +apply (case_tac a, auto split:if_splits option.splits) +done -lemma cph2spshs_createshm: + +lemma cph2spshs_attach: "valid (Attach p h flag # s) \ cph2spshs (Attach p h flag # s) = (cph2spshs s) (p := (case (ch2sshm s h) of Some sh \ cph2spshs s p \ {(sh, flag)} | _ \ cph2spshs s p) )" apply (frule vd_cons, frule vt_grant_os, rule ext) -unfolding cph2spshs_def using ch2sshm_other[where e = "Attach p h flag" and s = s] apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits -dest!:current_shm_has_sh' -) -thm current_shm_has_sshm -apply (rule set_eqI, rule iffI) -apply (case_tac xa, simp split:if_splits) -apply (case_tac xa, erule CollectE) -apply (simp, (erule conjE|erule bexE)+) -apply ( -apply auto +dest!:current_shm_has_sh' dest: procs_of_shm_prop1 simp:cph2spshs_def) +apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) +apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) +apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) +apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) +apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) +apply (case_tac "ha = h", simp, frule procs_of_shm_prop1, simp) +apply (rule_tac x = ha in exI, simp) +apply (rule_tac x = ha in exI, simp, drule procs_of_shm_prop1, simp, simp) +apply (rule_tac x = ha in exI, simp) +apply (frule procs_of_shm_prop1, simp, simp) +apply (rule impI, simp) +done + +lemma cph2spshs_detach: "valid (Detach p h # s) \ + cph2spshs (Detach p h # s) = (cph2spshs s) (p := + (case (ch2sshm s h) of + Some sh \ if (\ h'. h' \ h \ p \ procs_of_shm s h' \ ch2sshm s p h' = Some sh) + then cph2spshs s p else cph2spshs s p - {(sh, flag) | flag. (sh, flag) \ cph2spshs s p} + | _ \ cph2spshs s p) )" +apply (frule vd_cons, frule vt_grant_os, rule ext) +using ch2sshm_other[where e = "Detach p h" and s = s] +apply (auto split del:t_open_must_flag.splits t_open_option_flag.splits split add:if_splits option.splits +dest!:current_shm_has_sh' dest: procs_of_shm_prop1 simp:cph2spshs_def) +apply (rule_tac x = ha in exI, frule_tac h = ha in procs_of_shm_prop1, simp, simp) +apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) +apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) +apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) +apply (erule_tac x = ha in allE, frule procs_of_shm_prop1, simp, simp) +apply (case_tac "ha = h", simp, frule procs_of_shm_prop1, simp) +apply (rule_tac x = ha in exI, simp) +apply (rule_tac x = ha in exI, simp, drule procs_of_shm_prop1, simp, simp) +apply (rule_tac x = ha in exI, simp) +apply (frule procs_of_shm_prop1, simp, simp) +apply (rule impI, simp) +done lemma cph2spshs_other: "\valid (e # s); " diff -r 7b5e9fbeaf93 -r cc1e46225a81 Flask.thy --- a/Flask.thy Thu May 30 09:15:19 2013 +0800 +++ b/Flask.thy Thu May 30 15:28:57 2013 +0800 @@ -157,6 +157,7 @@ and init_fileflag_valid: "\ p fd flags. init_oflags_of_proc_fd p fd = Some flags \ \ f. init_file_of_proc_fd p fd = Some f" and init_procs_has_shm: "\ p h flag. (p,flag) \ init_procs_of_shm h \ p \ init_procs \ h \ init_shms" + and init_procshm_valid: "\ p h flag flag'. \(p,flag) \ init_procs_of_shm h; (p, flag') \ init_procs_of_shm h\ \ flag = flag'" and init_msgs_distinct: "\ q. distinct (init_msgs_of_queue q)" and init_msgs_valid: "\ m q. m \ set (init_msgs_of_queue q) \ q \ init_msgqs" @@ -426,7 +427,7 @@ | "current_shms (DeleteShM p s # \) = (current_shms \) - {s}" | "current_shms (e # \) = current_shms \ " -fun procs_of_shm :: "t_state \ t_shm \ (t_process \ t_shm_attach_flag) set" +fun procs_of_shm : "t_state \ t_shm \ (t_process \ t_shm_attach_flag) set" where "procs_of_shm [] = init_procs_of_shm" | "procs_of_shm (CreateShM p h # \) =