Theory Friend_Request
theory Friend_Request
imports "../Observation_Setup" Friend_Request_Value_Setup
begin
subsection ‹Declassification bound›
fun T :: "(state,act,out) trans ⇒ bool"
where "T (Trans _ _ _ _) = False"
text ‹Friendship updates form an alternating sequence of friending and unfriending,
and every successful friend creation is preceded by one or two friendship requests.›
fun validValSeq :: "value list ⇒ bool ⇒ bool ⇒ bool ⇒ bool" where
"validValSeq [] _ _ _ = True"
| "validValSeq (FRVal U1 req # vl) st r1 r2 ⟷ (¬st) ∧ (¬r1) ∧ validValSeq vl st True r2"
| "validValSeq (FRVal U2 req # vl) st r1 r2 ⟷ (¬st) ∧ (¬r2) ∧ validValSeq vl st r1 True"
| "validValSeq (FVal True # vl) st r1 r2 ⟷ (¬st) ∧ (r1 ∨ r2) ∧ validValSeq vl True False False"
| "validValSeq (FVal False # vl) st r1 r2 ⟷ st ∧ (¬r1) ∧ (¬r2) ∧ validValSeq vl False False False"
| "validValSeq (OVal True # vl) st r1 r2 ⟷ validValSeq vl st r1 r2"
| "validValSeq (OVal False # vl) st r1 r2 ⟷ validValSeq vl st r1 r2"
abbreviation validValSeqFrom :: "value list ⇒ state ⇒ bool"
where "validValSeqFrom vl s
≡ validValSeq vl (friends12 s) (UID1 ∈∈ pendingFReqs s UID2) (UID2 ∈∈ pendingFReqs s UID1)"
text ‹With respect to the friendship status updates, we use the same
``while-or-last-before'' bound as for friendship status confidentiality.›
inductive BO :: "value list ⇒ value list ⇒ bool"
and BC :: "value list ⇒ value list ⇒ bool"
where
BO_FVal[simp,intro!]:
"BO (map FVal fs) (map FVal fs)"
|BO_BC[intro]:
"BC vl vl1 ⟹
BO (map FVal fs @ OVal False # vl) (map FVal fs @ OVal False # vl1)"
|BC_FVal[simp,intro!]:
"BC (map FVal fs) (map FVal fs1)"
|BC_BO[intro]:
"BO vl vl1 ⟹ (fs = [] ⟷ fs1 = []) ⟹ (fs ≠ [] ⟹ last fs = last fs1) ⟹
BC (map FVal fs @ OVal True # vl)
(map FVal fs1 @ OVal True # vl1)"
text ‹Taking into account friendship requests, two value sequences ‹vl› and ‹vl1› are in the bound if
▪ ‹vl1› (with friendship requests) forms a valid value sequence,
▪ ‹vl› and ‹vl1› are in ‹BO› (without friendship requests),
▪ ‹vl1› is empty if ‹vl› is empty, and
▪ ‹vl1› begins with \<^term>‹OVal False› if ‹vl› begins with \<^term>‹OVal False›.
The last two points are due to the fact that \<^term>‹UID1› and \<^term>‹UID1› might not exist yet
if ‹vl› is empty (or before \<^term>‹OVal False›), in which case the observer can deduce that no
friendship request has happened yet.›
definition "B vl vl1 ≡ BO (filter (Not o isFRVal) vl) (filter (Not o isFRVal) vl1) ∧
validValSeqFrom vl1 istate ∧
(vl = [] ⟶ vl1 = []) ∧
(vl ≠ [] ∧ hd vl = OVal False ⟶ vl1 ≠ [] ∧ hd vl1 = OVal False)"
lemma BO_Nil_iff: "BO vl vl1 ⟹ vl = [] ⟷ vl1 = []"
by (cases rule: BO.cases) auto
unbundle no relcomp_syntax