diff -r 0753309adfc7 -r 6f9a588bcfc4 Static.thy --- a/Static.thy Thu Oct 24 09:42:35 2013 +0800 +++ b/Static.thy Wed Oct 30 08:18:40 2013 +0800 @@ -122,6 +122,122 @@ definition "init_static_state \ {sobj. \ obj. init_alive obj \ init_obj2sobj obj = Some sobj}" +(**************** translation from dynamic to static *******************) + +definition cf2sfile :: "t_state \ t_file \ t_sfile option" +where + "cf2sfile s f \ + case (parent f) of + None \ Some sroot + | Some pf \ if (is_file s f) + then (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 \ is_init_file f) then Init f else Created, sec, Some psec, set asecs) + | _ \ None) + else (case (sectxt_of_obj s (O_dir f), sectxt_of_obj s (O_dir pf), get_parentfs_ctxts s pf) of + (Some sec, Some psec, Some asecs) \ + Some (if (\ deleted (O_dir f) s \ is_init_dir f) then Init f else Created, sec, Some psec, set asecs) + | _ \ None)" + +definition cf2sfiles :: "t_state \ t_file \ t_sfile set" +where + "cf2sfiles s f \ {sf. \ f' \ (same_inode_files s f). cf2sfile s f' = Some sf}" + +(* here cf2sfile is passed with True, because, process' fds are only for files not dirs *) +definition cfd2sfd :: "t_state \ t_process \ t_fd \ t_sfd option" +where + "cfd2sfd s p fd \ + (case (file_of_proc_fd s p fd, flags_of_proc_fd s p fd, sectxt_of_obj s (O_fd p fd)) of + (Some f, Some flags, Some sec) \ (case (cf2sfile s f) of + Some sf \ Some (sec, flags, sf) + | _ \ None) + | _ \ None)" + + +definition cpfd2sfds :: "t_state \ t_process \ t_sfd set" +where + "cpfd2sfds s p \ {sfd. \ fd \ proc_file_fds s p. cfd2sfd s p fd = Some sfd}" + +definition ch2sshm :: "t_state \ t_shm \ t_sshm option" +where + "ch2sshm s h \ (case (sectxt_of_obj s (O_shm h)) of + Some sec \ + Some (if (\ deleted (O_shm h) s \ h \ init_shms) then Init h else Created, sec) + | _ \ None)" + +definition cph2spshs :: "t_state \ t_process \ t_sproc_sshm set" +where + "cph2spshs s p \ {(sh, flag)| sh flag h. (p, flag) \ procs_of_shm s h \ ch2sshm s h = Some sh}" + +definition cp2sproc :: "t_state \ t_process \ t_sproc option" +where + "cp2sproc s p \ (case (sectxt_of_obj s (O_proc p)) of + Some sec \ + Some (if (\ deleted (O_proc p) s \ p \ init_procs) then Init p else Created, sec, + cpfd2sfds s p, cph2spshs s p) + | _ \ None)" + +definition cm2smsg :: "t_state \ t_msgq \ t_msg \ t_smsg option" +where + "cm2smsg s q m \ (case (sectxt_of_obj s (O_msg q m)) of + Some sec \ + Some (if (\ deleted (O_msg q m) s \ m \ set (init_msgs_of_queue q)) then Init m else Created, + sec, O_msg q m \ tainted s) + | _ \ None)" + +fun cqm2sms:: "t_state \ t_msgq \ t_msg list \ (t_smsg list) option" +where + "cqm2sms s q [] = Some []" +| "cqm2sms s q (m#ms) = + (case (cqm2sms s q ms, cm2smsg s q m) of + (Some sms, Some sm) \ Some (sm#sms) + | _ \ None)" + +definition cq2smsgq :: "t_state \ t_msgq \ t_smsgq option" +where + "cq2smsgq s q \ (case (sectxt_of_obj s (O_msgq q), cqm2sms s q (msgs_of_queue s q)) of + (Some sec, Some sms) \ + Some (if (\ deleted (O_msgq q) s \ q \ init_msgqs) then Init q else Created, + sec, sms) + | _ \ None)" + +fun co2sobj :: "t_state \ t_object \ t_sobject option" +where + "co2sobj s (O_proc p) = + (case (cp2sproc s p) of + Some sp \ Some (S_proc sp (O_proc p \ tainted s)) + | _ \ None)" +| "co2sobj s (O_file f) = + Some (S_file (cf2sfiles s f) (O_file f \ tainted s))" +| "co2sobj s (O_dir f) = + (case (cf2sfile s f) of + Some sf \ Some (S_dir sf) + | _ \ None)" +| "co2sobj s (O_msgq q) = + (case (cq2smsgq s q) of + Some sq \ Some (S_msgq sq) + | _ \ None)" +| "co2sobj s (O_shm h) = + (case (ch2sshm s h) of + Some sh \ Some (S_shm sh) + | _ \ None)" +(* +| "co2sobj s (O_msg q m) = + (case (cq2smsgq s q, cm2smsg s q m) of + (Some sq, Some sm) \ Some (S_msg sq sm) + | _ \ None)" +*) +| "co2sobj s _ = None" + + +definition s2ss :: "t_state \ t_static_state" +where + "s2ss s \ {sobj. \ obj. alive s obj \ co2sobj s obj = Some sobj}" + + +(* ******************************************************** *) + + fun is_init_sfile :: "t_sfile \ bool" where "is_init_sfile (Init _, sec, psec,asec) = True" @@ -204,18 +320,6 @@ else (ss - {S_dir sf tag}) \ {S_dir sf' tag'})" *) -definition update_ss :: "t_static_state \ t_sobject \ t_sobject \ t_static_state" -where - "update_ss ss so so' \ if (is_many so) then ss \ {so'} else (ss - {so}) \ {so'}" - -definition add_ss :: "t_static_state \ t_sobject \ t_static_state" -where - "add_ss ss so \ ss \ {so}" - -definition del_ss :: "t_static_state \ t_sobject \ t_static_state" -where - "del_ss ss so \ if (is_many so) then ss else ss - {so}" - (* fun sparent :: "t_sfile \ t_sfile option" where @@ -487,147 +591,6 @@ | "\info_flow_sshm sp (pi,sec,fds,shms); info_flow_sproc_sshms shms shms'\ \ info_flow_sshm sp (pi',sec',fds',shms')" -definition update_ss_shms :: "t_static_state \ t_sproc \ bool \ t_static_state" -where - "update_ss_shms ss spfrom tag \ {sobj. \ sobj' \ ss. - (case sobj' of - S_proc sp tagp \ if (info_flow_sshm spfrom sp) - then if (is_many_sproc sp) - then (sobj = S_proc sp tagp \ sobj = S_proc sp (tagp \ tag)) - else (sobj = S_proc sp (tagp \ tag)) - else (sobj = S_proc sp tag) - | _ \ sobj = sobj')}" - - - -(* all reachable static states(sobjects set) *) -inductive_set static :: "t_static_state set" -where - s_init: "init_static_state \ static" -| s_execve: "\ss \ static; S_proc (pi, pctxt, fds, shms) tagp \ ss; S_file sfs tagf \ ss; - (fi,fsec,pfsec,asecs) \ sfs; npctxt_execve pctxt fsec = Some pctxt'; - grant_execve pctxt fsec pctxt'; search_check_s pctxt (fi,fsec,pfsec,asecs) True; - inherit_fds_check_s pctxt' fds'; fds' \ fds\ - \ (update_ss ss (S_proc (pi, pctxt, fds, shms) tagp) - (S_proc (pi, pctxt', fds', {}) (tagp \ tagf))) \ static" -| s_clone: "\ss \ static; S_proc (pi, pctxt, fds, shms) tagp \ ss; - permission_check pctxt pctxt C_process P_fork; - inherit_fds_check_s pctxt fds'; fds' \ fds; - inherit_shms_check_s pctxt shms'; shms' \ shms\ - \ (add_ss ss (S_proc (Created, pctxt, fds', shms') tagp)) \ static" -| s_kill: "\ss \ static; S_proc (pi, pctxt, fds, shms) tagp \ ss; - S_proc (pi', pctxt', fds', shms') tagp' \ ss; - permission_check pctxt pctxt' C_process P_sigkill\ - \ (del_ss ss (S_proc (pi', pctxt', fds', shms') tagp')) \ static" -| s_ptrace: "\ss \ static; S_proc sp tagp \ ss; S_proc sp' tagp' \ ss; - permission_check (sextxt_of_sproc sp) (sectxt_of_sproc sp') C_process P_ptrace\ - \ (update_ss_shms (update_ss_shms ss sp tagp) sp' tagp') \ static" -| s_exit: "\ss \ static; S_proc sp tagp \ ss\ \ (del_ss ss (S_proc sp tagp)) \ static" -| s_open: "\ss \ static; S_proc (pi, pctxt, fds, shms) tagp \ ss; S_file sfs tagf \ ss; sf \ sfs; - search_check_s pctxt sf True; \ is_creat_excl_flag flags; - oflags_check flags pctxt (sectxt_of_sfile sf); permission_check pctxt pctxt C_fd P_setattr\ - \ (update_ss ss (S_proc (pi, pctxt, fds, shms) tagp) - (S_proc (pi, pctxt, fds \ {(pctxt,flags,sf)}, shms) tagp)) \ static" -| s_open': "\ss \ static; S_proc (pi, pctxt, fds, shms) tagp \ ss; is_creat_excl_flag flags; - S_dir (pfi,fsec,pfsec,asecs) \ ss; search_check_s pctxt (pfi,fsec,pfsec,asecs) False; - nfsec = nfctxt_create pctxt fsec C_file; oflags_check flags pctxt nfsec; - permission_check pctxt fsec C_dir P_add_rename; permission_check pctxt pctxt C_fd P_setattr\ - \ (update_ss (add_ss ss (S_file {(Created, nfsec, Some fsec, asecs \ {fsec})} tagp)) - (S_proc (pi, pctxt, fds, shms) tagp) - (S_proc (pi, pctxt, fds \ {(pctxt, flags, (Created, nfsec, Some fsec, asecs \ {fsec}))}, shms) tagp) - ) \ static" -| S_readf: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; (fdctxt,flags,sf) \ fds; - permission_check pctxt fdctxt C_fd P_setattr; S_file sfs tagf \ ss; sf \ sfs; - permission_check pctxt (sectxt_of_sfile sf) C_file P_read; is_read_flag flags\ - \ (update_ss_shms ss (pi, pctxt,fds,shms) (tagp \ tagf)) \ static" -| S_writef: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; (fdctxt,flags,sf) \ fds; - permission_check pctxt fdctxt C_fd P_setattr; sf \ sfs; S_file sfs tagf \ ss; - permission_check pctxt (sectxt_of_sfile sf) C_file P_write; is_write_flag flags\ - \ (update_ss ss (S_file sfs tagf) (S_file sfs (tagp \ tagf))) \ static" -| S_unlink: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_file sfs tagf \ ss; - (Init f,fsec,Some pfsec,asecs) \ sfs; - search_check_s pctxt (Init f,fsec,Some pfsec,asecs) True; - permission_check pctxt fsec C_file P_unlink; - permission_check pctxt pfsec C_dir P_remove_name\ - \ ((ss - {S_file sfs tagf}) \ {S_file (sfs - {(Init f,fsec,Some pfsec,asecs)}) tagf}) \ static" -| S_rmdir: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; - S_dir (fi,fsec,Some pfsec,asecs) \ ss; - search_check_s pctxt (fi,fsec,Some pfsec,asecs) False; - permission_check pctxt fsec C_dir P_rmdir; - permission_check pctxt pfsec C_dir P_remove_name\ - \ (del_ss ss (S_dir (fi,fsec,Some pfsec,asecs))) \ static" -| S_mkdir: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_dir (fi,fsec,pfsec,asecs) \ ss; - search_check_s pctxt (fi,fsec,pfsec,asecs) False; - permission_check pctxt (nfctxt_create pctxt fsec C_dir) C_dir P_create; - permission_check pctxt fsec C_dir P_add_name\ - \ (add_ss ss (S_dir (Created,nfctxt_create pctxt fsec C_dir,Some fsec,asecs \ {fsec}))) \ static" -| s_link: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_dir (pfi,pfsec,ppfsec,asecs) \ ss; - S_file sfs tagf \ ss; sf \ sfs; nfsec = nfctxt_create pctxt pfsec C_file; - search_check_s pctxt (pfi,pfsec,ppfsec,asecs) False; search_check_s pctxt sf True; - permission_check pctxt (sectxt_of_sfile sf) C_file P_link; - permission_check pctxt pfsec C_dir P_add_name\ - \ (update_ss ss (S_file sfs tagf) - (S_file (sfs \ {(Created,nfsec,Some pfsec, asecs \ {pfsec})}) tagf)) \ static" -| s_trunc: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_file sfs tagf \ ss; sf \ sfs; - search_check_s pctxt sf True; permission_check pctxt (sectxt_of_sfile sf) C_file P_setattr\ - \ (update_ss ss (S_file sfs tagf) (S_file sfs (tagf \ tagp))) \ static" -(* -| s_rename: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_file sfs tagf \ ss; - (sf#spf') \ sfs; S_dir spf tagpf \ ss; \((sf#spf') \ (sf#spf)); - search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True; - sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt; - permission_check pctxt fctxt C_file P_rename; - permission_check pctxt pfctxt C_dir P_add_name; - ss_rename ss (sf#spf') (sf#spf) ss'; - ss_rename_no_same_fname ss (sf#spf') (sf#spf)\ - \ ss' \ static" -| s_rename': "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_dir (sf#spf') tagf \ ss; - S_dir spf tagpf \ ss; \((sf#spf') \ (sf#spf)); - search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True; - sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt; - permission_check pctxt fctxt C_dir P_reparent; - permission_check pctxt pfctxt C_dir P_add_name; - ss_rename ss (sf#spf') (sf#spf) ss'; - ss_rename_no_same_fname ss (sf#spf') (sf#spf)\ - \ ss' \ static" -*) -| s_createq: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; - permission_check pctxt pctxt C_msgq P_associate; - permission_check pctxt pctxt C_msgq P_create\ - \ (add_ss ss (S_msgq (Created,pctxt,[]))) \ static" -| s_sendmsg: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_msgq (qi,qctxt,sms) \ ss; - permission_check pctxt qctxt C_msgq P_enqueue; - permission_check pctxt qctxt C_msgq P_write; - permission_check pctxt pctxt C_msg P_create\ - \ (update_ss ss (S_msgq (qi,qctxt,sms)) - (S_msgq (qi,qctxt,sms @ [(Created, pctxt, tagp)]))) \ static" -| s_recvmsg: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; - S_msgq (qi,qctxt,(mi,mctxt,tagm)#sms) \ ss; - permission_check pctxt qctxt C_msgq P_read; - permission_check pctxt mctxt C_msg P_receive\ - \ (update_ss (update_ss_shms ss (pi,pctxt,fds,shms) (tagp \ tagm)) - (S_msgq (qi, qctxt, (mi, mctxt, tagm)#sms)) - (S_msgq (qi, qctxt, sms))) \ static" -| s_removeq: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_msgq (qi,qctxt,sms) \ ss; - permission_check pctxt qctxt C_msgq P_destroy\ - \ (del_ss ss (S_msgq (qi,qctxt,sms))) \ static" -| s_createh: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; - permission_check pctxt pctxt C_shm P_associate; - permission_check pctxt pctxt C_shm P_create\ - \ (add_ss ss (S_shm (Created, pctxt))) \ static" -| s_attach: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_shm (hi,hctxt) \ ss; - if flag = SHM_RDONLY then permission_check pctxt hctxt C_shm P_read - else (permission_check pctxt hctxt C_shm P_read \ - permission_check pctxt hctxt C_shm P_write)\ - \ (update_ss ss (S_proc (pi,pctxt,fds,shms) tagp) - (S_proc (pi,pctxt,fds,shms \ {((hi,hctxt),flag)}) tagp)) \ static" -| s_detach: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_shm sh \ ss; - (sh,flag) \ shms; \ is_many_sshm sh\ - \ (update_ss ss (S_proc (pi,pctxt,fds,shms) tagp) - (S_proc (pi,pctxt,fds,shms - {(sh,flag)}) tagp)) \ static" -| s_deleteh: "\ss \ static; S_proc (pi,pctxt,fds,shms) tagp \ ss; S_shm (hi,hctxt) \ ss; - permission_check pctxt hctxt C_shm P_destroy; \ is_many_sshm sh\ - \ (remove_sproc_sshm (del_ss ss (S_shm (hi,hctxt))) (hi,hctxt)) \ static" (* fun smsg_related :: "t_msg \ t_smsg list \ bool" @@ -672,144 +635,6 @@ *) | "init_obj_related _ _ = False" -fun tainted_s :: "t_static_state \ t_sobject \ bool" -where - "tainted_s ss (S_proc sp tag) = (S_proc sp tag \ ss \ tag = True)" -| "tainted_s ss (S_file sfs tag) = (S_file sfs tag \ ss \ tag = True)" -(* -| "tainted_s ss (S_msg (qi, sec, sms) (mi, secm, tag)) = - (S_msgq (qi, sec, sms) \ ss \ (mi,secm,tag) \ set sms \ tag = True)" -*) -| "tainted_s ss _ = False" - -(* -fun tainted_s :: "t_object \ t_static_state \ bool" -where - "tainted_s (O_proc p) ss = (\ sp. S_proc sp True \ ss \ sproc_related p sp)" -| "tainted_s (O_file f) ss = (\ sfs sf. S_file sfs True \ ss \ sf \ sfs \ sfile_related f sf)" -| "tainted_s (O_msg q m) ss = (\ sq. S_msgq sq \ ss \ smsgq_smsg_relatainted q m sq)" -| "tainted_s _ ss = False" -*) - -definition taintable_s :: "t_object \ bool" -where - "taintable_s obj \ \ ss \ static. \ sobj. tainted_s ss sobj \ init_obj_related sobj obj \ init_alive obj" - -definition deletable_s :: "t_object \ bool" -where - "deletable_s obj \ init_alive obj \ (\ ss \ static. \ sobj \ ss. \ init_obj_related sobj obj)" - -definition undeletable_s :: "t_object \ bool" -where - "undeletable_s obj \ init_alive obj \ (\ ss \ static. \ sobj \ ss. init_obj_related sobj obj)" - - -(**************** translation from dynamic to static *******************) - -definition cf2sfile :: "t_state \ t_file \ t_sfile option" -where - "cf2sfile s f \ - case (parent f) of - None \ Some sroot - | Some pf \ if (is_file s f) - then (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 \ is_init_file f) then Init f else Created, sec, Some psec, set asecs) - | _ \ None) - else (case (sectxt_of_obj s (O_dir f), sectxt_of_obj s (O_dir pf), get_parentfs_ctxts s pf) of - (Some sec, Some psec, Some asecs) \ - Some (if (\ deleted (O_dir f) s \ is_init_dir f) then Init f else Created, sec, Some psec, set asecs) - | _ \ None)" - -definition cf2sfiles :: "t_state \ t_file \ t_sfile set" -where - "cf2sfiles s f \ {sf. \ f' \ (same_inode_files s f). cf2sfile s f' = Some sf}" - -(* here cf2sfile is passed with True, because, process' fds are only for files not dirs *) -definition cfd2sfd :: "t_state \ t_process \ t_fd \ t_sfd option" -where - "cfd2sfd s p fd \ - (case (file_of_proc_fd s p fd, flags_of_proc_fd s p fd, sectxt_of_obj s (O_fd p fd)) of - (Some f, Some flags, Some sec) \ (case (cf2sfile s f) of - Some sf \ Some (sec, flags, sf) - | _ \ None) - | _ \ None)" - - -definition cpfd2sfds :: "t_state \ t_process \ t_sfd set" -where - "cpfd2sfds s p \ {sfd. \ fd \ proc_file_fds s p. cfd2sfd s p fd = Some sfd}" - -definition ch2sshm :: "t_state \ t_shm \ t_sshm option" -where - "ch2sshm s h \ (case (sectxt_of_obj s (O_shm h)) of - Some sec \ - Some (if (\ deleted (O_shm h) s \ h \ init_shms) then Init h else Created, sec) - | _ \ None)" - -definition cph2spshs :: "t_state \ t_process \ t_sproc_sshm set" -where - "cph2spshs s p \ {(sh, flag)| sh flag h. (p, flag) \ procs_of_shm s h \ ch2sshm s h = Some sh}" - -definition cp2sproc :: "t_state \ t_process \ t_sproc option" -where - "cp2sproc s p \ (case (sectxt_of_obj s (O_proc p)) of - Some sec \ - Some (if (\ deleted (O_proc p) s \ p \ init_procs) then Init p else Created, sec, - cpfd2sfds s p, cph2spshs s p) - | _ \ None)" - -definition cm2smsg :: "t_state \ t_msgq \ t_msg \ t_smsg option" -where - "cm2smsg s q m \ (case (sectxt_of_obj s (O_msg q m)) of - Some sec \ - Some (if (\ deleted (O_msg q m) s \ m \ set (init_msgs_of_queue q)) then Init m else Created, - sec, O_msg q m \ tainted s) - | _ \ None)" - -fun cqm2sms:: "t_state \ t_msgq \ t_msg list \ (t_smsg list) option" -where - "cqm2sms s q [] = Some []" -| "cqm2sms s q (m#ms) = - (case (cqm2sms s q ms, cm2smsg s q m) of - (Some sms, Some sm) \ Some (sm#sms) - | _ \ None)" - -definition cq2smsgq :: "t_state \ t_msgq \ t_smsgq option" -where - "cq2smsgq s q \ (case (sectxt_of_obj s (O_msgq q), cqm2sms s q (msgs_of_queue s q)) of - (Some sec, Some sms) \ - Some (if (\ deleted (O_msgq q) s \ q \ init_msgqs) then Init q else Created, - sec, sms) - | _ \ None)" - -fun co2sobj :: "t_state \ t_object \ t_sobject option" -where - "co2sobj s (O_proc p) = - (case (cp2sproc s p) of - Some sp \ Some (S_proc sp (O_proc p \ tainted s)) - | _ \ None)" -| "co2sobj s (O_file f) = - Some (S_file (cf2sfiles s f) (O_file f \ tainted s))" -| "co2sobj s (O_dir f) = - (case (cf2sfile s f) of - Some sf \ Some (S_dir sf) - | _ \ None)" -| "co2sobj s (O_msgq q) = - (case (cq2smsgq s q) of - Some sq \ Some (S_msgq sq) - | _ \ None)" -| "co2sobj s (O_shm h) = - (case (ch2sshm s h) of - Some sh \ Some (S_shm sh) - | _ \ None)" -(* -| "co2sobj s (O_msg q m) = - (case (cq2smsgq s q, cm2smsg s q m) of - (Some sq, Some sm) \ Some (S_msg sq sm) - | _ \ None)" -*) -| "co2sobj s _ = None" (***************** for backward proof when Instancing static objects ******************) @@ -863,10 +688,6 @@ where "new_childf pf \ = next_fname pf \ # pf" -definition s2ss :: "t_state \ t_static_state" -where - "s2ss s \ {sobj. \ obj. alive s obj \ co2sobj s obj = Some sobj}" - end end \ No newline at end of file