Removed references to system generated names.

This commit is contained in:
Achim D. Brucker 2017-01-08 22:12:57 +00:00
parent 23edf34523
commit 0108e7575b
10 changed files with 683 additions and 605 deletions

View File

@ -6,8 +6,9 @@
* 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
* 2008-2015 Achim D. Brucker, Germany
* 2009-2017 Université Paris-Sud, France
* 2015-2017 The University of Sheffield, UK
*
* All rights reserved.
*
@ -39,7 +40,6 @@
* (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: Analysis.thy 10953 2014-11-24 11:23:40Z wolff $ *)
section{* Properties on Policies *}
@ -88,13 +88,14 @@ by(auto split : option.split decision.split)
lemma more_permissive_trans : "p \<sqsubseteq>\<^sub>A p' \<Longrightarrow> p' \<sqsubseteq>\<^sub>A p'' \<Longrightarrow> p \<sqsubseteq>\<^sub>A p''"
unfolding more_permissive_def
apply(auto split : option.split decision.split)
subgoal for x y
apply(erule_tac x = x and
P = "\<lambda>x. case p'' x of \<bottom> \<Rightarrow> True
| \<lfloor>allow y\<rfloor> \<Rightarrow> \<exists>z. p' x = \<lfloor>allow z\<rfloor>
| \<lfloor>deny y\<rfloor> \<Rightarrow> True" in allE)
apply(simp, elim exE)
by(erule_tac x = x in allE, simp)
done
text{* Policy p is more rejective than q: *}
definition more_rejective :: "('a \<mapsto> 'b) \<Rightarrow> ('a \<mapsto> 'b) \<Rightarrow> bool" (infixl "\<sqsubseteq>\<^sub>D" 60)
@ -106,13 +107,14 @@ where " p \<sqsubseteq>\<^sub>D q = (\<forall> x. (case q x of \<lfloor>deny y\<
lemma more_rejective_trans : "p \<sqsubseteq>\<^sub>D p' \<Longrightarrow> p' \<sqsubseteq>\<^sub>D p'' \<Longrightarrow> p \<sqsubseteq>\<^sub>D p''"
unfolding more_rejective_def
apply(auto split : option.split decision.split)
subgoal for x y
apply(erule_tac x = x and
P = "\<lambda>x. case p'' x of \<bottom> \<Rightarrow> True
| \<lfloor>allow y\<rfloor> \<Rightarrow> True
| \<lfloor>deny y\<rfloor> \<Rightarrow> \<exists>z. p' x = \<lfloor>deny z\<rfloor>" in allE)
apply(simp, elim exE)
by(erule_tac x = x in allE, simp)
done
lemma more_rejective_refl : "p \<sqsubseteq>\<^sub>D p "
@ -153,15 +155,18 @@ and B: "p' \<sqsubseteq>\<^bsub>f'\<^esub>\<^sub>,\<^bsub>g'\<^esub> p''"
shows "p \<sqsubseteq>\<^bsub>f o f'\<^esub>\<^sub>,\<^bsub>g o g'\<^esub> p''"
apply(insert A B)
unfolding policy_refinement_def
apply(auto split: option.split decision.split simp: o_def)
apply(erule_tac x="f (f' a')" in allE, simp)
apply(erule_tac x="f' a'" in allE, auto)
apply(erule_tac x=" (f' a')" in allE, auto)
apply(erule_tac x="f (f' a')" in allE, simp)
apply(erule_tac x="f' a'" in allE, auto)
apply(erule_tac x=" (f' a')" in allE, auto)
apply(auto split: option.split decision.split simp: o_def)[1]
subgoal for a a'
apply(erule_tac x="f (f' a')" in allE, simp)[1]
apply(erule_tac x="f' a'" in allE, auto)[1]
apply(erule_tac x=" (f' a')" in allE, auto)[1]
done
subgoal for a a'
apply(erule_tac x="f (f' a')" in allE, simp)[1]
apply(erule_tac x="f' a'" in allE, auto)[1]
apply(erule_tac x=" (f' a')" in allE, auto)[1]
done
done
subsection {* Equivalence of Policies *}
subsubsection{* Equivalence over domain D *}

View File

@ -6,8 +6,9 @@
* 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
* 2008-2015 Achim D. Brucker, Germany
* 2009-2017 Université Paris-Sud, France
* 2015-2017 The University of Sheffield, UK
*
* All rights reserved.
*
@ -39,7 +40,6 @@
* (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
@ -105,11 +105,13 @@ lemma AllD_apply[simp]: "(\<forall>Dx. Some (P x)) x = \<lfloor>deny (P x)\<rflo
lemma neq_Allow_Deny: "pf \<noteq> \<emptyset> \<Longrightarrow> (deny_pfun pf) \<noteq> (allow_pfun pf)"
apply (erule contrapos_nn)
apply (rule ext)
subgoal for x
apply (drule_tac x=x in fun_cong)
apply (auto simp: deny_pfun_def allow_pfun_def)
apply (case_tac "pf x = \<bottom>")
apply (auto)
done
done
subsection{* Common Instances *}
@ -138,12 +140,10 @@ definition
text{* ... and resulting properties: *}
lemma "A\<^sub>I \<Oplus> empty = A\<^sub>I"
apply simp
done
by simp
lemma "A\<^sub>f f \<Oplus> empty = A\<^sub>f f"
apply simp
done
by simp
lemma "allow_pfun empty = empty"
apply (rule ext)
@ -158,8 +158,7 @@ lemma allow_left_cancel :"dom pf = UNIV \<Longrightarrow> (allow_pfun pf) \<Oplu
lemma deny_left_cancel :"dom pf = UNIV \<Longrightarrow> (deny_pfun pf) \<Oplus> x = (deny_pfun pf)"
apply (rule ext)+
apply (auto simp: deny_pfun_def option.splits)
done
by (auto simp: deny_pfun_def option.splits)
subsection{* Domain, Range, and Restrictions *}
@ -193,24 +192,28 @@ text{*
However, some properties are specific to policy concepts:
*}
lemma sub_ran : "ran p \<subseteq> Allow \<union> Deny"
apply (auto simp: Allow_def Deny_def ran_def full_SetCompr_eq[symmetric])
apply (auto simp: Allow_def Deny_def ran_def full_SetCompr_eq[symmetric])[1]
subgoal for x a
apply (case_tac "x")
apply (simp_all)
apply (rename_tac \<alpha>)
apply (erule_tac x="\<alpha>" in allE)
apply (simp)
done
done
lemma dom_allow_pfun [simp]:"dom(allow_pfun f) = dom f"
apply (auto simp: allow_pfun_def)
subgoal for x y
apply (case_tac "f x", simp_all)
done
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 (auto simp: deny_pfun_def)[1]
apply (case_tac "f x")
apply (simp_all)
done
@ -221,12 +224,17 @@ lemma dom_deny_all: " dom(D\<^sub>f f) = UNIV"
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 (auto)[1]
subgoal for x a
apply (case_tac "f a")
apply (auto simp: image_def)
apply (auto simp: image_def)[1]
apply (auto simp: image_def)[1]
done
subgoal for xa a
apply (rule_tac x=a in exI)
apply (simp)
done
done
lemma ran_allow_all: "ran(A\<^sub>f id) = Allow"
apply (simp add: allow_all_fun_def Allow_def o_def)
@ -237,12 +245,17 @@ 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 (auto)[1]
subgoal for x a
apply (case_tac "f a")
apply (auto simp: image_def)
apply (auto simp: image_def)[1]
apply (auto simp: image_def)[1]
done
subgoal for xa a
apply (rule_tac x=a in exI)
apply (simp)
done
done
lemma ran_deny_all: "ran(D\<^sub>f id) = Deny"
apply (simp add: deny_all_fun_def Deny_def o_def)
@ -272,11 +285,15 @@ where "S \<triangleleft> p \<equiv> (\<lambda>x. if x \<in> S then p x else
lemma dom_dom_restrict[simp] : "dom(S \<triangleleft> p) = S \<inter> dom p"
apply (auto simp: dom_restrict_def)
subgoal for x y
apply (case_tac "x \<in> S")
apply (simp_all)
done
subgoal for x y
apply (case_tac "x \<in> S")
apply (simp_all)
done
done
lemma dom_restrict_idem[simp] : "(dom p) \<triangleleft> p = p"
apply (rule ext)
@ -299,11 +316,12 @@ where "p \<triangleright>2 S \<equiv> (\<lambda>x. if (the (p x)) \<in> (S)
lemma "ran_restrict = ran_restrict2"
apply (rule ext)+
apply (simp add: ran_restrict_def ran_restrict2_def)
subgoal for x xa xb
apply (case_tac "x xb")
apply simp_all
apply (metis inj_Some inj_image_mem_iff)
done
done
lemma ran_ran_restrict[simp] : "ran(p \<triangleright> S) = S \<inter> ran p"

View File

@ -6,8 +6,9 @@
* This file is part of HOL-TestGen.
*
* Copyright (c) 2005-2012 ETH Zurich, Switzerland
* 2009-2014 Univ. Paris-Sud, France
* 2009-2014 Achim D. Brucker, Germany
* 2009-2017 Univ. Paris-Sud, France
* 2009-2015 Achim D. Brucker, Germany
* 2015-2017 The University of Sheffield, UK
*
* All rights reserved.
*
@ -39,7 +40,6 @@
* (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: Monads.thy 10922 2014-11-10 15:41:49Z wolff $ *)
section {* Basic Monad Theory for Sequential Computations *}
theory
@ -114,16 +114,22 @@ done
lemma bind_right_unit: "(x \<leftarrow> m; return x) = m"
apply (simp add: unit_SE_def bind_SE_def)
apply (rule ext)
subgoal for "\<sigma>"
apply (case_tac "m \<sigma>")
apply ( simp_all)
done
done
lemma bind_assoc: "(y \<leftarrow> (x \<leftarrow> m; k); h) = (x \<leftarrow> m; (y \<leftarrow> k; h))"
apply (simp add: unit_SE_def bind_SE_def)
apply (rule ext)
subgoal for "\<sigma>"
apply (case_tac "m \<sigma>", simp_all)
subgoal for a
apply (case_tac "a", simp_all)
done
done
done
text{*
In order to express test-sequences also on the object-level and to make our theory amenable to
@ -189,18 +195,24 @@ An alternative is the fail-strict variant @{text "mbind'"} defined below. *}
lemma mbind_unit [simp]: "mbind [] f = (return [])"
by(rule ext, simp add: unit_SE_def)
lemma mbind_nofailure [simp]: "mbind S f \<sigma> \<noteq> None"
apply (rule_tac x=\<sigma> in spec)
apply (induct S)
apply (auto simp:unit_SE_def)
using mbind.simps(1) apply force
apply(simp add:unit_SE_def)
apply(safe)[1]
subgoal for a S x
apply (case_tac "f a x")
apply ( auto)
apply(simp)
apply(safe)[1]
subgoal for aa b
apply (erule_tac x="b" in allE)
apply (erule exE)
apply (erule exE)
apply (erule exE)+
apply (simp)
done
done
done
text{* The fail-strict version of @{text mbind'} looks as follows: *}
fun mbind' :: "'\<iota> list \<Rightarrow> ('\<iota> \<Rightarrow> ('o,'\<sigma>) MON\<^sub>S\<^sub>E) \<Rightarrow> ('o list,'\<sigma>) MON\<^sub>S\<^sub>E"
@ -237,12 +249,17 @@ lemma mbind_try:
else (x \<leftarrow> mbind S F; M (the a' # x)))"
apply (rule ext)
apply (simp add: bind_SE_def try_SE_def)
subgoal for x
apply (case_tac "F a x")
apply (auto)
apply(simp)
apply (safe)[1]
apply (simp add: bind_SE_def try_SE_def)
subgoal for aa b
apply (case_tac "mbind S F b")
apply (auto)
done
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. *}
@ -346,50 +363,55 @@ done
lemma bind_right_unit_SBE: "(x :\<equiv> m; returning x) = m"
apply (rule ext)
apply (simp add: unit_SBE_def bind_SBE_def)
subgoal for x
apply (case_tac "m x")
apply (simp_all add:Let_def)
apply (rule HOL.ccontr)
apply (simp add: Set.image_iff)
done
done
lemmas aux = trans[OF HOL.neq_commute,OF Option.not_None_eq]
lemma bind_assoc_SBE: "(y :\<equiv> (x :\<equiv> m; k); h) = (x :\<equiv> m; (y :\<equiv> k; h))"
proof (rule ext, simp add: unit_SBE_def bind_SBE_def,
case_tac "m x", simp_all add: Let_def Set.image_iff, safe)
case goal1 then show ?case
by(rule_tac x="(a, b)" in bexI, simp_all)
proof (rule ext, simp add: unit_SBE_def bind_SBE_def, rename_tac x,
case_tac "m x", simp_all add: Let_def Set.image_iff, safe,goal_cases)
case (1 x a aa b ab ba a b)
then show ?case by(rule_tac x="(a, b)" in bexI, simp_all)
next
case goal2 then show ?case
case (2 x a aa b ab ba)
then show ?case
apply (rule_tac x="(aa, b)" in bexI, simp_all add:split_def)
apply (erule_tac x="(aa,b)" in ballE)
apply (auto simp: aux image_def split_def intro!: rev_bexI)
done
next
case goal3 then show ?case
by(rule_tac x="(a, b)" in bexI, simp_all)
case (3 x a a b)
then show ?case by(rule_tac x="(a, b)" in bexI, simp_all)
next
case goal4 then show ?case
case (4 x a aa b)
then show ?case
apply (erule_tac Q="None = X" for X in contrapos_pp)
apply (erule_tac x="(aa,b)" and P="\<lambda> x. None \<noteq> case_prod (\<lambda>out. k) x" in ballE)
apply (auto simp: aux (*Option.not_None_eq*) image_def split_def intro!: rev_bexI)
apply (auto simp: aux image_def split_def intro!: rev_bexI)
done
next
case goal5 then show ?case
apply simp apply ((erule_tac x="(ab,ba)" in ballE)+)
apply (simp_all add: aux (* Option.not_None_eq *), (erule exE)+, simp add:split_def)
case (5 x a aa b ab ba a b)
then show ?case apply simp apply ((erule_tac x="(ab,ba)" in ballE)+)
apply (simp_all add: aux, (erule exE)+, simp add:split_def)
apply (erule rev_bexI, case_tac "None\<in>(\<lambda>p. h(snd p))`y",auto simp:split_def)
done
next
case goal6 then show ?case
apply simp apply ((erule_tac x="(a,b)" in ballE)+)
apply (simp_all add: aux (* Option.not_None_eq *), (erule exE)+, simp add:split_def)
case (6 x a aa b a b)
then show ?case apply simp apply ((erule_tac x="(a,b)" in ballE)+)
apply (simp_all add: aux, (erule exE)+, simp add:split_def)
apply (erule rev_bexI, case_tac "None\<in>(\<lambda>p. h(snd p))`y",auto simp:split_def)
done
qed
subsection{* Valid Test Sequences in the State Exception Monad *}
text{*
This is still an unstructured merge of executable monad concepts and specification oriented
@ -465,11 +487,10 @@ lemma valid_propagate_2: "\<sigma> \<Turnstile> ((s \<leftarrow> A ; M s)) \<Lon
apply (simp_all)
apply (drule_tac x="A \<sigma>" and f=the in arg_cong)
apply (simp)
apply (rename_tac a b aa )
apply (rule_tac x="fst aa" in exI)
apply (rule_tac x="snd aa" in exI)
apply (auto)
done
by (auto)
lemma valid_propagate_2': "\<sigma> \<Turnstile> ((s \<leftarrow> A ; M s)) \<Longrightarrow> \<exists> a. (A \<sigma>) = Some a \<and> (snd a) \<Turnstile> (M (fst a))"
apply (auto simp: valid_SE_def unit_SE_def bind_SE_def)
@ -478,19 +499,19 @@ lemma valid_propagate_2': "\<sigma> \<Turnstile> ((s \<leftarrow> A ; M s)) \<Lo
apply (simp_all split: prod.splits)
apply (drule_tac x="A \<sigma>" and f=the in arg_cong)
apply (simp)
apply (rename_tac a b aa x1 x2)
apply (rule_tac x="fst aa" in exI)
apply (rule_tac x="snd aa" in exI)
apply (auto)
done
lemma valid_propagate_2'': "\<sigma> \<Turnstile> ((s \<leftarrow> A ; M s)) \<Longrightarrow> \<exists> v \<sigma>'. A \<sigma> = Some(v,\<sigma>') \<and> \<sigma>' \<Turnstile> (M v)"
apply (auto simp: valid_SE_def unit_SE_def bind_SE_def)
apply (cases "A \<sigma>")
apply (simp_all)
apply (drule_tac x="A \<sigma>" and f=the in arg_cong)
apply (simp)
apply (rename_tac a b aa )
apply (rule_tac x="fst aa" in exI)
apply (rule_tac x="snd aa" in exI)
apply (auto)
@ -517,11 +538,13 @@ lemma assert_D : "(\<sigma> \<Turnstile> (x \<leftarrow> assert\<^sub>S\<^sub>E
lemma assume_D : "(\<sigma> \<Turnstile> (x \<leftarrow> assume\<^sub>S\<^sub>E P; M x)) \<Longrightarrow> \<exists> \<sigma>. (P \<sigma> \<and> \<sigma> \<Turnstile> (M ()))"
apply (auto simp: bind_SE_def assume_SE_def valid_SE_def split: HOL.if_split_asm)
apply (rule_tac x="Eps P" in exI)
apply (auto)
apply (auto)[1]
subgoal for x a b
apply (rule_tac x="True" in exI, rule_tac x="b" in exI)
apply (subst Hilbert_Choice.someI)
apply (assumption)
apply (simp)
done
apply (subst Hilbert_Choice.someI,assumption)
apply (simp)
done
@ -572,5 +595,4 @@ lemma assert_intro: "Q \<sigma> \<Longrightarrow> \<sigma> \<Turnstile>\<^sub>S\
by(simp add: valid_SBE_def assume_SBE_def assert_SBE_def bind_SBE_def)
(* legacy : lemmas valid_failure''=valid_failure *)
end

View File

@ -5,8 +5,9 @@
* 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
* 2008-2015 Achim D. Brucker, Germany
* 2009-2017 Université Paris-Sud, France
* 2015-2017 The University of Sheffield, UK
*
* All rights reserved.
*
@ -38,7 +39,6 @@
* (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: Normalisation.thy 10879 2014-10-26 11:35:31Z brucker $ *)
section{* Policy Transformations *}
theory

View File

@ -5,8 +5,9 @@
* 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
* 2008-2015 Achim D. Brucker, Germany
* 2009-2017 Université Paris-Sud, France
* 2015-2017 The University of Sheffield, UK
*
* All rights reserved.
*
@ -38,7 +39,6 @@
* (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: ParallelComposition.thy 10879 2014-10-26 11:35:31Z brucker $ *)
section{* Parallel Composition*}
theory
@ -96,7 +96,7 @@ done
lemma prod_orA_quasi_commute: "p2 \<Otimes>\<^sub>\<or>\<^sub>A p1 = (((\<lambda>(x,y). (y,x)) o_f (p1 \<Otimes>\<^sub>\<or>\<^sub>A p2))) o (\<lambda>(a,b).(b,a))"
apply (rule ext)
apply (simp add: prod_orA_def policy_range_comp_def o_def)
apply (auto)
apply (auto)[1]
apply (simp split: option.splits decision.splits)
done
@ -116,7 +116,7 @@ where "p1 \<Otimes>\<^sub>\<or>\<^sub>D p2 =
lemma prod_orD_mt[simp]:"p \<Otimes>\<^sub>\<or>\<^sub>D \<emptyset> = \<emptyset>"
apply (rule ext)
apply (simp add: prod_orD_def)
apply (auto)
apply (auto)[1]
apply (simp split: option.splits decision.splits)
done
@ -128,7 +128,7 @@ done
lemma prod_orD_quasi_commute: "p2 \<Otimes>\<^sub>\<or>\<^sub>D p1 = (((\<lambda>(x,y). (y,x)) o_f (p1 \<Otimes>\<^sub>\<or>\<^sub>D p2))) o (\<lambda>(a,b).(b,a))"
apply (rule ext)
apply (simp add: prod_orD_def policy_range_comp_def o_def)
apply (auto)
apply (auto)[1]
apply (simp split: option.splits decision.splits)
done
@ -152,7 +152,7 @@ where "p1 \<Otimes>\<^sub>1 p2 \<equiv>
lemma prod_1_mt[simp]:"p \<Otimes>\<^sub>1 \<emptyset> = \<emptyset>"
apply (rule ext)
apply (simp add: prod_1_def)
apply (auto)
apply (auto)[1]
apply (simp split: option.splits decision.splits)
done
@ -177,7 +177,7 @@ where "p1 \<Otimes>\<^sub>2 p2 \<equiv>
lemma prod_2_mt[simp]:"p \<Otimes>\<^sub>2 \<emptyset> = \<emptyset>"
apply (rule ext)
apply (simp add: prod_2_def)
apply (auto)
apply (auto)[1]
apply (simp split: option.splits decision.splits)
done
@ -264,60 +264,72 @@ lemma distr_or1_a: "(F = F1 \<Oplus> F2) \<Longrightarrow> (((N \<Otimes>\<^su
apply (rule ext)
apply (simp add: prod_1_def map_add_def
split: decision.splits option.splits)
subgoal for x
apply (case_tac "f x")
apply (simp_all add: prod_1_def map_add_def
split: decision.splits option.splits)
done
done
lemma distr_or1: "(F = F1 \<Oplus> F2) \<Longrightarrow> ((g o_f ((N \<Otimes>\<^sub>1 F) o f)) =
((g o_f ((N \<Otimes>\<^sub>1 F1) o f)) \<Oplus> (g o_f ((N \<Otimes>\<^sub>1 F2) o f)))) "
apply (rule ext)+
apply (simp add: prod_1_def map_add_def policy_range_comp_def
split: decision.splits option.splits)
subgoal for x
apply (case_tac "f x")
apply (simp_all add: prod_1_def map_add_def
split: decision.splits option.splits)
done
done
lemma distr_or2_a: "(F = F1 \<Oplus> F2) \<Longrightarrow> (((N \<Otimes>\<^sub>2 F) o f) =
(((N \<Otimes>\<^sub>2 F1) o f) \<Oplus> ((N \<Otimes>\<^sub>2 F2) o f))) "
apply (rule ext)
apply (simp add: prod_2_id_def prod_2_def map_add_def
split: decision.splits option.splits)
subgoal for x
apply (case_tac "f x")
apply (simp_all add: prod_2_def map_add_def
split: decision.splits option.splits)
done
done
lemma distr_or2: "(F = F1 \<Oplus> F2) \<Longrightarrow> ((r o_f ((N \<Otimes>\<^sub>2 F) o f)) =
((r o_f ((N \<Otimes>\<^sub>2 F1) o f)) \<Oplus> (r o_f ((N \<Otimes>\<^sub>2 F2) o f)))) "
apply (rule ext)
apply (simp add: prod_2_id_def prod_2_def map_add_def policy_range_comp_def
split: decision.splits option.splits)
subgoal for x
apply (case_tac "f x")
apply (simp_all add: prod_2_def map_add_def
split: decision.splits option.splits)
done
done
lemma distr_orA: "(F = F1 \<Oplus> F2) \<Longrightarrow> ((g o_f ((N \<Otimes>\<^sub>\<or>\<^sub>A F) o f)) =
((g o_f ((N \<Otimes>\<^sub>\<or>\<^sub>A F1) o f)) \<Oplus> (g o_f ((N \<Otimes>\<^sub>\<or>\<^sub>A F2) o f)))) "
apply (rule ext)+
apply (simp add: prod_orA_def map_add_def policy_range_comp_def
split: decision.splits option.splits)
subgoal for x
apply (case_tac "f x")
apply (simp_all add: map_add_def
split: decision.splits option.splits)
done
done
lemma distr_orD: "(F = F1 \<Oplus> F2) \<Longrightarrow> ((g o_f ((N \<Otimes>\<^sub>\<or>\<^sub>D F) o f)) =
((g o_f ((N \<Otimes>\<^sub>\<or>\<^sub>D F1) o f)) \<Oplus> (g o_f ((N \<Otimes>\<^sub>\<or>\<^sub>D F2) o f)))) "
apply (rule ext)+
apply (simp add: prod_orD_def map_add_def policy_range_comp_def
split: decision.splits option.splits)
subgoal for x
apply (case_tac "f x")
apply (simp_all add: map_add_def
split: decision.splits option.splits)
done
done
lemma coerc_assoc: "(r o_f P) o d = r o_f (P o d)"
apply (simp add: policy_range_comp_def)

View File

@ -5,8 +5,9 @@
* 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
* 2008-2015 Achim D. Brucker, Germany
* 2009-2017 Université Paris-Sud, France
* 2015-2017 The University of Sheffield, UK
*
* All rights reserved.
*
@ -38,7 +39,6 @@
* (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: SeqComposition.thy 10879 2014-10-26 11:35:31Z brucker $ *)
section{* Sequential Composition *}
theory
@ -164,10 +164,11 @@ where "lift f (deny s) = (case f s of
lemma lift_mt [simp]: "lift \<emptyset> = \<emptyset>"
apply (rule ext)
subgoal for x
apply (case_tac "x")
apply (simp_all)
done
done
text{*
Since policies are maps, we inherit a composition on them. However, this results in nestings

View File

@ -5,8 +5,9 @@
* This file is part of HOL-TestGen.
*
* Copyright (c) 2010-2012 ETH Zurich, Switzerland
* 2010-2014 Achim D. Brucker, Germany
* 2010-2014 Université Paris-Sud, France
* 2010-2015 Achim D. Brucker, Germany
* 2010-2017 Université Paris-Sud, France
* 2015-2017 The Univeristy of Sheffield, UK
*
* All rights reserved.
*
@ -38,7 +39,6 @@
* (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: Service.thy 10945 2014-11-21 12:50:43Z wolff $ *)
section {* Secure Service Specification *}
theory

View File

@ -38,7 +38,6 @@
* (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: ServiceExample.thy 10954 2014-11-24 12:43:29Z wolff $ *)
section {* Instantiating Our Secure Service Example *}
theory
@ -114,11 +113,14 @@ lemmas PolSimps = valid_SE_def unit_SE_def bind_SE_def if_splits policy2MON_def
lemma "SE_LR_RBAC_Policy ((createSCR alice Clerical patient1),\<sigma>0)= Some (deny ())"
by (simp add: PolSimps)
lemma exBool[simp]: "\<exists>a::bool. a" by auto
lemma exBool[simp]: "\<exists>a::bool. a"
by auto
lemma deny_allow[simp]: " \<lfloor>deny ()\<rfloor> \<notin> Some ` range allow" by auto
lemma deny_allow[simp]: " \<lfloor>deny ()\<rfloor> \<notin> Some ` range allow"
by auto
lemma allow_deny[simp]: " \<lfloor>allow ()\<rfloor> \<notin> Some ` range deny" by auto
lemma allow_deny[simp]: " \<lfloor>allow ()\<rfloor> \<notin> Some ` range deny"
by auto
text{* Policy as monad. Alice using her first urp can read the SCR of patient1. *}
lemma

View File

@ -40,7 +40,6 @@
* (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: UPF.thy 10879 2014-10-26 11:35:31Z brucker $ *)
section {* Putting Everything Together: UPF *}
theory

View File

@ -6,8 +6,9 @@
* 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
* 2008-2015 Achim D. Brucker, Germany
* 2009-2017 Université Paris-Sud, France
* 2015-2017 The University of Sheffield, UK
*
* All rights reserved.
*
@ -39,7 +40,6 @@
* (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: UPFCore.thy 10951 2014-11-21 21:54:46Z wolff $ *)
section{* The Core of the Unified Policy Framework (UPF) *}
theory
@ -230,11 +230,15 @@ lemma override_A_empty[simp]: "p \<Oplus>\<^sub>A \<emptyset> = p"
lemma empty_override_A[simp]: "\<emptyset> \<Oplus>\<^sub>A p = p"
apply (rule ext)
apply (simp add:override_A_def)
subgoal for x
apply (case_tac "p x")
apply (simp_all)
subgoal for a
apply (case_tac a)
apply (simp_all)
done
done
done
lemma override_A_assoc: "p1 \<Oplus>\<^sub>A (p2 \<Oplus>\<^sub>A p3) = (p1 \<Oplus>\<^sub>A p2) \<Oplus>\<^sub>A p3"
@ -260,9 +264,13 @@ lemma override_D_empty[simp]: "p \<Oplus>\<^sub>D \<emptyset> = p"
lemma empty_override_D[simp]: "\<emptyset> \<Oplus>\<^sub>D p = p"
apply (rule ext)
apply (simp add:override_D_def)
subgoal for x
apply (case_tac "p x", simp_all)
subgoal for a
apply (case_tac a, simp_all)
done
done
done
lemma override_D_assoc: "p1 \<Oplus>\<^sub>D (p2 \<Oplus>\<^sub>D p3) = (p1 \<Oplus>\<^sub>D p2) \<Oplus>\<^sub>D p3"
apply (rule ext)
@ -323,11 +331,15 @@ lemma range_split_charn:
| \<bottom> \<Rightarrow> \<bottom>)"
apply (simp add: range_split_def)
apply (rule ext)
subgoal for x
apply (case_tac "p x")
apply (simp_all)
subgoal for a
apply (case_tac "a")
apply (simp_all)
done
done
done
text{*
The connection between these two becomes apparent if considering the following lemma:
@ -339,21 +351,28 @@ lemma range_split_vs_range_compose: "(f,f) \<nabla> p = f o\<^sub>f p"
lemma range_split_id [simp]: "(id,id) \<nabla> p = p"
apply (rule ext)
apply (simp add: range_split_charn id_def)
subgoal for x
apply (case_tac "p x")
apply (simp_all)
subgoal for a
apply (case_tac "a")
apply (simp_all)
done
done
done
lemma range_split_bi_compose [simp]: "(f1,f2) \<nabla> (g1,g2) \<nabla> p = (f1 o g1,f2 o g2) \<nabla> p"
apply (rule ext)
apply (simp add: range_split_charn comp_def)
subgoal for x
apply (case_tac "p x")
apply (simp_all)
subgoal for a
apply (case_tac "a")
apply (simp_all)
done
done
done
text{*
The next three operators are rather exotic and in most cases not used.