# HG changeset patch # User chunhan # Date 1377741684 -28800 # Node ID 6884b3c9284b0c0255dc5d08bc7147c4aacb056e # Parent 705e1e41faf78ef572ec0e16cbeac67dd8c02d03 fix bug of static.thy for the static of recvmsg case diff -r 705e1e41faf7 -r 6884b3c9284b S2ss_prop.thy --- a/S2ss_prop.thy Wed Aug 28 08:59:12 2013 +0800 +++ b/S2ss_prop.thy Thu Aug 29 10:01:24 2013 +0800 @@ -4,4 +4,9 @@ begin (*>*) -context tainting_s begin \ No newline at end of file +context tainting_s begin + +(* simpset for s2ss*) + +lemma s2ss_execve: + "valid (Execve p f fds # s) \ s2ss (Execve p f fds # s) = s2ss s " \ No newline at end of file diff -r 705e1e41faf7 -r 6884b3c9284b Static.thy --- a/Static.thy Wed Aug 28 08:59:12 2013 +0800 +++ b/Static.thy Thu Aug 29 10:01:24 2013 +0800 @@ -592,7 +592,9 @@ 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_shms ss (pi,pctxt,fds,shms) (tagp \ tagm)) \ static" + \ (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" @@ -786,10 +788,12 @@ (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"