where " p \<sqsubseteq>\<^sub>A q = (\<forall> x. (case q x of \<lfloor>allow y\<rfloor> \<Rightarrow> (\<exists> z. (p x = \<lfloor>allow z\<rfloor>))
| \<lfloor>deny y\<rfloor> \<Rightarrow> True
@ -97,7 +97,7 @@ lemma more_permissive_trans : "p \<sqsubseteq>\<^sub>A p' \<Longrightarrow> p' \
@@ -97,7 +97,7 @@ lemma more_permissive_trans : "p \<sqsubseteq>\<^sub>A p' \<Longrightarrow> p' \
by(erule_tac x = x in allE, simp)
done
text{* Policy p is more rejective than q: *}
text\<open>Policy p is more rejective than q:\<close>
same operation will occur; this form of side-conditions have to be expressed
inside \HOL. From the user perspective, this will not make much difference,
since junk-data resulting from too weak typing can be ruled out by adopted
front-ends. *}
front-ends.\<close>
text{* Note that the subsequent notion of a test-sequence allows the io stepping
text\<open>Note that the subsequent notion of a test-sequence allows the io stepping
function (and the special case of a program under test) to stop execution
\emph{within} the sequence; such premature terminations are characterized by an
output list which is shorter than the input list. Note that our primary
notion of multiple execution ignores failure and reports failure
steps only by missing results ... *}
steps only by missing results ...\<close>
fun mbind :: "'\<iota> list \<Rightarrow> ('\<iota> \<Rightarrow> ('o,'\<sigma>) MON\<^sub>S\<^sub>E) \<Rightarrow> ('o list,'\<sigma>) MON\<^sub>S\<^sub>E"
@ -188,9 +188,9 @@ fun mbind :: "'\<iota> list \<Rightarrow> ('\<iota> \<Rightarrow> ('o,'\<si
@@ -188,9 +188,9 @@ fun mbind :: "'\<iota> list \<Rightarrow> ('\<iota> \<Rightarrow> ('o,'\<si
text{* As mentioned, this definition is fail-safe; in case of an exception,
text\<open>As mentioned, this definition is fail-safe; in case of an exception,
the current state is maintained, no result is reported.
An alternative is the fail-strict variant @{text "mbind'"} defined below. *}
An alternative is the fail-strict variant \<open>mbind'\<close> defined below.\<close>
lemma mbind_unit [simp]: "mbind [] f = (return [])"
by(rule ext, simp add: unit_SE_def)
@ -214,7 +214,7 @@ lemma mbind_nofailure [simp]: "mbind S f \<sigma> \<noteq> None"
@@ -214,7 +214,7 @@ lemma mbind_nofailure [simp]: "mbind S f \<sigma> \<noteq> None"
done
done
text{* The fail-strict version of @{text mbind'} looks as follows: *}
text\<open>The fail-strict version of \<open>mbind'\<close> looks as follows:\<close>
fun mbind' :: "'\<iota> list \<Rightarrow> ('\<iota> \<Rightarrow> ('o,'\<sigma>) MON\<^sub>S\<^sub>E) \<Rightarrow> ('o list,'\<sigma>) MON\<^sub>S\<^sub>E"
where "mbind' [] iostep \<sigma> = Some([], \<sigma>)" |
@ -48,75 +48,75 @@ theory
@@ -48,75 +48,75 @@ theory
ParallelComposition
begin
text{*
text\<open>
This theory provides the formalisations required for the transformation of UPF policies.
A typical usage scenario can be observed in the firewall case
study~\cite{brucker.ea:formal-fw-testing:2014}.
*}
\<close>
subsection{* Elementary Operators *}
text{*
subsection\<open>Elementary Operators\<close>
text\<open>
We start by providing several operators and theorems useful when reasoning about a list of
rules which should eventually be interpreted as combined using the standard override operator.
*}
\<close>
text{*
text\<open>
The following definition takes as argument a list of rules and returns a policy where the
rules are combined using the standard override operator.
*}
\<close>
definition list2policy::"('a \<mapsto> 'b) list \<Rightarrow> ('a \<mapsto> 'b)" where
"list2policy l = foldr (\<lambda> x y. (x \<Oplus> y)) l \<emptyset>"
text{*
text\<open>
Determine the position of element of a list.
*}
\<close>
fun position :: "'\<alpha> \<Rightarrow> '\<alpha> list \<Rightarrow> nat" where
"position a [] = 0"
|"(position a (x#xs)) = (if a = x then 1 else (Suc (position a xs)))"
text{*
text\<open>
Provides the first applied rule of a policy given as a list of rules.
*}
\<close>
fun applied_rule where
"applied_rule C a (x#xs) = (if a \<in> dom (C x) then (Some x)
else (applied_rule C a xs))"
|"applied_rule C a [] = None"
text {*
text \<open>
The following is used if the list is constructed backwards.
*}
\<close>
definition applied_rule_rev where
"applied_rule_rev C a x = applied_rule C a (rev x)"
text{*
text\<open>
The following is a typical policy transformation. It can be applied to any type of policy and
removes all the rules from a policy with an empty domain. It takes two arguments: a semantic
interpretation function and a list of rules.
*}
\<close>
fun rm_MT_rules where
"rm_MT_rules C (x#xs) = (if dom (C x)= {}
then rm_MT_rules C xs
else x#(rm_MT_rules C xs))"
|"rm_MT_rules C [] = []"
text {*
text \<open>
The following invariant establishes that there are no rules with an empty domain in a list
of rules.
*}
\<close>
fun none_MT_rules where
"none_MT_rules C (x#xs) = (dom (C x) \<noteq> {} \<and> (none_MT_rules C xs))"
|"none_MT_rules C [] = True"
text{*
text\<open>
The following related invariant establishes that the policy has not a completely empty domain.
*}
\<close>
fun not_MT where
"not_MT C (x#xs) = (if (dom (C x) = {}) then (not_MT C xs) else True)"
|"not_MT C [] = False"
text{*
text\<open>
Next, a few theorems about the two invariants and the transformation:
*}
\<close>
lemma none_MT_rules_vs_notMT: "none_MT_rules C p \<Longrightarrow> p \<noteq> [] \<Longrightarrow> not_MT C p"
apply (induct p)
apply (simp_all)
@ -174,7 +174,7 @@ lemma NMPrm: "not_MT C p \<Longrightarrow> not_MT C (rm_MT_rules C p)"
@@ -174,7 +174,7 @@ lemma NMPrm: "not_MT C p \<Longrightarrow> not_MT C (rm_MT_rules C p)"
apply (simp_all)
done
text{* Next, a few theorems about applied\_rule: *}
text\<open>Next, a few theorems about applied\_rule:\<close>
lemma mrconc: "applied_rule_rev C x p = Some a \<Longrightarrow> applied_rule_rev C x (b#p) = Some a"
proof (induct p rule: rev_induct)
case Nil show ?case using Nil
@ -236,8 +236,8 @@ next
@@ -236,8 +236,8 @@ next
qed
subsection{* Distributivity of the Transformation. *}
text{*
subsection\<open>Distributivity of the Transformation.\<close>
text\<open>
The scenario is the following (can be applied iteratively):
\begin{itemize}
\item Two policies are combined using one of the parallel combinators
@ -277,10 +277,10 @@ lemma list2listNMT: "x \<noteq> [] \<Longrightarrow> map sem x \<noteq> []"
@@ -277,10 +277,10 @@ lemma list2listNMT: "x \<noteq> [] \<Longrightarrow> map sem x \<noteq> []"
lemma two_conc: "(prod_list x (y#ys) p r d) = ((r o_f ((p x y) o d))#(prod_list x ys p r d))"
by simp
text{*
text\<open>
The following two invariants establish if the law of distributivity holds for a combinator
and if an operator is strict regarding undefinedness.
*}
\<close>
definition is_distr where
"is_distr p = (\<lambda> g f. (\<forall> N P1 P2. ((g o_f ((p N (P1 \<Oplus> P2)) o f)) =
((g o_f ((p N P1) o f)) \<Oplus> (g o_f ((p N P2) o f))))))"
@ -320,9 +320,9 @@ lemma notDom: "x \<in> dom A \<Longrightarrow> \<not> A x = None"
@@ -320,9 +320,9 @@ lemma notDom: "x \<in> dom A \<Longrightarrow> \<not> A x = None"
apply auto
done
text{*
text\<open>
The following theorems are crucial: they establish the correctness of the distribution.
*}
\<close>
lemma Norm_Distr_1: "((r o_f (((\<Otimes>\<^sub>1) P1 (list2policy P2)) o d)) x =
((list2policy ((P1 \<Otimes>\<^sub>L P2) (\<Otimes>\<^sub>1) r d)) x))"
proof (induct P2)
@ -395,7 +395,7 @@ next
@@ -395,7 +395,7 @@ next
qed
qed
text {* Some domain reasoning *}
text \<open>Some domain reasoning\<close>
lemma domSubsetDistr1: "dom A = UNIV \<Longrightarrow> dom ((\<lambda>(x, y). x) o_f (A \<Otimes>\<^sub>1 B) o (\<lambda> x. (x,x))) = dom B"
@ -244,7 +244,7 @@ where "P \<Otimes>\<^sub>\<nabla> p \<equiv> \<lambda>x. case p (fst x) of
@@ -244,7 +244,7 @@ where "P \<Otimes>\<^sub>\<nabla> p \<equiv> \<lambda>x. case p (fst x) of