diff --git a/.ci/Jenkinsfile b/.ci/Jenkinsfile index 5878762..7f464b9 100644 --- a/.ci/Jenkinsfile +++ b/.ci/Jenkinsfile @@ -3,7 +3,7 @@ pipeline { stages { stage('Build') { steps { - sh 'docker run -v $PWD/UPF:/UPF logicalhacking:isabelle2018 isabelle build -D /UPF' + sh 'docker run -v $PWD/UPF:/UPF logicalhacking:isabelle2019 isabelle build -D /UPF' } } } diff --git a/UPF/Analysis.thy b/UPF/Analysis.thy index b416307..71f3d01 100644 --- a/UPF/Analysis.thy +++ b/UPF/Analysis.thy @@ -42,7 +42,7 @@ ******************************************************************************) -section{* Properties on Policies *} +section\Properties on Policies\ theory Analysis imports @@ -50,19 +50,19 @@ theory SeqComposition begin -text {* +text \ In this theory, several standard policy properties are paraphrased in UPF terms. -*} +\ -subsection{* Basic Properties *} +subsection\Basic Properties\ -subsubsection{* A Policy Has no Gaps *} +subsubsection\A Policy Has no Gaps\ definition gap_free :: "('a \ 'b) \ bool" where "gap_free p = (dom p = UNIV)" -subsubsection{* Comparing Policies *} -text {* Policy p is more defined than q: *} +subsubsection\Comparing Policies\ +text \Policy p is more defined than q:\ definition more_defined :: "('a \ 'b) \('a \ 'b) \bool" where "more_defined p q = (dom q \ dom p)" @@ -74,7 +74,7 @@ lemma strictly_more_vs_more: "strictly_more_defined p q \ more_d unfolding more_defined_def strictly_more_defined_def by auto -text{* Policy p is more permissive than q: *} +text\Policy p is more permissive than q:\ definition more_permissive :: "('a \ 'b) \ ('a \ 'b) \ bool" (infixl "\\<^sub>A" 60) where " p \\<^sub>A q = (\ x. (case q x of \allow y\ \ (\ z. (p x = \allow z\)) | \deny y\ \ True @@ -97,7 +97,7 @@ lemma more_permissive_trans : "p \\<^sub>A p' \ p' \ by(erule_tac x = x in allE, simp) done -text{* Policy p is more rejective than q: *} +text\Policy p is more rejective than q:\ definition more_rejective :: "('a \ 'b) \ ('a \ 'b) \ bool" (infixl "\\<^sub>D" 60) where " p \\<^sub>D q = (\ x. (case q x of \deny y\ \ (\ z. (p x = \deny z\)) | \allow y\ \ True @@ -130,7 +130,7 @@ lemma "A\<^sub>I \\<^sub>A p" unfolding more_permissive_def allow_all_fun_def allow_pfun_def allow_all_id_def by(auto split: option.split decision.split) -subsection{* Combined Data-Policy Refinement *} +subsection\Combined Data-Policy Refinement\ definition policy_refinement :: "('a \ 'b) \ ('a' \ 'a) \('b' \ 'b) \ ('a' \ 'b') \ bool" @@ -168,13 +168,13 @@ theorem polref_trans: done done -subsection {* Equivalence of Policies *} -subsubsection{* Equivalence over domain D *} +subsection \Equivalence of Policies\ +subsubsection\Equivalence over domain D\ definition p_eq_dom :: "('a \ 'b) \ 'a set \ ('a \ 'b) \bool" ("_ \\<^bsub>_\<^esub> _" [60,60,60]60) where "p \\<^bsub>D\<^esub> q = (\x\D. p x = q x)" -text{* p and q have no conflicts *} +text\p and q have no conflicts\ definition no_conflicts :: "('a \ 'b) \('a \ 'b) \bool" where "no_conflicts p q = (dom p = dom q \ (\x\(dom p). (case p x of \allow y\ \ (\z. q x = \allow z\) @@ -195,7 +195,7 @@ lemma policy_eq: apply (metis)+ done -subsubsection{* Miscellaneous *} +subsubsection\Miscellaneous\ lemma dom_inter: "\dom p \ dom q = {}; p x = \y\\ \ q x = \" by (auto) diff --git a/UPF/ElementaryPolicies.thy b/UPF/ElementaryPolicies.thy index 5299f40..780678d 100644 --- a/UPF/ElementaryPolicies.thy +++ b/UPF/ElementaryPolicies.thy @@ -41,20 +41,20 @@ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) -section{* Elementary Policies *} +section\Elementary Policies\ theory ElementaryPolicies imports UPFCore begin -text{* +text\ In this theory, we introduce the elementary policies of UPF that build the basis for more complex policies. These complex policies, respectively, embedding of well-known access control or security models, are build by composing the elementary policies defined in this theory. -*} +\ -subsection{* The Core Policy Combinators: Allow and Deny Everything *} +subsection\The Core Policy Combinators: Allow and Deny Everything\ definition deny_pfun :: "('\ \'\) \ ('\ \ '\)" ("AllD") @@ -113,7 +113,7 @@ lemma neq_Allow_Deny: "pf \ \ \ (deny_pfun pf) done done -subsection{* Common Instances *} +subsection\Common Instances\ definition allow_all_fun :: "('\ \ '\) \ ('\ \ '\)" ("A\<^sub>f") where "allow_all_fun f = allow_pfun (Some o f)" @@ -137,7 +137,7 @@ definition deny_all :: "('\ \ unit)" ("D\<^sub>U") where "deny_all p = \deny ()\" -text{* ... and resulting properties: *} +text\... and resulting properties:\ lemma "A\<^sub>I \ Map.empty = A\<^sub>I" by simp @@ -160,9 +160,9 @@ lemma deny_left_cancel :"dom pf = UNIV \ (deny_pfun pf) \ apply (rule ext)+ by (auto simp: deny_pfun_def option.splits) -subsection{* Domain, Range, and Restrictions *} +subsection\Domain, Range, and Restrictions\ -text{* +text\ Since policies are essentially maps, we inherit the basic definitions for domain and range on Maps: \\ \verb+Map.dom_def+ : @{thm Map.dom_def} \\ @@ -186,11 +186,11 @@ text{* \item\verb+Map.dom_if+ @{thm Map.dom_if} \item\verb+Map.dom_map_add+ @{thm Map.dom_map_add} \end{itemize} -*} +\ -text{* +text\ However, some properties are specific to policy concepts: -*} +\ lemma sub_ran : "ran p \ Allow \ Deny" apply (auto simp: Allow_def Deny_def ran_def full_SetCompr_eq[symmetric])[1] subgoal for x a @@ -261,7 +261,7 @@ lemma ran_deny_all: "ran(D\<^sub>f id) = Deny" done -text{* +text\ Reasoning over \verb+dom+ is most crucial since it paves the way for simplification and reordering of policies composed by override (i.e. by the normal left-to-right rule composition method. @@ -275,7 +275,7 @@ text{* \item \verb+Map.map_add_upd_left+ @{thm Map.map_add_upd_left} \end{itemize} The latter rule also applies to allow- and deny-override. -*} +\ definition dom_restrict :: "['\ set, '\\'\] \ '\\'\" (infixr "\" 55) where "S \ p \ (\x. if x \ S then p x else \)" diff --git a/UPF/Monads.thy b/UPF/Monads.thy index e29db24..db76679 100644 --- a/UPF/Monads.thy +++ b/UPF/Monads.thy @@ -41,15 +41,15 @@ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) -section {* Basic Monad Theory for Sequential Computations *} +section \Basic Monad Theory for Sequential Computations\ theory Monads imports Main begin -subsection{* General Framework for Monad-based Sequence-Test *} -text{* +subsection\General Framework for Monad-based Sequence-Test\ +text\ As such, Higher-order Logic as a purely functional specification formalism has no built-in mechanism for state and state-transitions. Forms of testing involving state require therefore explicit mechanisms for their treatment inside the logic; a well-known technique to model @@ -67,9 +67,9 @@ text{* \item non-deterministic i/o automata, and \item labelled transition systems (LTS) \end{enumerate} -*} +\ -subsubsection{* State Exception Monads *} +subsubsection\State Exception Monads\ type_synonym ('o, '\) MON\<^sub>S\<^sub>E = "'\ \ ('o \ '\)" definition bind_SE :: "('o,'\)MON\<^sub>S\<^sub>E \ ('o \ ('o','\)MON\<^sub>S\<^sub>E) \ ('o','\)MON\<^sub>S\<^sub>E" @@ -103,9 +103,9 @@ definition if_SE :: "['\ \ bool, ('\, '\)MON\<^ where "if_SE c E F = (\\. if c \ then E \ else F \)" notation if_SE ("if\<^sub>S\<^sub>E") -text{* +text\ The standard monad theorems about unit and associativity: -*} +\ lemma bind_left_unit : "(x \ return a; k) = k" apply (simp add: unit_SE_def bind_SE_def) @@ -131,7 +131,7 @@ lemma bind_assoc: "(y \ (x \ m; k); h) = (x \ m done done -text{* +text\ In order to express test-sequences also on the object-level and to make our theory amenable to formal reasoning over test-sequences, we represent them as lists of input and generalize the bind-operator of the state-exception monad accordingly. The approach is straightforward, but @@ -147,9 +147,9 @@ text{* 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. -*} +\ -text{* +text\ In order to express test-sequences also on the object-level and to make our theory amenable to formal reasoning over test-sequences, we represent them as lists of input and generalize the bind-operator of the state-exception monad accordingly. Thus, the notion of test-sequence @@ -168,15 +168,15 @@ text{* 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.\ -text{* Note that the subsequent notion of a test-sequence allows the io stepping +text\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 ...\ fun mbind :: "'\ list \ ('\ \ ('o,'\) MON\<^sub>S\<^sub>E) \ ('o list,'\) MON\<^sub>S\<^sub>E" @@ -188,9 +188,9 @@ fun mbind :: "'\ list \ ('\ \ ('o,'\ Some([out],\') | Some(outs,\'') \ Some(out#outs,\'')))" -text{* As mentioned, this definition is fail-safe; in case of an exception, +text\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 \mbind'\ defined below.\ 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 \ \ None" done done -text{* The fail-strict version of @{text mbind'} looks as follows: *} +text\The fail-strict version of \mbind'\ looks as follows:\ fun mbind' :: "'\ list \ ('\ \ ('o,'\) MON\<^sub>S\<^sub>E) \ ('o list,'\) MON\<^sub>S\<^sub>E" where "mbind' [] iostep \ = Some([], \)" | "mbind' (a#H) iostep \ = @@ -224,21 +224,21 @@ where "mbind' [] iostep \ = Some([], \)" | None \ None \ \fail-strict\ | Some(outs,\'') \ Some(out#outs,\'')))" -text{* +text\ mbind' as failure strict operator can be seen as a foldr on bind---if the types would match \ldots -*} +\ definition try_SE :: "('o,'\) MON\<^sub>S\<^sub>E \ ('o option,'\) MON\<^sub>S\<^sub>E" where "try_SE ioprog = (\\. case ioprog \ of None \ Some(None, \) | Some(outs, \') \ Some(Some outs, \'))" -text{* In contrast @{term mbind} as a failure safe operator can roughly be seen +text\In contrast @{term mbind} as a failure safe operator can roughly be seen as a @{term foldr} on bind - try: - @{text "m1 ; try m2 ; try m3; ..."}. Note, that the rough equivalence only holds for + \m1 ; try m2 ; try m3; ...\. Note, that the rough equivalence only holds for certain predicates in the sequence - length equivalence modulo None, for example. However, if a conditional is added, the equivalence - can be made precise: *} + can be made precise:\ lemma mbind_try: @@ -261,8 +261,8 @@ lemma mbind_try: done done -text{* On this basis, a symbolic evaluation scheme can be established - that reduces @{term mbind}-code to @{term try_SE}-code and If-cascades. *} +text\On this basis, a symbolic evaluation scheme can be established + that reduces @{term mbind}-code to @{term try_SE}-code and If-cascades.\ definition alt_SE :: "[('o, '\)MON\<^sub>S\<^sub>E, ('o, '\)MON\<^sub>S\<^sub>E] \ ('o, '\)MON\<^sub>S\<^sub>E" (infixl "\\<^sub>S\<^sub>E" 10) @@ -279,13 +279,13 @@ lemma malt_SE_mt [simp]: "\\<^sub>S\<^sub>E [] = fail\<^sub>S\<^sub>E" lemma malt_SE_cons [simp]: "\\<^sub>S\<^sub>E (a # S) = (a \\<^sub>S\<^sub>E (\\<^sub>S\<^sub>E S))" by(simp add: malt_SE_def) -subsubsection{* State-Backtrack Monads *} -text{*This subsection is still rudimentary and as such an interesting +subsubsection\State-Backtrack Monads\ +text\This subsection is still rudimentary and as such an interesting formal analogue to the previous monad definitions. It is doubtful that it is interesting for testing and as a computational structure at all. Clearly more relevant is ``sequence'' instead of ``set,'' which would rephrase Isabelle's internal tactic concept. -*} +\ type_synonym ('o, '\) MON\<^sub>S\<^sub>B = "'\ \ ('o \ '\) set" @@ -318,13 +318,13 @@ lemma bind_assoc_SB: "(y := (x := m; k); h) = (x := m; (y := k; h))" apply (simp add: unit_SB_def bind_SB_def split_def) done -subsubsection{* State Backtrack Exception Monad *} -text{* +subsubsection\State Backtrack Exception Monad\ +text\ The following combination of the previous two Monad-Constructions allows for the semantic foundation of a simple generic assertion language in the style of Schirmer's Simpl-Language or Rustan Leino's Boogie-PL language. The key is to use the exceptional element None for violations of the assert-statement. -*} +\ type_synonym ('o, '\) MON\<^sub>S\<^sub>B\<^sub>E = "'\ \ (('o \ '\) set) option" definition bind_SBE :: "('o,'\)MON\<^sub>S\<^sub>B\<^sub>E \ ('o \ ('o','\)MON\<^sub>S\<^sub>B\<^sub>E) \ ('o','\)MON\<^sub>S\<^sub>B\<^sub>E" @@ -412,20 +412,20 @@ qed -subsection{* Valid Test Sequences in the State Exception Monad *} -text{* +subsection\Valid Test Sequences in the State Exception Monad\ +text\ This is still an unstructured merge of executable monad concepts and specification oriented high-level properties initiating test procedures. -*} +\ definition valid_SE :: "'\ \ (bool,'\) MON\<^sub>S\<^sub>E \ bool" (infix "\" 15) where "(\ \ m) = (m \ \ None \ fst(the (m \)))" -text{* +text\ This notation consideres failures as valid---a definition inspired by I/O conformance. Note that it is not possible to define this concept once and for all in a Hindley-Milner type-system. For the moment, we present it only for the state-exception monad, although for the same definition, this notion is applicable to other monads as well. -*} +\ lemma syntax_test : "\ \ (os \ (mbind \s ioprog); return(length \s = length os))" @@ -435,7 +435,7 @@ oops lemma valid_true[simp]: "(\ \ (s \ return x ; return (P s))) = P x" by(simp add: valid_SE_def unit_SE_def bind_SE_def) -text{* Recall mbind\_unit for the base case. *} +text\Recall mbind\_unit for the base case.\ lemma valid_failure: "ioprog a \ = None \ (\ \ (s \ mbind (a#S) ioprog ; M s)) = @@ -549,12 +549,12 @@ lemma assume_D : "(\ \ (x \ assume\<^sub>S\<^sub>E apply (simp) done -text{* +text\ These two rule prove that the SE Monad in connection with the notion of valid sequence is actually sufficient for a representation of a Boogie-like language. The SBE monad with explicit sets of states---to be shown below---is strictly speaking not necessary (and will therefore be discontinued in the development). -*} +\ lemma if_SE_D1 : "P \ \ (\ \ if\<^sub>S\<^sub>E P B\<^sub>1 B\<^sub>2) = (\ \ B\<^sub>1)" by(auto simp: if_SE_def valid_SE_def) @@ -576,17 +576,17 @@ lemma [code]: "(\ \ m) = (case (m \) of None \Valid Test Sequences in the State Exception Backtrack Monad\ +text\ This is still an unstructured merge of executable monad concepts and specification oriented high-level properties initiating test procedures. -*} +\ definition valid_SBE :: "'\ \ ('a,'\) MON\<^sub>S\<^sub>B\<^sub>E \ bool" (infix "\\<^sub>S\<^sub>B\<^sub>E" 15) where "\ \\<^sub>S\<^sub>B\<^sub>E m \ (m \ \ None)" -text{* +text\ This notation considers all non-failures as valid. -*} +\ lemma assume_assert: "(\ \\<^sub>S\<^sub>B\<^sub>E ( _ :\ assume\<^sub>S\<^sub>B\<^sub>E P ; assert\<^sub>S\<^sub>B\<^sub>E Q)) = (P \ \ Q \)" by(simp add: valid_SBE_def assume_SBE_def assert_SBE_def bind_SBE_def) diff --git a/UPF/Normalisation.thy b/UPF/Normalisation.thy index ecae0e4..beb6317 100644 --- a/UPF/Normalisation.thy +++ b/UPF/Normalisation.thy @@ -40,7 +40,7 @@ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) -section{* Policy Transformations *} +section\Policy Transformations\ theory Normalisation imports @@ -48,75 +48,75 @@ theory ParallelComposition begin -text{* +text\ 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}. -*} +\ -subsection{* Elementary Operators *} -text{* +subsection\Elementary Operators\ +text\ 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. -*} +\ -text{* +text\ The following definition takes as argument a list of rules and returns a policy where the rules are combined using the standard override operator. -*} +\ definition list2policy::"('a \ 'b) list \ ('a \ 'b)" where "list2policy l = foldr (\ x y. (x \ y)) l \" -text{* +text\ Determine the position of element of a list. -*} +\ fun position :: "'\ \ '\ list \ nat" where "position a [] = 0" |"(position a (x#xs)) = (if a = x then 1 else (Suc (position a xs)))" -text{* +text\ Provides the first applied rule of a policy given as a list of rules. -*} +\ fun applied_rule where "applied_rule C a (x#xs) = (if a \ dom (C x) then (Some x) else (applied_rule C a xs))" |"applied_rule C a [] = None" -text {* +text \ The following is used if the list is constructed backwards. -*} +\ definition applied_rule_rev where "applied_rule_rev C a x = applied_rule C a (rev x)" -text{* +text\ 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. -*} +\ 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 \ The following invariant establishes that there are no rules with an empty domain in a list of rules. -*} +\ fun none_MT_rules where "none_MT_rules C (x#xs) = (dom (C x) \ {} \ (none_MT_rules C xs))" |"none_MT_rules C [] = True" -text{* +text\ The following related invariant establishes that the policy has not a completely empty domain. -*} +\ 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\ Next, a few theorems about the two invariants and the transformation: -*} +\ lemma none_MT_rules_vs_notMT: "none_MT_rules C p \ p \ [] \ not_MT C p" apply (induct p) apply (simp_all) @@ -174,7 +174,7 @@ lemma NMPrm: "not_MT C p \ not_MT C (rm_MT_rules C p)" apply (simp_all) done -text{* Next, a few theorems about applied\_rule: *} +text\Next, a few theorems about applied\_rule:\ lemma mrconc: "applied_rule_rev C x p = Some a \ 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 qed -subsection{* Distributivity of the Transformation. *} -text{* +subsection\Distributivity of the Transformation.\ +text\ The scenario is the following (can be applied iteratively): \begin{itemize} \item Two policies are combined using one of the parallel combinators @@ -246,12 +246,12 @@ text{* \item policies that are semantically equivalent to the original policy if \item combined from left to right using the override operator. \end{itemize} -*} +\ -text{* +text\ The following function is crucial for the distribution. Its arguments are a policy, a list of policies, a parallel combinator, and a range and a domain coercion function. -*} +\ fun prod_list :: "('\ \'\) \ (('\ \'\) list) \ (('\ \'\) \ ('\ \'\) \ (('\ \ '\) \ ('\ \ '\))) \ (('\ \ '\) \ 'y) \ ('x \ ('\ \ '\)) \ @@ -260,9 +260,9 @@ fun prod_list :: "('\ \'\) \ (('\ \ An instance, as usual there are four of them. -*} +\ definition prod_2_list :: "[('\ \'\), (('\ \'\) list)] \ (('\ \ '\) \ 'y) \ ('x \ ('\ \ '\)) \ @@ -277,10 +277,10 @@ lemma list2listNMT: "x \ [] \ map sem x \ []" 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\ The following two invariants establish if the law of distributivity holds for a combinator and if an operator is strict regarding undefinedness. -*} +\ definition is_distr where "is_distr p = (\ g f. (\ N P1 P2. ((g o_f ((p N (P1 \ P2)) o f)) = ((g o_f ((p N P1) o f)) \ (g o_f ((p N P2) o f))))))" @@ -320,9 +320,9 @@ lemma notDom: "x \ dom A \ \ A x = None" apply auto done -text{* +text\ The following theorems are crucial: they establish the correctness of the distribution. -*} +\ lemma Norm_Distr_1: "((r o_f (((\\<^sub>1) P1 (list2policy P2)) o d)) x = ((list2policy ((P1 \\<^sub>L P2) (\\<^sub>1) r d)) x))" proof (induct P2) @@ -395,7 +395,7 @@ next qed qed -text {* Some domain reasoning *} +text \Some domain reasoning\ lemma domSubsetDistr1: "dom A = UNIV \ dom ((\(x, y). x) o_f (A \\<^sub>1 B) o (\ x. (x,x))) = dom B" apply (rule set_eqI) apply (rule iffI) diff --git a/UPF/NormalisationTestSpecification.thy b/UPF/NormalisationTestSpecification.thy index 70281c1..c7490b1 100644 --- a/UPF/NormalisationTestSpecification.thy +++ b/UPF/NormalisationTestSpecification.thy @@ -40,14 +40,14 @@ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) -section {* Policy Transformation for Testing *} +section \Policy Transformation for Testing\ theory NormalisationTestSpecification imports Normalisation begin -text{* +text\ This theory provides functions and theorems which are useful if one wants to test policy which are transformed. Most exist in two versions: one where the domains of the rules of the list (which is the result of a transformation) are pairwise disjoint, and one where @@ -55,11 +55,11 @@ text{* The examples in the firewall case study provide a good documentation how these theories can be applied. -*} +\ -text{* +text\ This invariant establishes that the domains of a list of rules are pairwise disjoint. -*} +\ fun disjDom where "disjDom (x#xs) = ((\y\(set xs). dom x \ dom y = {}) \ disjDom xs)" |"disjDom [] = True" @@ -110,11 +110,11 @@ lemma distrPUTL: apply (auto) done -text{* +text\ It makes sense to cater for the common special case where the normalisation returns a list where the last element is a default-catch-all rule. It seems easier to cater for this globally, rather than to require the normalisation procedures to do this. -*} +\ fun gatherDomain_aux where "gatherDomain_aux (x#xs) = (dom x \ (gatherDomain_aux xs))" diff --git a/UPF/ParallelComposition.thy b/UPF/ParallelComposition.thy index 02dd666..e766d74 100644 --- a/UPF/ParallelComposition.thy +++ b/UPF/ParallelComposition.thy @@ -40,14 +40,14 @@ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) -section{* Parallel Composition*} +section\Parallel Composition\ theory ParallelComposition imports ElementaryPolicies begin -text{* +text\ The following combinators are based on the idea that two policies are executed in parallel. Since both input and the output can differ, we chose to pair them. @@ -60,13 +60,13 @@ text{* In any case, although we have strictly speaking a pairing of decisions and not a nesting of them, we will apply the same notational conventions as for the latter, i.e. as for flattening. -*} +\ -subsection{* Parallel Combinators: Foundations *} -text {* +subsection\Parallel Combinators: Foundations\ +text \ There are four possible semantics how the decision can be combined, thus there are four parallel composition operators. For each of them, we prove several properties. -*} +\ definition prod_orA ::"['\\'\, '\ \'\] \ ('\\'\ \ '\\'\)" (infixr "\\<^sub>\\<^sub>A" 55) where "p1 \\<^sub>\\<^sub>A p2 = @@ -132,9 +132,9 @@ lemma prod_orD_quasi_commute: "p2 \\<^sub>\\<^sub>D p1 = (((\ The following two combinators are by definition non-commutative, but still strict. -*} +\ definition prod_1 :: "['\\'\, '\ \'\] \ ('\\'\ \ '\\'\)" (infixr "\\<^sub>1" 55) where "p1 \\<^sub>1 p2 \ @@ -212,11 +212,11 @@ lemma mt_prod_2_id[simp]:"\ \\<^sub>2\<^sub>I p = \" apply (simp add: prod_2_id_def prod_2_def) done -subsection{* Combinators for Transition Policies *} -text {* +subsection\Combinators for Transition Policies\ +text \ For constructing transition policies, two additional combinators are required: one combines state transitions by pairing the states, the other works equivalently on general maps. -*} +\ definition parallel_map :: "('\ \ '\) \ ('\ \ '\) \ ('\ \ '\ \ '\ \ '\)" (infixr "\\<^sub>M" 60) @@ -231,11 +231,11 @@ where "p1 \\<^sub>S p2 = (p1 \\<^sub>M p2) o (\ (a,b,c). ((a,b),a,c))" -subsection{* Range Splitting *} -text{* +subsection\Range Splitting\ +text\ The following combinator is a special case of both a parallel composition operator and a range splitting operator. Its primary use case is when combining a policy with state transitions. -*} +\ definition comp_ran_split :: "[('\ \ '\) \ ('\ \'\), 'd \ '\] \ ('d \ '\) \ ('\ \ '\)" (infixr "\\<^sub>\" 100) @@ -244,7 +244,7 @@ where "P \\<^sub>\ p \ \x. case p (fst x) of | \deny y\ \ (case ((snd P) (snd x)) of \ \ \ | \z\ \ \deny (y,z)\) | \ \ \" -text{* An alternative characterisation of the operator is as follows: *} +text\An alternative characterisation of the operator is as follows:\ lemma comp_ran_split_charn: "(f, g) \\<^sub>\ p = ( (((p \ Allow)\\<^sub>\\<^sub>A (A\<^sub>p f)) \ @@ -257,7 +257,7 @@ lemma comp_ran_split_charn: apply (auto) done -subsection {* Distributivity of the parallel combinators *} +subsection \Distributivity of the parallel combinators\ lemma distr_or1_a: "(F = F1 \ F2) \ (((N \\<^sub>1 F) o f) = (((N \\<^sub>1 F1) o f) \ ((N \\<^sub>1 F2) o f))) " diff --git a/UPF/ROOT b/UPF/ROOT index c1c3e30..e920cca 100644 --- a/UPF/ROOT +++ b/UPF/ROOT @@ -1,7 +1,7 @@ chapter AFP session "UPF-devel" (AFP) = HOL + - description {* The Unified Policy Framework (UPF) *} + description "The Unified Policy Framework (UPF) " options [timeout = 300] theories Monads diff --git a/UPF/SeqComposition.thy b/UPF/SeqComposition.thy index 55f5a5b..c9c1b7f 100644 --- a/UPF/SeqComposition.thy +++ b/UPF/SeqComposition.thy @@ -40,23 +40,23 @@ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) -section{* Sequential Composition *} +section\Sequential Composition\ theory SeqComposition imports ElementaryPolicies begin -text{* +text\ Sequential composition is based on the idea that two policies are to be combined by applying the second policy to the output of the first one. Again, there are four possibilities how the - decisions can be combined. *} + decisions can be combined.\ -subsection {* Flattening *} -text{* +subsection \Flattening\ +text\ A key concept of sequential policy composition is the flattening of nested decisions. There are four possibilities, and these possibilities will give the various flavours of policy composition. -*} +\ fun flat_orA :: "('\ decision) decision \ ('\ decision)" where "flat_orA(allow(allow y)) = allow y" |"flat_orA(allow(deny y)) = allow y" @@ -149,10 +149,10 @@ lemma flat_2_deny[dest]: "flat_2 x = deny y \ x = deny(deny y) apply (case_tac "\", simp_all)[1] done -subsection{* Policy Composition *} -text{* +subsection\Policy Composition\ +text\ The following definition allows to compose two policies. Denies and allows are transferred. -*} +\ fun lift :: "('\ \ '\) \ ('\ decision \'\ decision)" where "lift f (deny s) = (case f s of @@ -170,10 +170,10 @@ lemma lift_mt [simp]: "lift \ = \" done done -text{* +text\ Since policies are maps, we inherit a composition on them. However, this results in nestings of decisions---which must be flattened. As we now that there are four different forms of - flattening, we have four different forms of policy composition: *} + flattening, we have four different forms of policy composition:\ definition comp_orA :: "['\\'\, '\\'\] \ '\\'\" (infixl "o'_orA" 55) where "p2 o_orA p1 \ (map_option flat_orA) o (lift p2 \\<^sub>m p1)" diff --git a/UPF/Service.thy b/UPF/Service.thy index 5160693..1695884 100644 --- a/UPF/Service.thy +++ b/UPF/Service.thy @@ -40,52 +40,52 @@ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) -section {* Secure Service Specification *} +section \Secure Service Specification\ theory Service imports UPF begin -text {* +text \ In this section, we model a simple Web service and its access control model that allows the staff in a hospital to access health care records of patients. -*} +\ -subsection{* Datatypes for Modelling Users and Roles*} -subsubsection {* Users *} -text{* +subsection\Datatypes for Modelling Users and Roles\ +subsubsection \Users\ +text\ First, we introduce a type for users that we use to model that each staff member has a unique id: -*} +\ type_synonym user = int (* Each NHS employee has a unique NHS_ID. *) -text {* +text \ Similarly, each patient has a unique id: -*} +\ type_synonym patient = int (* Each patient gets a unique id *) -subsubsection {* Roles and Relationships*} -text{* In our example, we assume three different roles for members of the clinical staff: *} +subsubsection \Roles and Relationships\ +text\In our example, we assume three different roles for members of the clinical staff:\ datatype role = ClinicalPractitioner | Nurse | Clerical -text{* +text\ We model treatment relationships (legitimate relationships) between staff and patients (respectively, their health records. This access control model is inspired by our detailed NHS model. -*} +\ type_synonym lr_id = int type_synonym LR = "lr_id \ (user set)" -text{* The security context stores all the existing LRs. *} +text\The security context stores all the existing LRs.\ type_synonym \ = "patient \ LR" -text{* The user context stores the roles the users are in. *} +text\The user context stores the roles the users are in.\ type_synonym \ = "user \ role" -subsection {* Modelling Health Records and the Web Service API*} -subsubsection {* Health Records *} -text {* The content and the status of the entries of a health record *} +subsection \Modelling Health Records and the Web Service API\ +subsubsection \Health Records\ +text \The content and the status of the entries of a health record\ datatype data = dummyContent datatype status = Open | Closed type_synonym entry_id = int @@ -93,8 +93,8 @@ type_synonym entry = "status \ user \ data" type_synonym SCR = "(entry_id \ entry)" type_synonym DB = "patient \ SCR" -subsubsection {* The Web Service API *} -text{* The operations provided by the service: *} +subsubsection \The Web Service API\ +text\The operations provided by the service:\ datatype Operation = createSCR user role patient | appendEntry user role patient entry_id entry | deleteEntry user role patient entry_id @@ -207,17 +207,17 @@ fun allContentStatic where |"allContentStatic [] = True" -subsection{* Modelling Access Control*} -text {* +subsection\Modelling Access Control\ +text \ In the following, we define a rather complex access control model for our scenario that extends traditional role-based access control (RBAC)~\cite{sandhu.ea:role-based:1996} with treatment relationships and sealed envelopes. Sealed envelopes (see~\cite{bruegger:generation:2012} for details) are a variant of break-the-glass access control (see~\cite{brucker.ea:extending:2009} for a general motivation and explanation of break-the-glass access control). -*} +\ -subsubsection {* Sealed Envelopes *} +subsubsection \Sealed Envelopes\ type_synonym SEPolicy = "(Operation \ DB \ unit) " @@ -259,7 +259,7 @@ definition SEPolicy :: SEPolicy where lemmas SEsimps = SEPolicy_def get_entry_def userHasAccess_def editEntrySE_def deleteEntrySE_def readEntrySE_def -subsubsection {* Legitimate Relationships *} +subsubsection \Legitimate Relationships\ type_synonym LRPolicy = "(Operation \ \, unit) policy" @@ -365,7 +365,7 @@ definition FunPolicy where removeLRFunPolicy \ readSCRFunPolicy \ addLRFunPolicy \ createFunPolicy \ A\<^sub>U" -subsubsection{* Modelling Core RBAC *} +subsubsection\Modelling Core RBAC\ type_synonym RBACPolicy = "Operation \ \ \ unit" @@ -389,9 +389,9 @@ definition RBACPolicy :: RBACPolicy where then \allow ()\ else \deny ()\)" -subsection {* The State Transitions and Output Function*} +subsection \The State Transitions and Output Function\ -subsubsection{* State Transition *} +subsubsection\State Transition\ fun OpSuccessDB :: "(Operation \ DB) \ DB" where "OpSuccessDB (createSCR u r p,S) = (case S p of \ \ \S(p\\)\ @@ -434,7 +434,7 @@ fun OpSuccessSigma :: "(Operation \ \) \ \" fun OpSuccessUC :: "(Operation \ \) \ \" where "OpSuccessUC (f,u) = \u\" -subsubsection {* Output *} +subsubsection \Output\ type_synonym Output = unit @@ -445,7 +445,7 @@ fun OpSuccessOutput :: "(Operation) \ Output" where fun OpFailOutput :: "Operation \ Output" where "OpFailOutput x = \()\" -subsection {* Combine All Parts *} +subsection \Combine All Parts\ definition SE_LR_Policy :: "(Operation \ DB \ \, unit) policy" where "SE_LR_Policy = (\(x,x). x) o\<^sub>f (SEPolicy \\<^sub>\\<^sub>D LR_Policy) o (\(a,b,c). ((a,b),a,c))" diff --git a/UPF/ServiceExample.thy b/UPF/ServiceExample.thy index f584640..4ab4bf1 100644 --- a/UPF/ServiceExample.thy +++ b/UPF/ServiceExample.thy @@ -39,19 +39,19 @@ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) -section {* Instantiating Our Secure Service Example *} +section \Instantiating Our Secure Service Example\ theory ServiceExample imports Service begin -text {* +text \ In the following, we briefly present an instantiations of our secure service example from the last section. We assume three different members of the health care staff and two patients: -*} +\ -subsection {* Access Control Configuration *} +subsection \Access Control Configuration\ definition alice :: user where "alice = 1" definition bob :: user where "bob = 2" definition charlie :: user where "charlie = 3" @@ -85,11 +85,11 @@ definition LR1 :: LR where definition \0 :: \ where "\0 = (Map.empty(patient1\LR1))" -subsection {* The Initial System State *} +subsection \The Initial System State\ definition \0 :: "DB \ \\\" where "\0 = (Spine0,\0,UC0)" -subsection{* Basic Properties *} +subsection\Basic Properties\ lemma [simp]: "(case a of allow d \ \X\ | deny d2 \ \Y\) = \ \ False" by (case_tac a,simp_all) @@ -122,13 +122,13 @@ lemma deny_allow[simp]: " \deny ()\ \ Some ` range allow" lemma allow_deny[simp]: " \allow ()\ \ Some ` range deny" by auto -text{* Policy as monad. Alice using her first urp can read the SCR of patient1. *} +text\Policy as monad. Alice using her first urp can read the SCR of patient1.\ lemma "(\0 \ (os \ mbind [(createSCR alice Clerical patient1)] (PolMon); (return (os = [(deny (Out) )]))))" by (simp add: PolMon_def MonSimps PolSimps) -text{* Presenting her other urp, she is not allowed to read it. *} +text\Presenting her other urp, she is not allowed to read it.\ lemma "SE_LR_RBAC_Policy ((appendEntry alice Clerical patient1 ei d),\0)= \deny ()\" by (simp add: PolSimps) diff --git a/UPF/UPF.thy b/UPF/UPF.thy index c4b68a3..424bb6d 100644 --- a/UPF/UPF.thy +++ b/UPF/UPF.thy @@ -41,7 +41,7 @@ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) -section {* Putting Everything Together: UPF *} +section \Putting Everything Together: UPF\ theory UPF imports @@ -50,10 +50,10 @@ theory Analysis begin -text{* +text\ This is the top-level theory for the Unified Policy Framework (UPF) and, thus, builds the base theory for using UPF. For the moment, we only define a set of lemmas for all core UPF definitions that is useful for using UPF: -*} +\ lemmas UPFDefs = UPFCoreDefs ParallelDefs ElementaryPoliciesDefs end diff --git a/UPF/UPFCore.thy b/UPF/UPFCore.thy index fba5c4e..dda007e 100644 --- a/UPF/UPFCore.thy +++ b/UPF/UPFCore.thy @@ -41,7 +41,7 @@ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) -section{* The Core of the Unified Policy Framework (UPF) *} +section\The Core of the Unified Policy Framework (UPF)\ theory UPFCore imports @@ -49,8 +49,8 @@ theory begin -subsection{* Foundation *} -text{* +subsection\Foundation\ +text\ The purpose of this theory is to formalize a somewhat non-standard view on the fundamental concept of a security policy which is worth outlining. This view has arisen from prior experience in the modelling of network @@ -74,37 +74,37 @@ text{* In more detail, we model policies as partial functions based on input data $\alpha$ (arguments, system state, security context, ...) to output data $\beta$: -*} +\ datatype '\ decision = allow '\ | deny '\ type_synonym ('\,'\) policy = "'\ \ '\ decision" (infixr "|->" 0) -text{*In the following, we introduce a number of shortcuts and alternative notations. -The type of policies is represented as: *} +text\In the following, we introduce a number of shortcuts and alternative notations. +The type of policies is represented as:\ translations (type) "'\ |-> '\" <= (type) "'\ \ '\ decision" type_notation "policy" (infixr "\" 0) -text{* ... allowing the notation @{typ "'\ \ '\"} for the policy type and the +text\... allowing the notation @{typ "'\ \ '\"} for the policy type and the alternative notations for @{term None} and @{term Some} of the \HOL library -@{typ "'\ option"} type:*} +@{typ "'\ option"} type:\ notation "None" ("\") notation "Some" ("\_\" 80) -text{* Thus, the range of a policy may consist of @{term "\accept x\"} data, +text\Thus, the range of a policy may consist of @{term "\accept x\"} data, of @{term "\deny x\"} data, as well as @{term "\"} modeling the undefinedness of a policy, i.e. a policy is considered as a partial function. Partial functions are used since we describe elementary policies by partial system behaviour, which are glued together by operators such as function override and functional composition. -*} +\ -text{* We define the two fundamental sets, the allow-set $Allow$ and the +text\We define the two fundamental sets, the allow-set $Allow$ and the deny-set $Deny$ (written $A$ and $D$ set for short), to characterize these two main sets of the range of a policy. -*} +\ definition Allow :: "('\ decision) set" where "Allow = range allow" @@ -112,13 +112,13 @@ definition Deny :: "('\ decision) set" where "Deny = range deny" -subsection{* Policy Constructors *} -text{* +subsection\Policy Constructors\ +text\ Most elementary policy constructors are based on the update operation @{thm [source] "Fun.fun_upd_def"} @{thm Fun.fun_upd_def} and the maplet-notation @{term "a(x \ y)"} used for @{term "a(x:=\y\)"}. - Furthermore, we add notation adopted to our problem domain: *} + Furthermore, we add notation adopted to our problem domain:\ nonterminal policylets and policylet @@ -137,14 +137,14 @@ translations "_MapUpd m (_policylet2 x y)" \ "m(x := CONST Some (CONST deny y))" "\" \ "CONST Map.empty" -text{* Here are some lemmas essentially showing syntactic equivalences: *} +text\Here are some lemmas essentially showing syntactic equivalences:\ lemma test: "\(x\\<^sub>+a, y\\<^sub>-b) = \(x \\<^sub>+ a, y \\<^sub>- b)" by simp lemma test2: "p(x\\<^sub>+a,x\\<^sub>-b) = p(x\\<^sub>-b)" by simp -text{* +text\ We inherit a fairly rich theory on policy updates from Map here. Some examples are: -*} +\ lemma pol_upd_triv1: "t k = \allow x\ \ t(k\\<^sub>+x) = t" by (rule ext) simp @@ -168,14 +168,14 @@ lemma pol_upd_neq1 [simp]: "m(a\\<^sub>+x) \ n(a\\<^sub>- by(auto dest: map_upd_eqD1) -subsection{* Override Operators *} -text{* +subsection\Override Operators\ +text\ Key operators for constructing policies are the override operators. There are four different versions of them, with one of them being the override operator from the Map theory. As it is common to compose policy rules in a ``left-to-right-first-fit''-manner, that one is taken as default, defined by a syntax translation from the provided override operator from the Map theory (which does it in reverse order). -*} +\ syntax "_policyoverride" :: "['a \ 'b, 'a \ 'b] \ 'a \ 'b" (infixl "\" 100) @@ -183,9 +183,9 @@ translations "p \ q" \ "q ++ p" -text{* +text\ Some elementary facts inherited from Map are: -*} +\ lemma override_empty: "p \ \ = p" by simp @@ -196,10 +196,10 @@ lemma empty_override: "\ \ p = p" lemma override_assoc: "p1 \ (p2 \ p3) = (p1 \ p2) \ p3" by simp -text{* +text\ The following two operators are variants of the standard override. For override\_A, an allow of wins over a deny. For override\_D, the situation is dual. -*} +\ definition override_A :: "['\\'\, '\\'\] \ '\\'\" (infixl "++'_A" 100) where "m2 ++_A m1 = @@ -268,15 +268,15 @@ lemma override_D_assoc: "p1 \\<^sub>D (p2 \\<^sub>D p3) = (p1 \Coercion Operators\ +text\ Often, especially when combining policies of different type, it is necessary to adapt the input or output domain of a policy to a more refined context. -*} +\ -text{* +text\ An analogous for the range of a policy is defined as follows: -*} +\ definition policy_range_comp :: "['\\'\, '\\'\] \ '\ \'\" (infixl "o'_f" 55) where @@ -296,10 +296,10 @@ lemma policy_range_comp_strict : "f o\<^sub>f \ = \" done -text{* +text\ A generalized version is, where separate coercion functions are applied to the result depending on the decision of the policy is as follows: -*} +\ definition range_split :: "[('\\'\)\('\\'\),'\ \ '\] \ '\ \ '\" (infixr "\" 100) @@ -331,9 +331,9 @@ lemma range_split_charn: done done -text{* +text\ The connection between these two becomes apparent if considering the following lemma: -*} +\ lemma range_split_vs_range_compose: "(f,f) \ p = f o\<^sub>f p" by(simp add: range_split_charn policy_range_comp_def) @@ -364,14 +364,14 @@ lemma range_split_bi_compose [simp]: "(f1,f2) \ (g1,g2) \ p = (f1 done done -text{* +text\ The next three operators are rather exotic and in most cases not used. -*} +\ -text {* +text \ The following is a variant of range\_split, where the change in the decision depends on the input instead of the output. -*} +\ definition dom_split2a :: "[('\ \ '\) \ ('\ \'\),'\ \ '\] \ '\ \ '\" (infixr "\a" 100) where "P \a p = (\x. case p x of @@ -391,11 +391,11 @@ where "P \2 p = (\x. case p x of | \deny y\ \ \deny (y,(snd P) x)\ | \ \ \)" -text{* +text\ The following operator is used for transition policies only: a transition policy is transformed into a state-exception monad. Such a monad can for example be used for test case generation using HOL-Testgen~\cite{brucker.ea:theorem-prover:2012}. -*} +\ definition policy2MON :: "('\\'\ \ 'o\'\) \ ('\ \('o decision,'\) MON\<^sub>S\<^sub>E)" where "policy2MON p = (\ \ \. case p (\,\) of diff --git a/UPF/document/root.bib b/UPF/document/root.bib index 520d028..2503e67 100644 --- a/UPF/document/root.bib +++ b/UPF/document/root.bib @@ -8,7 +8,7 @@ # {\providecommand{\isbn}{\textsc{isbn}} } # {\providecommand{\Cpp}{C++} } # {\providecommand{\Specsharp}{Spec\#} } - # {\providecommand{\doi}[1]{\href{http://dx.doi.org/#1}{doi: + # {\providecommand{\doi}[1]{\href{https://doi.org/#1}{doi: {\urlstyle{rm}\nolinkurl{#1}}}}} } @STRING{conf-sacmat="ACM symposium on access control models and technologies (SACMAT)" } @@ -319,7 +319,7 @@ revocation are provided, and proofs are given for the important properties of our delegation framework.}, issn = {0306-4379}, - doi = {http://dx.doi.org/10.1016/j.is.2005.11.008}, + doi = {https://doi.org/10.1016/j.is.2005.11.008}, publisher = pub-elsevier, address = {Oxford, UK, UK}, tags = {ReadingList, SoKNOS},