(***************************************************************************** * HOL-TestGen --- theorem-prover based test case generation * http://www.brucker.ch/projects/hol-testgen/ * * UPF * This file is part of HOL-TestGen. * * Copyright (c) 2005-2012 ETH Zurich, Switzerland * 2008-2014 Achim D. Brucker, Germany * 2009-2014 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) (* $Id: ElementaryPolicies.thy 10945 2014-11-21 12:50:43Z wolff $ *) section{* Elementary Policies *} theory ElementaryPolicies imports UPFCore begin 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 *} definition deny_pfun :: "('\ \'\) \ ('\ \ '\)" ("AllD") where "deny_pfun pf \ (\ x. case pf x of \y\ \ \deny (y)\ |\ \ \)" definition allow_pfun :: "('\ \'\) \ ('\ \ '\)" ( "AllA") where "allow_pfun pf \ (\ x. case pf x of \y\ \ \allow (y)\ |\ \ \)" syntax (xsymbols) "_allow_pfun" :: "('\ \'\) \ ('\ \ '\)" ("A\<^sub>p") translations "A\<^sub>p f" \ "AllA f" syntax (xsymbols) "_deny_pfun" :: "('\ \'\) \ ('\ \ '\)" ("D\<^sub>p") translations "D\<^sub>p f" \ "AllD f" notation (xsymbols) "deny_pfun" (binder "\D" 10) and "allow_pfun" (binder "\A" 10) lemma AllD_norm[simp]: "deny_pfun (id o (\x. \x\)) = (\Dx. \x\)" by(simp add:id_def comp_def) lemma AllD_norm2[simp]: "deny_pfun (Some o id) = (\Dx. \x\)" by(simp add:id_def comp_def) lemma AllA_norm[simp]: "allow_pfun (id o Some) = (\Ax. \x\)" by(simp add:id_def comp_def) lemma AllA_norm2[simp]: "allow_pfun (Some o id) = (\Ax. \x\)" by(simp add:id_def comp_def) lemma AllA_apply[simp]: "(\Ax. Some (P x)) x = \allow (P x)\" by(simp add:allow_pfun_def) lemma AllD_apply[simp]: "(\Dx. Some (P x)) x = \deny (P x)\" by(simp add:deny_pfun_def) lemma neq_Allow_Deny: "pf \ \ \ (deny_pfun pf) \ (allow_pfun pf)" apply (erule contrapos_nn) apply (rule ext) apply (drule_tac x=x in fun_cong) apply (auto simp: deny_pfun_def allow_pfun_def) apply (case_tac "pf x = \") apply (auto) done subsection{* Common Instances *} definition allow_all_fun :: "('\ \ '\) \ ('\ \ '\)" ("A\<^sub>f") where "allow_all_fun f = allow_pfun (Some o f)" definition deny_all_fun :: "('\ \ '\) \ ('\ \ '\)" ("D\<^sub>f") where "deny_all_fun f \ deny_pfun (Some o f)" definition deny_all_id :: "'\ \ '\" ("D\<^sub>I") where "deny_all_id \ deny_pfun (id o Some)" definition allow_all_id :: "'\ \ '\" ("A\<^sub>I") where "allow_all_id \ allow_pfun (id o Some)" definition allow_all :: "('\ \ unit)" ("A\<^sub>U") where "allow_all p = \allow ()\" definition deny_all :: "('\ \ unit)" ("D\<^sub>U") where "deny_all p = \deny ()\" text{* ... and resulting properties: *} lemma "A\<^sub>I \ empty = A\<^sub>I" apply simp done lemma "A\<^sub>f f \ empty = A\<^sub>f f" apply simp done lemma "allow_pfun empty = empty" apply (rule ext) apply (simp add: allow_pfun_def) done lemma allow_left_cancel :"dom pf = UNIV \ (allow_pfun pf) \ x = (allow_pfun pf)" apply (rule ext)+ apply (auto simp: allow_pfun_def option.splits) done lemma deny_left_cancel :"dom pf = UNIV \ (deny_pfun pf) \ x = (deny_pfun pf)" apply (rule ext)+ apply (auto simp: deny_pfun_def option.splits) done subsection{* Domain, Range, and Restrictions *} 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} \\ whereas range is just an abrreviation for image: \begin{verbatim} abbreviation range :: "('a => 'b) => 'b set" where -- "of function" "range f == f ` UNIV" \end{verbatim} As a consequence, we inherit the following properties on policies: \begin{itemize} \item \verb+Map.domD+ @{thm Map.domD} \item\verb+Map.domI+ @{thm Map.domI} \item\verb+Map.domIff+ @{thm Map.domIff} \item\verb+Map.dom_const+ @{thm Map.dom_const} \item\verb+Map.dom_def+ @{thm Map.dom_def} \item\verb+Map.dom_empty+ @{thm Map.dom_empty} \item\verb+Map.dom_eq_empty_conv+ @{thm Map.dom_eq_empty_conv} \item\verb+Map.dom_eq_singleton_conv+ @{thm Map.dom_eq_singleton_conv} \item\verb+Map.dom_fun_upd+ @{thm Map.dom_fun_upd} \item\verb+Map.dom_if+ @{thm Map.dom_if} \item\verb+Map.dom_map_add+ @{thm Map.dom_map_add} \end{itemize} *} 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]) apply (case_tac "x") apply (simp_all) apply (rename_tac \) apply (erule_tac x="\" in allE) apply (simp) done lemma dom_allow_pfun [simp]:"dom(allow_pfun f) = dom f" apply (auto simp: allow_pfun_def) apply (case_tac "f x", simp_all) done lemma dom_allow_all: "dom(A\<^sub>f f) = UNIV" by(auto simp: allow_all_fun_def o_def) lemma dom_deny_pfun [simp]:"dom(deny_pfun f) = dom f" apply (auto simp: deny_pfun_def) apply (case_tac "f x") apply (simp_all) done lemma dom_deny_all: " dom(D\<^sub>f f) = UNIV" by(auto simp: deny_all_fun_def o_def) lemma ran_allow_pfun [simp]:"ran(allow_pfun f) = allow `(ran f)" apply (simp add: allow_pfun_def ran_def) apply (rule set_eqI) apply (auto) apply (case_tac "f a") apply (auto simp: image_def) apply (rule_tac x=a in exI) apply (simp) done lemma ran_allow_all: "ran(A\<^sub>f id) = Allow" apply (simp add: allow_all_fun_def Allow_def o_def) apply (rule set_eqI) apply (auto simp: image_def ran_def) done lemma ran_deny_pfun[simp]: "ran(deny_pfun f) = deny ` (ran f)" apply (simp add: deny_pfun_def ran_def) apply (rule set_eqI) apply (auto) apply (case_tac "f a") apply (auto simp: image_def) apply (rule_tac x=a in exI) apply (simp) done lemma ran_deny_all: "ran(D\<^sub>f id) = Deny" apply (simp add: deny_all_fun_def Deny_def o_def) apply (rule set_eqI) apply (auto simp: image_def ran_def) done 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. \begin{itemize} \item \verb+Map.dom_map_add+ @{thm Map.dom_map_add} \item \verb+Map.inj_on_map_add_dom+ @{thm Map.inj_on_map_add_dom} \item \verb+Map.map_add_comm+ @{thm Map.map_add_comm} \item \verb+Map.map_add_dom_app_simps(1)+ @{thm Map.map_add_dom_app_simps(1)} \item \verb+Map.map_add_dom_app_simps(2)+ @{thm Map.map_add_dom_app_simps(2)} \item \verb+Map.map_add_dom_app_simps(3)+ @{thm Map.map_add_dom_app_simps(3)} \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 \)" lemma dom_dom_restrict[simp] : "dom(S \ p) = S \ dom p" apply (auto simp: dom_restrict_def) apply (case_tac "x \ S") apply (simp_all) apply (case_tac "x \ S") apply (simp_all) done lemma dom_restrict_idem[simp] : "(dom p) \ p = p" apply (rule ext) apply (auto simp: dom_restrict_def dest: neq_commute[THEN iffD1,THEN not_None_eq [THEN iffD1]]) done lemma dom_restrict_inter[simp] : "T \ S \ p = T \ S \ p" apply (rule ext) apply (auto simp: dom_restrict_def dest: neq_commute[THEN iffD1,THEN not_None_eq [THEN iffD1]]) done definition ran_restrict :: "['\\'\,'\ decision set] \ '\ \'\" (infixr "\" 55) where "p \ S \ (\x. if p x \ (Some`S) then p x else \)" definition ran_restrict2 :: "['\\'\,'\ decision set] \ '\ \'\" (infixr "\2" 55) where "p \2 S \ (\x. if (the (p x)) \ (S) then p x else \)" lemma "ran_restrict = ran_restrict2" apply (rule ext)+ apply (simp add: ran_restrict_def ran_restrict2_def) apply (case_tac "x xb") apply simp_all apply (metis inj_Some inj_image_mem_iff) done lemma ran_ran_restrict[simp] : "ran(p \ S) = S \ ran p" by(auto simp: ran_restrict_def image_def ran_def) lemma ran_restrict_idem[simp] : "p \ (ran p) = p" apply (rule ext) apply (auto simp: ran_restrict_def image_def Ball_def ran_def) apply (erule contrapos_pp) apply (auto dest!: neq_commute[THEN iffD1,THEN not_None_eq [THEN iffD1]]) done lemma ran_restrict_inter[simp] : "(p \ S) \ T = p \ T \ S" apply (rule ext) apply (auto simp: ran_restrict_def dest: neq_commute[THEN iffD1,THEN not_None_eq [THEN iffD1]]) done lemma ran_gen_A[simp] : "(\Ax. \P x\) \ Allow = (\Ax. \P x\)" apply (rule ext) apply (auto simp: Allow_def ran_restrict_def) done lemma ran_gen_D[simp] : "(\Dx. \P x\) \ Deny = (\Dx. \P x\)" apply (rule ext) apply (auto simp: Deny_def ran_restrict_def) done lemmas ElementaryPoliciesDefs = deny_pfun_def allow_pfun_def allow_all_fun_def deny_all_fun_def allow_all_id_def deny_all_id_def allow_all_def deny_all_def dom_restrict_def ran_restrict_def end