Theory Post_COMPOSE2
theory Post_COMPOSE2
imports
"Post_ISSUER"
"Post_RECEIVER"
"BD_Security_Compositional.Composing_Security"
begin
subsection ‹Confidentiality for the (binary) issuer-receiver composition›
type_synonym ttrans = "(state, act, out) trans"
type_synonym value1 = Post_ISSUER.value type_synonym value2 = Post_RECEIVER.value
type_synonym obs1 = Post_Observation_Setup_ISSUER.obs
type_synonym obs2 = Post_Observation_Setup_RECEIVER.obs
datatype cval = PValC post
type_synonym cobs = "obs1 × obs2"
locale Post_COMPOSE2 =
Iss: Post_ISSUER UIDs PID +
Rcv: Post_RECEIVER UIDs2 PID AID1
for UIDs :: "userID set" and UIDs2 :: "userID set" and
AID1 :: "apiID" and PID :: "postID"
+ fixes AID2 :: "apiID"
begin
abbreviation "φ1 ≡ Iss.φ" abbreviation "f1 ≡ Iss.f" abbreviation "γ1 ≡ Iss.γ" abbreviation "g1 ≡ Iss.g"
abbreviation "T1 ≡ Iss.T" abbreviation "B1 ≡ Iss.B"
abbreviation "φ2 ≡ Rcv.φ" abbreviation "f2 ≡ Rcv.f" abbreviation "γ2 ≡ Rcv.γ" abbreviation "g2 ≡ Rcv.g"
abbreviation "T2 ≡ Rcv.T" abbreviation "B2 ≡ Rcv.B"
fun isCom1 :: "ttrans ⇒ bool" where
"isCom1 (Trans s (COMact ca1) ou1 s') = (ou1 ≠ outErr)"
|"isCom1 _ = False"
fun isCom2 :: "ttrans ⇒ bool" where
"isCom2 (Trans s (COMact ca2) ou2 s') = (ou2 ≠ outErr)"
|"isCom2 _ = False"
fun isComV1 :: "value1 ⇒ bool" where
"isComV1 (Iss.PValS aid1 txt1) = True"
|"isComV1 _ = False"
fun isComV2 :: "value2 ⇒ bool" where
"isComV2 (Rcv.PValR txt2) = True"
fun syncV :: "value1 ⇒ value2 ⇒ bool" where
"syncV (Iss.PValS aid1 txt1) (Rcv.PValR txt2) = (txt1 = txt2)"
|"syncV _ _ = False"
fun cmpV :: "value1 ⇒ value2 ⇒ cval" where
"cmpV (Iss.PValS aid1 txt1) (Rcv.PValR txt2) = PValC txt1"
|"cmpV _ _ = undefined"
fun isComO1 :: "obs1 ⇒ bool" where
"isComO1 (COMact ca1, ou1) = (ou1 ≠ outErr)"
|"isComO1 _ = False"
fun isComO2 :: "obs2 ⇒ bool" where
"isComO2 (COMact ca2, ou2) = (ou2 ≠ outErr)"
|"isComO2 _ = False"