diff --git a/thys/CSP_RefTK/Assertions_ext.thy b/thys/CSP_RefTK/Assertions_ext.thy new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/Assertions_ext.thy @@ -0,0 +1,818 @@ +(*<*) +\\ ******************************************************************** + * Project : CSP-RefTK - A Refinement Toolkit for HOL-CSP + * Version : 1.0 + * + * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * + * This file : Assertions on DF and LF and their Relations + * + * Copyright (c) 2020 Université Paris-Saclay, 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. + ******************************************************************************\ +(*>*) + +chapter\Basic CSP\_RefTk-Theories\ + +theory Assertions_ext + imports "HOL-CSP.Assertions" +begin + +section \Preliminaries\ + +lemma DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold : "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A = ((\ z \ A \ DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) \ SKIP)" + by(simp add: DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule trans, rule fix_eq, simp) + +section \All refinements definitions\ + +thm failure_divergence_refine_def[simplified le_ref_def] trace_refine_def failure_refine_def + +definition divergence_refine :: "'a process \ 'a process \ bool" (infix "\\<^sub>D" 60) + where "P \\<^sub>D Q \ D Q \ D P" + +definition trace_divergence_refine :: "'a process \ 'a process \ bool" (infix "\\<^sub>D\<^sub>T" 60) + where "P \\<^sub>D\<^sub>T Q \ P \\<^sub>T Q \ P \\<^sub>D Q" + +section \Relations between refinements\ + +lemma le_F_T: "P \\<^sub>F Q \ P \\<^sub>T Q" + by (simp add: F_subset_imp_T_subset failure_refine_def trace_refine_def) + +lemma FD_F: "P \\<^sub>F\<^sub>D Q \ P \\<^sub>F Q" + by (simp add: failure_divergence_refine_def failure_refine_def le_ref_def) + +lemma FD_D: "P \\<^sub>F\<^sub>D Q \ P \\<^sub>D Q" + by (simp add: divergence_refine_def failure_divergence_refine_def le_ref_def) + +lemma DT_D: "P \\<^sub>D\<^sub>T Q \ P \\<^sub>D Q" + by (simp add: trace_divergence_refine_def) + +lemma DT_T: "P \\<^sub>D\<^sub>T Q \ P \\<^sub>T Q" + by (simp add: trace_divergence_refine_def) + +lemma F_D_FD:"P \\<^sub>F Q \ P \\<^sub>D Q \ P \\<^sub>F\<^sub>D Q" + by (simp add: divergence_refine_def failure_divergence_refine_def failure_refine_def le_ref_def) + +lemma D_T_DT:"P \\<^sub>D Q \ P \\<^sub>T Q \ P \\<^sub>D\<^sub>T Q" + by (simp add: trace_divergence_refine_def) + +section \Some obvious refinements\ + +lemma bot_FD[simp]: "\ \\<^sub>F\<^sub>D Q" + by (simp add: failure_divergence_refine_def) + +corollary bot_F[simp]: "\ \\<^sub>F Q" + and bot_D[simp]: "\ \\<^sub>D Q" + and bot_T[simp]: "\ \\<^sub>T Q" + and bot_DT[simp]: "\ \\<^sub>D\<^sub>T Q" + by (simp_all add: FD_F FD_D le_F_T D_T_DT) + +lemma STOP_leDT[simp]: "P \\<^sub>D\<^sub>T STOP" + by (simp add: D_STOP D_T_DT Nil_elem_T T_STOP divergence_refine_def trace_refine_def) + +\\For refinement proofs, we need admissibility and monotony starting with idempotency and + transitivity\ + +section \Idempotency\ + +lemma idem_F[simp]: "P \\<^sub>F P" + by (simp add: failure_refine_def) + +lemma idem_D[simp]: "P \\<^sub>D P" + by (simp add: divergence_refine_def) + +lemma idem_T[simp]: "P \\<^sub>T P" + by (simp add: trace_refine_def) + +lemma idem_FD[simp]: "P \\<^sub>F\<^sub>D P" + by (simp add: failure_divergence_refine_def) + +lemma idem_DT[simp]: "P \\<^sub>D\<^sub>T P" + by (simp add: trace_divergence_refine_def) + +section \Transitivity\ + +lemma trans_F: "P \\<^sub>F Q \ Q \\<^sub>F S \ P \\<^sub>F S" + by (meson failure_refine_def order_trans) + +lemma trans_D: "P \\<^sub>D Q \ Q \\<^sub>D S \ P \\<^sub>D S" + by (meson divergence_refine_def order_trans) + +lemma trans_T: "P \\<^sub>T Q \ Q \\<^sub>T S \ P \\<^sub>T S" + by (meson trace_refine_def order_trans) + +lemma trans_FD: "P \\<^sub>F\<^sub>D Q \ Q \\<^sub>F\<^sub>D S \ P \\<^sub>F\<^sub>D S" + by (meson failure_divergence_refine_def order_trans) + +lemma trans_DT: "P \\<^sub>D\<^sub>T Q \ Q \\<^sub>D\<^sub>T S \ P \\<^sub>D\<^sub>T S" + by (meson trace_divergence_refine_def order_trans trans_D trans_T) + +section \Admissibility\ + +lemma le_F_adm: "cont (u::('a::cpo) \ 'b process) \ monofun v \ adm(\x. u x \\<^sub>F v x)" +proof(auto simp add:cont2contlubE adm_def failure_refine_def) + fix Y a b + assume 1:"cont u" + and 2:"monofun v" + and 3:"chain Y" + and 4:"\i. F (v (Y i)) \ F (u (Y i))" + and 5:" (a, b) \ F (v (\x. Y x))" + hence "v (Y i) \ v (\i. Y i)" for i by (simp add: is_ub_thelub monofunE) + hence "F (v (\i. Y i)) \ F (u (Y i))" for i using 4 le_approx_lemma_F by blast + then show "(a, b) \ F (\i. u (Y i))" + using F_LUB[OF ch2ch_cont[OF 1 3]] limproc_is_thelub[OF ch2ch_cont[OF 1 3]] 5 by force +qed + +lemma le_T_adm: "cont (u::('a::cpo) \ 'b process) \ monofun v \ adm(\x. u x \\<^sub>T v x)" +proof(auto simp add:cont2contlubE adm_def trace_refine_def) + fix Y x + assume 1:"cont u" + and 2:"monofun v" + and 3:"chain Y" + and 4:"\i. T (v (Y i)) \ T (u (Y i))" + and 5:" x \ T (v (\i. Y i))" + hence "v (Y i) \ v (\i. Y i)" for i by (simp add: is_ub_thelub monofunE) + hence "T (v (\i. Y i)) \ T (u (Y i))" for i using 4 le_approx_lemma_T by blast + then show "x \ T (\i. u (Y i))" + using T_LUB[OF ch2ch_cont[OF 1 3]] limproc_is_thelub[OF ch2ch_cont[OF 1 3]] 5 by force +qed + +lemma le_D_adm: "cont (u::('a::cpo) \ 'b process) \ monofun v \ adm(\x. u x \\<^sub>D v x)" +proof(auto simp add:cont2contlubE adm_def divergence_refine_def) + fix Y x + assume 1:"cont u" + and 2:"monofun v" + and 3:"chain Y" + and 4:"\i. D (v (Y i)) \ D (u (Y i))" + and 5:" x \ D (v (\i. Y i))" + hence "v (Y i) \ v (\i. Y i)" for i by (simp add: is_ub_thelub monofunE) + hence "D (v (\i. Y i)) \ D (u (Y i))" for i using 4 le_approx1 by blast + then show "x \ D (\i. u (Y i))" + using D_LUB[OF ch2ch_cont[OF 1 3]] limproc_is_thelub[OF ch2ch_cont[OF 1 3]] 5 by force +qed + +lemma le_DT_adm: "cont (u::('a::cpo) \ 'b process) \ monofun v \ adm(\x. u x \\<^sub>D\<^sub>T v x)" + using adm_conj[OF le_T_adm[of u v] le_D_adm[of u v]] by (simp add:trace_divergence_refine_def) + +lemmas le_FD_adm = le_adm[simplified failure_divergence_refine_def[symmetric]] + +section \Monotonicity\ + +lemma mono_det_D[simp]: "\P \\<^sub>D P'; S \\<^sub>D S'\ \ (P \ S) \\<^sub>D (P' \ S')" + by (metis D_det Un_mono divergence_refine_def) + +lemma mono_det_T[simp]: "\P \\<^sub>T P'; S \\<^sub>T S'\ \ (P \ S) \\<^sub>T (P' \ S')" + by (metis T_det Un_mono trace_refine_def) + +corollary mono_det_DT[simp]: "\P \\<^sub>D\<^sub>T P'; S \\<^sub>D\<^sub>T S'\ \ (P \ S) \\<^sub>D\<^sub>T (P' \ S')" + by (simp_all add: trace_divergence_refine_def) + +lemmas mono_det_FD[simp]= mono_det_FD[simplified failure_divergence_refine_def[symmetric]] + +\\Deterministic choice monotony doesn't hold for \\\<^sub>F\\ + + +lemma mono_ndet_F[simp]: "\P \\<^sub>F P'; S \\<^sub>F S'\ \ (P \ S) \\<^sub>F (P' \ S')" + by (metis F_ndet Un_mono failure_refine_def) + +lemma mono_ndet_D[simp]: "\P \\<^sub>D P'; S \\<^sub>D S'\ \ (P \ S) \\<^sub>D (P' \ S')" + by (metis D_ndet Un_mono divergence_refine_def) + +lemma mono_ndet_T[simp]: "\P \\<^sub>T P'; S \\<^sub>T S'\ \ (P \ S) \\<^sub>T (P' \ S')" + by (metis T_ndet Un_mono trace_refine_def) + +corollary mono_ndet_DT[simp]: "\P \\<^sub>D\<^sub>T P'; S \\<^sub>D\<^sub>T S'\ \ (P \ S) \\<^sub>D\<^sub>T (P' \ S')" + by (auto simp add: trace_divergence_refine_def) + +lemmas mono_ndet_FD[simp]= + mono_ndet_FD[simplified failure_divergence_refine_def[symmetric]] + +lemma mono_ndet_F_left[simp]: "P \\<^sub>F Q \ (P \ S) \\<^sub>F Q" + by (simp add: F_ndet failure_refine_def order_trans) + +lemma mono_ndet_D_left[simp]: "P \\<^sub>D Q \ (P \ S) \\<^sub>D Q" + by (simp add: D_ndet divergence_refine_def order_trans) + +lemma mono_ndet_T_left[simp]: "P \\<^sub>T Q \ (P \ S) \\<^sub>T Q" + by (simp add: T_ndet trace_refine_def order_trans) + +corollary mono_ndet_DT_left[simp]: "P \\<^sub>D\<^sub>T Q \ (P \ S) \\<^sub>D\<^sub>T Q" + and mono_ndet_F_right[simp]: "P \\<^sub>F Q \ (S \ P) \\<^sub>F Q" + and mono_ndet_D_right[simp]: "P \\<^sub>D Q \ (S \ P) \\<^sub>D Q" + and mono_ndet_T_right[simp]: "P \\<^sub>T Q \ (S \ P) \\<^sub>T Q" + and mono_ndet_DT_right[simp]: "P \\<^sub>D\<^sub>T Q \ (S \ P) \\<^sub>D\<^sub>T Q" + by (simp_all add: trace_divergence_refine_def Ndet_commute) + +lemmas +mono_ndet_FD_left[simp] = + mono_ndet_FD_left[simplified failure_divergence_refine_def[symmetric]] and +mono_ndet_FD_right[simp] = + mono_ndet_FD_right[simplified failure_divergence_refine_def[symmetric]] + +lemma mono_ndet_det_FD[simp]: "(P \ S) \\<^sub>F\<^sub>D (P \ S)" + by (metis det_id failure_divergence_refine_def mono_det_FD mono_ndet_FD_left + mono_ndet_FD_right order_refl) + +corollary mono_ndet_det_F[simp]: "(P \ S) \\<^sub>F (P \ S)" + and mono_ndet_det_D[simp]: "(P \ S) \\<^sub>D (P \ S)" + and mono_ndet_det_T[simp]: "(P \ S) \\<^sub>T (P \ S)" + and mono_ndet_det_DT[simp]: "(P \ S) \\<^sub>D\<^sub>T (P \ S)" + by (simp_all add: FD_F FD_D le_F_T D_T_DT) + + +lemma mono_seq_F_right[simp]: "S \\<^sub>F S' \ (P `;` S) \\<^sub>F (P `;` S')" + apply (auto simp: failure_refine_def F_seq append_single_T_imp_tickFree) + using NF_ND by fastforce+ + +lemma mono_seq_D_right[simp]: "S \\<^sub>D S' \ (P `;` S) \\<^sub>D (P `;` S')" + by (auto simp: divergence_refine_def D_seq) + +lemma mono_seq_T_right[simp]: "S \\<^sub>T S' \ (P `;` S) \\<^sub>T (P `;` S')" + apply (auto simp: trace_refine_def T_seq append_single_T_imp_tickFree) + using D_T by fastforce+ + +corollary mono_seq_DT_right[simp]: "S \\<^sub>D\<^sub>T S' \ (P `;` S) \\<^sub>D\<^sub>T (P `;` S')" + by (simp add: trace_divergence_refine_def) + +lemma mono_seq_DT_left[simp]: "P \\<^sub>D\<^sub>T P' \ (P `;` S) \\<^sub>D\<^sub>T (P' `;` S)" + apply (auto simp: trace_divergence_refine_def divergence_refine_def trace_refine_def D_seq T_seq) + by (metis (no_types, lifting) Nil_elem_T Process.F_T T_F append.right_neutral + is_processT5_S3 subset_eq) + +\\left sequence monotony doesn't hold for \\\<^sub>F\, \\\<^sub>D\ and \\\<^sub>T\\ + +corollary mono_seq_DT[simp]: "P \\<^sub>D\<^sub>T P' \ S \\<^sub>D\<^sub>T S' \ (P `;` S) \\<^sub>D\<^sub>T (P' `;` S')" + using mono_seq_DT_left mono_seq_DT_right trans_DT by blast + +lemmas mono_seq_FD[simp]= mono_seq_FD[simplified failure_divergence_refine_def[symmetric]] + + +lemma mono_mprefix_F_process[simp]: "\x. P x \\<^sub>F Q x \ Mprefix A P \\<^sub>F Mprefix A Q" + by (auto simp: failure_refine_def F_Mprefix) + +lemma mono_mprefix_D_process[simp]: "\x. P x \\<^sub>D Q x \ Mprefix A P \\<^sub>D Mprefix A Q" + by (auto simp: divergence_refine_def D_Mprefix) + +lemma mono_mprefix_T_process[simp]: "\x. P x \\<^sub>T Q x \ Mprefix A P \\<^sub>T Mprefix A Q" + by (auto simp: trace_refine_def T_Mprefix) + +corollary mono_mprefix_DT_process[simp]: "\x. P x \\<^sub>D\<^sub>T Q x \ Mprefix A P \\<^sub>D\<^sub>T Mprefix A Q" + by (simp add: trace_divergence_refine_def) + +lemmas +mono_mprefix_FD_process[simp] = + mono_mprefix_FD[simplified failure_divergence_refine_def[symmetric]] + +lemma mono_mprefix_DT_set[simp]: "A \ B \ Mprefix B P \\<^sub>D\<^sub>T Mprefix A P" + by (auto simp add:T_Mprefix D_Mprefix trace_divergence_refine_def + trace_refine_def divergence_refine_def) + +corollary mono_mprefix_D_set[simp]: "A \ B \ Mprefix B P \\<^sub>D Mprefix A P" + and mono_mprefix_T_set[simp]: "A \ B \ Mprefix B P \\<^sub>T Mprefix A P" + by (simp_all add: DT_D DT_T) + +\\Mprefix set monotony doesn't hold for \\\<^sub>F\ and \\\<^sub>F\<^sub>D\ where it holds for deterministic choice\ + + +lemma mono_hide_F[simp]: "P \\<^sub>F Q \ P \\ A \\<^sub>F Q \\ A" + apply(cases "A={}", simp_all add: hide_set_empty failure_refine_def F_hiding, intro conjI, blast) +proof(subst (2) Un_commute, rule subsetI, rule UnCI, auto, goal_cases) + case (1 b t u) + then obtain a where a:"a \ A" by blast + define f where A:"f = rec_nat t (\i t. t @ [ev a])" + hence AA:"f (Suc i) = (f i) @ [ev a]" for i by simp + hence B:"strict_mono f" by (simp add:strict_mono_def lift_Suc_mono_less_iff) + from A have C:"t \ range f" by (metis (mono_tags, lifting) old.nat.simps(6) range_eqI) + { fix i + have "f i \ D Q \ tickFree (f i) \ trace_hide (f i) (ev ` A) = (trace_hide t (ev ` A))" + proof(induct i) + case 0 + then show ?case by (simp add: 1(4) 1(7) A) + next + case (Suc i) + then show ?case + apply (simp add:AA a) + using is_processT7[rule_format, of "f i" Q "[ev a]"] front_tickFree_single by blast + qed + } + with B C have "isInfHiddenRun f P A \ t \ range f" + by (metis 1(1) D_T F_subset_imp_T_subset subsetD) + with 1 show ?case by metis +next + case (2 b u f x) + then show ?case using F_subset_imp_T_subset by blast +qed + +lemma mono_hide_T[simp]: "P \\<^sub>T Q \ P \\ A \\<^sub>T Q \\ A" + apply(cases "A={}", simp add: hide_set_empty, simp add:trace_refine_def T_hiding, intro conjI) +proof(goal_cases) + case 1 + then show ?case + proof(subst Un_commute, rule_tac subsetI, rule_tac UnCI, auto, goal_cases) + case 11:(1 t a) + hence tt:"t \ T P" by (meson Process.F_T subset_eq) + with 11(1) inf_hidden[of A t P] obtain f where "isInfHiddenRun f P A \ t \ range f" by auto + with 11(4)[rule_format, of t "[]"] show ?case + by (metis 11(1) tt append_Nil2 front_tickFree_Nil is_processT2_TR + nonTickFree_n_frontTickFree tick_T_F) + qed +next + case 2 + then show ?case + proof(subst Un_commute, auto, goal_cases) + case 21:(1 t u a) + define f where A:"f = rec_nat t (\i t. t @ [ev a])" + hence AA:"f (Suc i) = (f i) @ [ev a]" for i by simp + hence B:"strict_mono f" by (simp add:strict_mono_def lift_Suc_mono_less_iff) + from A have C:"t \ range f" + by (metis (mono_tags, lifting) old.nat.simps(6) range_eqI) + { fix i + have "f i \ D Q \ tickFree (f i) \ trace_hide (f i) (ev ` A) = (trace_hide t (ev ` A))" + proof(induct i) + case 0 + then show ?case by (simp add: 21(4) 21(7) A) + next + case (Suc i) + then show ?case + apply (simp add:AA 21(6)) + using is_processT7[rule_format, of "f i" Q "[ev a]"] front_tickFree_single by blast + qed + } + with B C have "isInfHiddenRun f P A \ t \ range f" by (metis 21(1) D_T subsetD) + with 21 show ?case by metis + next + case 22:(2 b u f x) + then show ?case by blast + qed +qed + +lemma mono_hide_DT[simp]: "P \\<^sub>D\<^sub>T Q \ P \\ A \\<^sub>D\<^sub>T Q \\ A" +proof - + assume as:"P \\<^sub>D\<^sub>T Q" + then have "P \\ A \\<^sub>D Q \\ A" + apply(auto simp:trace_divergence_refine_def divergence_refine_def + trace_refine_def D_hiding T_hiding) + by blast+ + with DT_T[THEN mono_hide_T, OF as] show ?thesis by (simp add: trace_divergence_refine_def) +qed + +lemmas mono_hide_FD[simp] = + mono_hide_FD[simplified failure_divergence_refine_def[symmetric]] + +\\Obviously, Hide monotony doesn't hold for \\\<^sub>D\\ + + +lemma mono_sync_DT[simp]: "P \\<^sub>D\<^sub>T P' \ Q \\<^sub>D\<^sub>T Q' \ (P \ A \ Q) \\<^sub>D\<^sub>T (P' \ A \ Q')" + by (simp add:trace_divergence_refine_def divergence_refine_def trace_refine_def T_sync D_sync) + blast + +lemmas mono_sync_FD[simp] = + mono_sync_FD[simplified failure_divergence_refine_def[symmetric]] + +\\synchronization monotony doesn't hold for \\\<^sub>F\, \\\<^sub>D\ and \\\<^sub>T\\ + + +lemma mono_mndet_F_process[simp]: "\x\A. P x \\<^sub>F Q x \ mndet A P \\<^sub>F mndet A Q" + by (cases "A = {}", auto simp: failure_refine_def F_mndet write0_def F_Mprefix) + +lemma mono_mndet_D_process[simp]: "\x\A. P x \\<^sub>D Q x \ mndet A P \\<^sub>D mndet A Q" + by (cases "A = {}", auto simp: divergence_refine_def D_mndet write0_def D_Mprefix) + +lemma mono_mndet_T_process[simp]: "\x\A. P x \\<^sub>T Q x \ mndet A P \\<^sub>T mndet A Q" + by (cases "A = {}", auto simp: trace_refine_def T_mndet write0_def T_Mprefix) + +corollary mono_mndet_DT_process[simp]: "\x\A. P x \\<^sub>D\<^sub>T Q x \ mndet A P \\<^sub>D\<^sub>T mndet A Q" + by (simp add: trace_divergence_refine_def) + +lemmas +mono_mndet_FD_process[simp] = + mono_mndet_FD[simplified failure_divergence_refine_def[symmetric]] + +lemmas +mono_mndet_FD_set[simp] = + mndet_FD_subset[simplified failure_divergence_refine_def[symmetric]] + +corollary mono_mndet_F_set[simp]: "A \ {} \ A \ B \ mndet B P \\<^sub>F mndet A P" + and mono_mndet_D_set[simp]: "A \ {} \ A \ B \ mndet B P \\<^sub>D mndet A P" + and mono_mndet_T_set[simp]: "A \ {} \ A \ B \ mndet B P \\<^sub>T mndet A P" + and mono_mndet_DT_set[simp]: "A \ {} \ A \ B \ mndet B P \\<^sub>D\<^sub>T mndet A P" + by (simp_all add: FD_F FD_D le_F_T D_T_DT) + +lemmas +Mprefix_refines_Mndet_FD[simp] = + Mprefix_refines_Mndet[simplified failure_divergence_refine_def[symmetric]] + +corollary Mprefix_refines_Mndet_F[simp]: "mndet A P \\<^sub>F Mprefix A P" + and Mprefix_refines_Mndet_D[simp]: "mndet A P \\<^sub>D Mprefix A P" + and Mprefix_refines_Mndet_T[simp]: "mndet A P \\<^sub>T Mprefix A P" + and Mprefix_refines_Mndet_DT[simp]: "mndet A P \\<^sub>D\<^sub>T Mprefix A P" + by (simp_all add: FD_F FD_D le_F_T D_T_DT) + + +section \Reference processes and their unfolding rules\ + +thm DF_def DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def RUN_def CHAOS_def + +definition CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P :: "'a set \ 'a process" + where "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \ \ X. (SKIP \ STOP \ (\ x \ A \ X))" + + +thm DF_unfold DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold + +lemma RUN_unfold : "RUN A = (\ z \ A \ RUN A)" + by(simp add: RUN_def, rule trans, rule fix_eq, simp) + +lemma CHAOS_unfold : "CHAOS A = (STOP \ (\ z \ A \ CHAOS A))" + by(simp add: CHAOS_def, rule trans, rule fix_eq, simp) + +lemma CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \ SKIP \ STOP \ (\ x \ A \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" + unfolding CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def using fix_eq[of "\ X. (SKIP \ STOP \ (\ x \ A \ X))"] by simp + +section \Process events and reference processes events\ + +definition events_of where "events_of P \ (\t\T P. {a. ev a \ set t})" + +lemma events_DF: "events_of (DF A) = A" +proof(auto simp add:events_of_def) + fix x t + show "t \ T (DF A) \ ev x \ set t \ x \ A" + proof(induct t) + case Nil + then show ?case by simp + next + case (Cons a t) + from Cons(2) have "a # t \ T (\z\A \ DF A)" using DF_unfold by metis + then obtain aa where "a = ev aa \ aa \ A \ t \ T (DF A)" + by (cases "A={}", auto simp add:T_mndet write0_def T_Mprefix T_STOP) + with Cons show ?case by auto + qed +next + fix x + show "x \ A \ \xa\T (DF A). ev x \ set xa" + apply(subst DF_unfold, cases "A={}", auto simp add:T_mndet write0_def T_Mprefix) + by (metis Nil_elem_T list.sel(1) list.sel(3) list.set_intros(1)) +qed + +lemma events_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P: "events_of (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = A" +proof(auto simp add:events_of_def) + fix x t + show "t \ T (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) \ ev x \ set t \ x \ A" + proof(induct t) + case Nil + then show ?case by simp + next + case (Cons a t) + from Cons(2) have "a # t \ T ((\z\A \ DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) \ SKIP)" using DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold by metis + with Cons obtain aa where "a = ev aa \ aa \ A \ t \ T (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" + by (cases "A={}", auto simp add:T_mndet write0_def T_Mprefix T_STOP T_SKIP T_ndet) + with Cons show ?case by auto + qed +next + fix x + show "x \ A \ \xa\T (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A). ev x \ set xa" + apply(subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, cases "A={}") + apply(auto simp add:T_mndet write0_def T_Mprefix T_SKIP T_ndet) + by (metis Nil_elem_T list.sel(1) list.sel(3) list.set_intros(1)) +qed + +lemma events_RUN: "events_of (RUN A) = A" +proof(auto simp add:events_of_def) + fix x t + show "t \ T (RUN A) \ ev x \ set t \ x \ A" + proof(induct t) + case Nil + then show ?case by simp + next + case (Cons a t) + from Cons(2) have "a # t \ T (\z\A \ RUN A)" using RUN_unfold by metis + then obtain aa where "a = ev aa \ aa \ A \ t \ T (RUN A)" by (auto simp add:T_Mprefix) + with Cons show ?case by auto + qed +next + fix x + show "x \ A \ \xa\T (RUN A). ev x \ set xa" + apply(subst RUN_unfold, simp add: T_Mprefix) + by (metis Nil_elem_T list.sel(1) list.sel(3) list.set_intros(1)) +qed + +lemma events_CHAOS: "events_of (CHAOS A) = A" +proof(auto simp add:events_of_def) + fix x t + show "t \ T (CHAOS A) \ ev x \ set t \ x \ A" + proof(induct t) + case Nil + then show ?case by simp + next + case (Cons a t) + from Cons(2) have "a # t \ T (STOP \ (\ z \ A \ CHAOS A))" using CHAOS_unfold by metis + then obtain aa where "a = ev aa \ aa \ A \ t \ T (CHAOS A)" + by (auto simp add:T_ndet T_Mprefix T_STOP) + with Cons show ?case by auto + qed +next + fix x + show "x \ A \ \xa\T (CHAOS A). ev x \ set xa" + apply(subst CHAOS_unfold, simp add:T_ndet T_Mprefix T_STOP) + by (metis Nil_elem_T list.sel(1) list.sel(3) list.set_intros(1)) +qed + +lemma events_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P: "events_of (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = A" +proof(auto simp add:events_of_def) + fix x t + show "t \ T (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) \ ev x \ set t \ x \ A" + proof(induct t) + case Nil + then show ?case by simp + next + case (Cons a t) + from Cons(2) have "a # t \ T (SKIP \ STOP \ (\ z \ A \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A))" + using CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold by metis + with Cons obtain aa where "a = ev aa \ aa \ A \ t \ T (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" + by (auto simp add:T_ndet T_Mprefix T_STOP T_SKIP) + with Cons show ?case by auto + qed +next + fix x + show "x \ A \ \xa\T (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A). ev x \ set xa" + apply(subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, simp add:T_ndet T_Mprefix T_STOP T_SKIP) + by (metis Nil_elem_T list.sel(1) list.sel(3) list.set_intros(1)) +qed + +lemma events_div: "D(P) \ {} \ events_of (P) = UNIV" +proof(auto simp add:events_of_def) + fix x xa + assume 1: "x \ D P" + show "\x\T P. ev xa \ set x" + proof(cases "tickFree x") + case True + hence "x@[ev xa] \ T P" + using 1 NT_ND front_tickFree_single is_processT7 by blast + then show ?thesis by force + next + case False + hence "(butlast x)@[ev xa] \ T P" + by (metis "1" D_T D_imp_front_tickFree append_single_T_imp_tickFree butlast_snoc + front_tickFree_single nonTickFree_n_frontTickFree process_charn) + then show ?thesis by force + qed +qed + + +thm DF_subset + +lemma DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_subset_FD: "A \ {} \ A \ B \ DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P B \\<^sub>F\<^sub>D DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" + apply(subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule fix_ind, rule le_FD_adm, simp_all add:monofunI, subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold) + by (rule mono_ndet_FD, simp_all) (meson mono_mndet_FD_process mono_mndet_FD_set trans_FD) + +lemma RUN_subset_DT: "A \ B \ RUN B \\<^sub>D\<^sub>T RUN A" + apply(subst RUN_def, rule fix_ind, rule le_DT_adm, simp_all add:monofunI, subst RUN_unfold) + by (meson mono_mprefix_DT_process mono_mprefix_DT_set trans_DT) + +lemma CHAOS_subset_FD: "A \ B \ CHAOS B \\<^sub>F\<^sub>D CHAOS A" + apply(subst CHAOS_def, rule fix_ind, rule le_FD_adm, simp_all add:monofunI, subst CHAOS_unfold) + by (auto simp add: failure_divergence_refine_def le_ref_def + D_Mprefix D_ndet F_STOP F_Mprefix F_ndet) + +lemma CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_subset_FD: "A \ B \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P B \\<^sub>F\<^sub>D CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" + apply(subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule fix_ind, rule le_FD_adm) + apply(simp_all add:monofunI, subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold) + by (auto simp add: failure_divergence_refine_def le_ref_def + D_Mprefix D_ndet F_STOP F_Mprefix F_ndet) + +section \Relations between refinements on reference processes\ + +lemma CHAOS_has_all_tickFree_failures: + "tickFree a \ {x. ev x \ set a} \ A \ (a,b) \ F (CHAOS A)" +proof(induct a) + case Nil + then show ?case + by (subst CHAOS_unfold, simp add:F_ndet F_STOP) +next + case (Cons a0 al) + hence "tickFree al" + by (metis append.left_neutral append_Cons front_tickFree_charn front_tickFree_mono) + with Cons show ?case + apply (subst CHAOS_unfold, simp add:F_ndet F_STOP F_Mprefix events_of_def) + using event_set by blast +qed + +lemma CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures: + assumes as:"(events_of P) \ A" + shows "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F P" +proof - + have "front_tickFree a \ set a \ (\t\T P. set t) \ (a,b) \ F (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" for a b + proof(induct a) + case Nil + then show ?case + by (subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, simp add:F_ndet F_STOP) + next + case (Cons a0 al) + hence "front_tickFree al" + by (metis append.left_neutral append_Cons front_tickFree_charn front_tickFree_mono) + with Cons show ?case + apply (subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, simp add:F_ndet F_STOP F_SKIP F_Mprefix events_of_def as) + apply(cases "a0=tick") + apply (metis append.simps(2) front_tickFree_charn + front_tickFree_mono list.distinct(1) tickFree_Cons) + using event_set image_iff as[simplified events_of_def] by fastforce + qed + thus ?thesis + by (simp add: F_T SUP_upper failure_refine_def process_charn subrelI) +qed + +corollary CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures_ev: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P (events_of P) \\<^sub>F P" + and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures_Un: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>F P" + by (simp_all add: CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures) + + +lemma DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF_refine_F: "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F DF A" + by (simp add:DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule fix_ind, simp_all add:le_F_adm monofunI, subst DF_unfold, simp) + +lemma DF_RUN_refine_F: "DF A \\<^sub>F RUN A" + apply (simp add:DF_def, rule fix_ind, simp_all add:le_F_adm monofunI, subst RUN_unfold) + by (meson Mprefix_refines_Mndet_F mono_mndet_F_process trans_F) + +lemma CHAOS_DF_refine_F: "CHAOS A \\<^sub>F DF A" + apply (simp add:CHAOS_def DF_def, rule parallel_fix_ind, simp_all add: monofunI) + apply (rule le_F_adm, simp_all add: monofun_snd) + by (cases "A={}", auto simp add:adm_def failure_refine_def F_mndet + F_Mprefix write0_def F_ndet F_STOP) + +corollary CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_CHAOS_refine_F: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F CHAOS A" + and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_refine_F: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" + by (simp_all add: CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures events_CHAOS events_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P + trans_F[OF CHAOS_DF_refine_F DF_RUN_refine_F]) + + +lemma div_free_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P: "D(DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = {}" +proof(simp add:DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def fix_def) + define Y where "Y \ \i. iterate i\(\ x. (\xa\A \ x) \ SKIP)\\" + hence a:"chain Y" by simp + have "D (Y (Suc i)) = {d. d \ [] \ hd d \ (ev ` A) \ tl d \ D(Y i)}" for i + by (cases "A={}", auto simp add:Y_def D_STOP D_SKIP D_mndet write0_def D_Mprefix D_ndet) + hence b:"d \ D(Y i) \ length d \ i" for d i + by (induct i arbitrary:d, simp_all add: Nitpick.size_list_simp(2)) + { fix x + assume c:"\ i. x \ D (Y i)" + from b have "x \ D (Y (Suc (length x)))" using Suc_n_not_le_n by blast + with c have False by simp + } + with a show "D (\i. Y i) = {}" + using D_LUB[OF a] limproc_is_thelub[OF a] by auto +qed + +lemma div_free_DF: "D(DF A) = {}" +proof - + have "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>D DF A" + by (simp add:DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule fix_ind, simp_all add:le_D_adm monofunI, subst DF_unfold, simp) + then show ?thesis using divergence_refine_def div_free_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P by blast +qed + +lemma div_free_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P: "D (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = {}" +proof - + have "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>D CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" + proof (simp add:DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule fix_ind, simp_all add:le_D_adm monofunI, subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold) + fix x + assume 1:"x \\<^sub>D CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" + have a:"((\xa\A \ x) \ SKIP) = (SKIP \ SKIP \ (\xa\A \ x))" + by (simp add: Ndet_commute ndet_id) + from 1 have b:"(SKIP \ SKIP \ (\xa\A \ x)) \\<^sub>D (SKIP \ STOP \ (\xa\A \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A))" + by (meson DT_D Mprefix_refines_Mndet_D STOP_leDT idem_D + mono_mprefix_D_process mono_ndet_D trans_D) + from a b show "((\xa\A \ x) |-| SKIP) \\<^sub>D (SKIP |-| STOP |-| Mprefix A (\x. CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A))" + by simp + qed + then show ?thesis using divergence_refine_def div_free_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P by blast +qed + +lemma div_free_CHAOS: "D(CHAOS A) = {}" +proof - + have "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>D CHAOS A" + apply (simp add:CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule fix_ind) + by (simp_all add:le_D_adm monofunI, subst CHAOS_unfold, simp) + then show ?thesis using divergence_refine_def div_free_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P by blast +qed + +lemma div_free_RUN: "D(RUN A) = {}" +proof - + have "CHAOS A \\<^sub>D RUN A" + by (simp add:CHAOS_def, rule fix_ind, simp_all add:le_D_adm monofunI, subst RUN_unfold, simp) + then show ?thesis using divergence_refine_def div_free_CHAOS by blast +qed + +corollary DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF_refine_FD: "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F\<^sub>D DF A" + and DF_RUN_refine_FD: "DF A \\<^sub>F\<^sub>D RUN A" + and CHAOS_DF_refine_FD: "CHAOS A \\<^sub>F\<^sub>D DF A" + and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_CHAOS_refine_FD: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F\<^sub>D CHAOS A" + and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_refine_FD: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F\<^sub>D DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" + using div_free_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P[of A] div_free_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P[of A] div_free_DF[of A] div_free_RUN[of A] + div_free_CHAOS[of A] + F_D_FD[OF DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF_refine_F, of A] F_D_FD[OF DF_RUN_refine_F, of A] + F_D_FD[OF CHAOS_DF_refine_F, of A] F_D_FD[OF CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_CHAOS_refine_F, of A] + F_D_FD[OF CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_refine_F, of A] + by (auto simp add:divergence_refine_def) + + +lemma traces_CHAOS_sub: "T(CHAOS A) \ {s. set s \ ev ` A}" +proof(auto) + fix s sa + assume "s \ T (CHAOS A)" and "sa \ set s" + then show "sa \ ev ` A" + apply (induct s, simp) + by (subst (asm) (2) CHAOS_unfold, cases "A={}", auto simp add:T_ndet T_STOP T_Mprefix) +qed + +lemma traces_RUN_sub: "{s. set s \ ev ` A} \ T(RUN A)" +proof(auto) + fix s + assume "set s \ ev ` A" + then show "s \ T (RUN A)" + by (induct s, simp add: Nil_elem_T) (subst RUN_unfold, auto simp add:T_Mprefix) +qed + +corollary RUN_all_tickfree_traces1: "T(RUN A) = {s. set s \ ev ` A}" + and DF_all_tickfree_traces1: "T(DF A) = {s. set s \ ev ` A}" + and CHAOS_all_tickfree_traces1: "T(CHAOS A) = {s. set s \ ev ` A}" + using DF_RUN_refine_F[THEN le_F_T, simplified trace_refine_def] + CHAOS_DF_refine_F[THEN le_F_T,simplified trace_refine_def] + traces_CHAOS_sub traces_RUN_sub by blast+ + +corollary RUN_all_tickfree_traces2: "tickFree s \ s \ T(RUN UNIV)" + and DF_all_tickfree_traces2: "tickFree s \ s \ T(DF UNIV)" + and CHAOS_all_tickfree_trace2: "tickFree s \ s \ T(CHAOS UNIV)" + apply(simp_all add:tickFree_def RUN_all_tickfree_traces1 + DF_all_tickfree_traces1 CHAOS_all_tickfree_traces1) + by (metis event_set insertE subsetI)+ + +lemma traces_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_sub: "T(CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) \ {s. front_tickFree s \ set s \ (ev ` A \ {tick})}" +proof(auto simp add:is_processT2_TR) + fix s sa + assume "s \ T (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" and "sa \ set s" and "sa \ ev ` A" + then show "sa = tick" + apply (induct s, simp) + by (subst (asm) (2) CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, cases "A={}", auto simp add:T_ndet T_STOP T_SKIP T_Mprefix) +qed + +lemma traces_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_sub: + "{s. front_tickFree s \ set s \ (ev ` A \ {tick})} \ T(DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A::'a process)" +proof(auto) + fix s + assume a:"front_tickFree s" and b:"set s \ insert tick (ev ` A)" + have c:"front_tickFree ((tick::'a event) # s) \ s = []" for s + by (metis butlast.simps(2) butlast_snoc front_tickFree_charn list.distinct(1) tickFree_Cons) + with a b show "s \ T (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" + apply (induct s, simp add: Nil_elem_T, subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, cases "A={}") + apply (subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, cases "A={}") + apply(auto simp add:T_Mprefix T_mndet write0_def T_SKIP T_ndet T_STOP) + apply (metis append_Cons append_Nil front_tickFree_charn front_tickFree_mono) + by (metis append_Cons append_Nil front_tickFree_mono) + qed + +corollary DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces1: + "T(DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = {s. front_tickFree s \ set s \ (ev ` A \ {tick})}" + and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces1: + "T(CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = {s. front_tickFree s \ set s \ (ev ` A \ {tick})}" + using CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_refine_F[THEN le_F_T, simplified trace_refine_def] + traces_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_sub traces_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_sub by blast+ + +corollary DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces2: "front_tickFree s \ s \ T(DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV)" + and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces2: "front_tickFree s \ s \ T(CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV)" + apply(simp_all add:tickFree_def DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces1 + CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces1) + by (metis event_set subsetI)+ + +corollary DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_traces: "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>T P" + and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_traces: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>T P" + apply (simp add:trace_refine_def DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces2 is_processT2_TR subsetI) + by (simp add:trace_refine_def CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces2 is_processT2_TR subsetI) + +end diff --git a/thys/CSP_RefTK/Conclusion.thy b/thys/CSP_RefTK/Conclusion.thy new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/Conclusion.thy @@ -0,0 +1,81 @@ +(*<*) +\\ ******************************************************************** + * Project : CSP-RefTK - A Refinement Toolkit for HOL-CSP + * Version : 1.0 + * + * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * + * This file : Conclusion + * + * Copyright (c) 2020 Université Paris-Saclay, 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. + ******************************************************************************\ +(*>*) + +chapter\Conclusion\ +(*<*) +theory Conclusion + imports HOLCF +begin +(*>*) + +text\ We presented a formalisation of the most comprehensive semantic model for CSP, a 'classical' +language for the specification and analysis of concurrent systems studied in a rich body of +literature. For this purpose, we ported @{cite "tej.ea:corrected:1997"} to a modern version +of Isabelle, restructured the proofs, and extended the resulting theory of the language +substantially. The result HOL-CSP 2 has been submitted to the Isabelle AFP @{cite "HOL-CSP-AFP"}, +thus a fairly sustainable format accessible to other researchers and tools. + +We developed a novel set of deadlock - and livelock inference proof principles based on +classical and denotational characterizations. In particular, we formally investigated the relations +between different refinement notions in the presence of deadlock - and livelock; an area where +traditional CSP literature skates over the nitty-gritty details. Finally, we demonstrated how to +exploit these results for deadlock/livelock analysis of protocols. + +We put a large body of abstract CSP laws and induction principles together to form +concrete verification technologies for generalized classical problems, which have been considered +so far from the perspective of data-independence or structural parametricity. The underlying novel +principle of ``trading rich structure against rich state'' allows one to convert processes +into classical transition systems for which established invariant techniques become applicable. + +Future applications of HOL-CSP 2 could comprise a combination with model checkers, where our theory +with its derived rules can be used to certify the output of a model-checker over CSP. In our experience, +labelled transition systems generated by model checkers may be used to steer inductions or to construct +the normalized processes \P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\\\<^sub>,\\\ automatically, thus combining efficient finite reasoning +over finite sub-systems with globally infinite systems in a logically safe way. +\ + + +(*<*) +end +(*>*) + diff --git a/thys/CSP_RefTK/CopyBuffer_props.thy b/thys/CSP_RefTK/CopyBuffer_props.thy new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/CopyBuffer_props.thy @@ -0,0 +1,129 @@ +(*<*) +\\ ******************************************************************** + * Project : CSP-RefTK - A Refinement Toolkit for HOL-CSP + * Version : 1.0 + * + * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * + * This file : The Copy-Buffer Example Revisited + * + * Copyright (c) 2020 Université Paris-Saclay, 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. + ******************************************************************************\ +(*>*) + +chapter\Examples\ + +section\CopyBuffer Refinement over an infinite alphabet\ + +theory CopyBuffer_props + imports "HOL-CSP.CopyBuffer" "Properties" +begin + +subsection\ The Copy-Buffer vs. reference processes \ + +lemma DF_COPY: "(DF (range left \ range right)) \\<^sub>F\<^sub>D COPY" + apply(simp add:DF_def,rule fix_ind2, simp_all) + unfolding failure_divergence_refine_def +proof - + show "adm (\a. a \ COPY)" by(rule le_adm, simp_all add:monofunI) +next + have 1:"(\xa\ range left \ range right \ \) \ (\xa\ range left \ \)" + using mndet_subset_FD by (metis UNIV_I empty_iff imageI) + have 2:"(\xa\ range left \ \) \ (left`?`x \ \)" + unfolding read_def + by (metis Mprefix_refines_Mndet comp_apply dual_order.antisym mono_mprefix_FD order_refl) + + show "(\x\range left \ range right \ \) \ COPY" + by (metis (mono_tags, lifting) 1 2 COPY_rec bot_less1 mono_read_FD order.trans) +next + fix P::"'a channel process" + assume *: "P \ COPY" and ** : "(\x\range left \ range right \ P) \ COPY" + have 1:"(\xa\ range left \ range right \ P) \ (\xa\ range right \ P)" + using mndet_subset_FD by (metis UNIV_I Un_commute empty_iff imageI) + have 2:"(\xa\ range right \ P) \ (right`!`x \ P)" for x + using mndet_subset_FD[of "{right x}" "range right"] + by(simp add:write_def write0_def mndet_unit) + from 1 2 have ab:"(\xa\ range left \ range right \ P) \ (right`!`x \ P)" for x + using dual_order.trans by blast + hence 3:"(left`?`x \ (\xa\ range left \ range right \ P)) \ (left`?`x \(right`!`x \ P))" + by (simp add: mono_read_FD) + have 4:"\X. (\xa\ range left \ range right \ X) \ (\xa\ range left \ X)" + using mndet_subset_FD by (metis UNIV_I empty_iff imageI) + have 5:"\X. (\xa\ range left \ X) \ (left`?`x \ X)" + unfolding read_def + by (metis Mprefix_refines_Mndet comp_apply dual_order.antisym mono_mprefix_FD order_refl) + from 3 4[of "(\xa\ range left \ range right \ P)"] + 5 [of "(\xa\ range left \ range right \ P)"] + have 6:"(\xa\ range left \ range right \ + (\xa\ range left \ range right \ P)) \ (left`?`x \ (right`!`x \ P))" + by (meson dual_order.trans) + from * ** have 7:"(left`?`x \ (right`!`x \ P)) \ (left`?`x \ (right`!`x \ COPY))" + by (simp add: mono_read_FD mono_write_FD) + + show "(\x\range left \ range right \ \x\range left \ range right \ P) \ COPY" + by (metis (mono_tags, lifting) 6 7 COPY_rec dual_order.trans) +qed + +subsection\ ... and abstract consequences \ + +corollary df_COPY: "deadlock_free COPY" + and lf_COPY: "lifelock_free COPY" + apply (meson DF_COPY DF_Univ_freeness UNIV_not_empty image_is_empty sup_eq_bot_iff) + by (meson CHAOS_DF_refine_FD DF_COPY DF_Univ_freeness UNIV_not_empty deadlock_free_def + image_is_empty lifelock_free_def sup_eq_bot_iff trans_FD) + +corollary df_v2_COPY: "deadlock_free_v2 COPY" + and lf_v2_COPY: "lifelock_free_v2 COPY" + and nt_COPY: "non_terminating COPY" + apply (simp add: df_COPY deadlock_free_is_deadlock_free_v2) + apply (simp add: lf_COPY lifelock_free_is_lifelock_free_v2) + using lf_COPY lifelock_free_is_non_terminating by blast + +lemma DF_SYSTEM: "DF UNIV \\<^sub>F\<^sub>D SYSTEM" + using DF_subset[of "(range left \ range right)" UNIV, simplified] + impl_refines_spec[THEN le_approx_implies_le_ref] DF_COPY + failure_divergence_refine_def trans_FD by blast + +corollary df_SYSTEM: "deadlock_free SYSTEM" + and lf_SYSTEM: "lifelock_free SYSTEM" + apply (simp add: DF_SYSTEM deadlock_free_def) + using CHAOS_DF_refine_FD DF_SYSTEM lifelock_free_def trans_FD by blast + +corollary df_v2_SYSTEM: "deadlock_free_v2 SYSTEM" + and lf_v2_SYSTEM: "lifelock_free_v2 SYSTEM" + and nt_SYSTEM: "non_terminating SYSTEM" + apply (simp add: df_SYSTEM deadlock_free_is_deadlock_free_v2) + apply (simp add: lf_SYSTEM lifelock_free_is_lifelock_free_v2) + using lf_SYSTEM lifelock_free_is_non_terminating by blast + +end diff --git a/thys/CSP_RefTK/DiningPhilosophers.thy b/thys/CSP_RefTK/DiningPhilosophers.thy new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/DiningPhilosophers.thy @@ -0,0 +1,970 @@ +(*<*) +\\ ******************************************************************** + * Project : CSP-RefTK - A Refinement Toolkit for HOL-CSP + * Version : 1.0 + * + * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * + * This file : Example on Structural Parameterisation: Dining Philosophers + * + * Copyright (c) 2020 Université Paris-Saclay, 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. + * + * TrHIS SOFTrWARE IS PROVIDED BY TrHE COPYRIGHTr HOLDERS AND CONTrRIBUTrORS + * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTrIES, INCLUDING, BUTr NOTr + * LIMITrED TrO, TrHE IMPLIED WARRANTrIES OF MERCHANTrABILITrY AND FITrNESS FOR + * A PARTrICULAR PURPOSE ARE DISCLAIMED. IN NO EVENTr SHALL TrHE COPYRIGHTr + * OWNER OR CONTrRIBUTrORS BE LIABLE FOR ANY DIRECTr, INDIRECTr, INCIDENTrAL, + * SPECIAL, EXEMPLARY, OR CONSEQUENTrIAL DAMAGES (INCLUDING, BUTr NOTr + * LIMITrED TrO, PROCUREMENTr OF SUBSTrITrUTrE GOODS OR SERVICES; LOSS OF USE, + * DATrA, OR PROFITrS; OR BUSINESS INTrERRUPTrION) HOWEVER CAUSED AND ON ANY + * TrHEORY OF LIABILITrY, WHETrHER IN CONTrRACTr, STrRICTr LIABILITrY, OR TrORTr + * (INCLUDING NEGLIGENCE OR OTrHERWISE) ARISING IN ANY WAY OUTr OF TrHE USE + * OF TrHIS SOFTrWARE, EVEN IF ADVISED OF TrHE POSSIBILITrY OF SUCH DAMAGE. + ******************************************************************************\ +(*>*) + +section\ Generalized Dining Philosophers \ + +theory DiningPhilosophers + imports "Process_norm" +begin + +subsection \Preliminary lemmas for proof automation\ + +lemma Suc_mod: "n > 1 \ i \ Suc i mod n" + by (metis One_nat_def mod_Suc mod_if mod_mod_trivial n_not_Suc_n) + +lemmas suc_mods = Suc_mod Suc_mod[symmetric] + +lemma l_suc: "n > 1 \ \ n \ Suc 0" + by simp + +lemma minus_suc: "n > 0 \ n - Suc 0 \ n" + by linarith + +lemma numeral_4_eq_4:"4 = Suc (Suc (Suc (Suc 0)))" + by simp + +lemma numeral_5_eq_5:"5 = Suc (Suc (Suc (Suc (Suc 0))))" + by simp + +subsection\The dining processes definition\ + +locale DiningPhilosophers = + + fixes N::nat + assumes N_g1[simp] : "N > 1" + +begin + +datatype dining_event = picks (phil:nat) (fork:nat) + | putsdown (phil:nat) (fork:nat) + +definition RPHIL:: "nat \ dining_event process" + where "RPHIL i = (\ X. (picks i i \ (picks i (i-1) \ (putsdown i (i-1) \ (putsdown i i \ X)))))" + +definition LPHIL0:: "dining_event process" + where "LPHIL0 = (\ X. (picks 0 (N-1) \ (picks 0 0 \ (putsdown 0 0 \ (putsdown 0 (N-1) \ X)))))" + +definition FORK :: "nat \ dining_event process" + where "FORK i = (\ X. (picks i i \ (putsdown i i \ X)) + \ (picks ((i+1) mod N) i \ (putsdown ((i+1) mod N) i \ X)))" + + +abbreviation "foldPHILs n \ fold (\ i P. P ||| RPHIL i) [1..< n] (LPHIL0)" +abbreviation "foldFORKs n \ fold (\ i P. P ||| FORK i) [1..< n] (FORK 0)" + +abbreviation "PHILs \ foldPHILs N" +abbreviation "FORKs \ foldFORKs N" + +corollary FORKs_def2: "FORKs = fold (\ i P. P ||| FORK i) [0..< N] SKIP" + using N_g1 by (subst (2) upt_rec, auto) (metis (no_types, lifting) Inter_commute Inter_skip1) + +corollary "N = 3 \ PHILs = (LPHIL0 ||| RPHIL 1 ||| RPHIL 2)" + by (subst upt_rec, auto simp add:numeral_2_eq_2)+ + + +definition DINING :: "dining_event process" + where "DINING = (FORKs || PHILs)" + +subsubsection \Unfolding rules\ + +lemma RPHIL_rec: + "RPHIL i = (picks i i \ (picks i (i-1) \ (putsdown i (i-1) \ (putsdown i i \ RPHIL i))))" + by (simp add:RPHIL_def write0_def, subst fix_eq, simp) + +lemma LPHIL0_rec: + "LPHIL0 = (picks 0 (N-1) \ (picks 0 0 \ (putsdown 0 0 \ (putsdown 0 (N-1) \ LPHIL0))))" + by (simp add:LPHIL0_def write0_def, subst fix_eq, simp) + +lemma FORK_rec: "FORK i = ( (picks i i \ (putsdown i i \ (FORK i))) + \ (picks ((i+1) mod N) i \ (putsdown ((i+1) mod N) i \ (FORK i))))" + by (simp add:FORK_def write0_def, subst fix_eq, simp) + +subsection \Translation into normal form\ + +lemma N_pos[simp]: "N > 0" + using N_g1 neq0_conv by blast + +lemmas N_pos_simps[simp] = suc_mods[OF N_g1] l_suc[OF N_g1] minus_suc[OF N_pos] + +text \The one-fork process\ + +type_synonym fork_id = nat +type_synonym fork_state = nat + +definition fork_transitions:: "fork_id \ fork_state \ dining_event set" ("Tr\<^sub>f") + where "Tr\<^sub>f i s = (if s = 0 then {picks i i} \ {picks ((i+1) mod N) i} + else if s = 1 then {putsdown i i} + else if s = 2 then {putsdown ((i+1) mod N) i} + else {})" +declare Un_insert_right[simp del] Un_insert_left[simp del] + +lemma ev_fork_idx[simp]: "e \ Tr\<^sub>f i s \ fork e = i" + by (auto simp add:fork_transitions_def split:if_splits) + +definition fork_state_update:: "fork_id \ fork_state \ dining_event \ fork_state" ("Up\<^sub>f") + where "Up\<^sub>f i s e = ( if e = (picks i i) then 1 + else if e = (picks ((i+1) mod N) i) then 2 + else 0 )" + +definition FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m:: "fork_id \ fork_state \ dining_event process" + where "FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m i = P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr\<^sub>f i ,Up\<^sub>f i\" + +lemma FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec: "FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m i = (\ s. \ e \ (Tr\<^sub>f i s) \ FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m i (Up\<^sub>f i s e))" + using fix_eq[of "\ X. (\s. Mprefix (Tr\<^sub>f i s) (\e. X (Up\<^sub>f i s e)))"] FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def by simp + +lemma FORK_refines_FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m: "FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m i 0 \ FORK i" +proof(unfold FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, + induct rule:fix_ind_k_skip[where k=2 and f="\ x.(\s. Mprefix (Tr\<^sub>f i s) (\e. x (Up\<^sub>f i s e)))"]) + show "(1::nat) \ 2" by simp +next + show "adm (\a. a 0 \ FORK i)" by (simp add: cont_fun) +next + case base_k_steps:3 + show ?case (is "\j<2. (iterate j\?f\\) 0 \ FORK i") + proof - + have less_2:"\j. (j::nat) < 2 = (j = 0 \ j = 1)" by linarith + moreover have "(iterate 0\?f\\) 0 \ FORK i" by simp + moreover have "(iterate 1\?f\\) 0 \ FORK i" + by (subst FORK_rec) + (simp add: write0_def + fork_transitions_def + mprefix_Un_distr mono_det_ref mono_mprefix_ref) + ultimately show ?thesis by simp + qed +next + case step:(4 x) + then show ?case (is "(iterate 2\?f\x) 0 \ FORK i") + by (subst FORK_rec) + (simp add: write0_def numeral_2_eq_2 + fork_transitions_def fork_state_update_def + mprefix_Un_distr mono_det_ref mono_mprefix_ref) +qed + +lemma FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_refines_FORK: "FORK i \ FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m i 0" +proof(unfold FORK_def, + induct rule:fix_ind_k_skip[where k=1]) + show "(1::nat) \ 1" by simp +next + show "adm (\a. a \ FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m i 0)" by simp +next + case base_k_steps:3 + show ?case by simp +next + case step:(4 x) + then show ?case (is "iterate 1\?f\x \ FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m i 0") + apply (subst FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec) + apply (simp add: write0_def + fork_transitions_def fork_state_update_def + mprefix_Un_distr) + by (subst (1 2) FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec) + (simp add: fork_transitions_def fork_state_update_def + mprefix_Un_distr mono_det_ref mono_mprefix_ref) +qed + +lemma FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_is_FORK: "FORK i = FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m i 0" + using FORK_refines_FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_refines_FORK below_antisym by blast + + +text \The all-forks process in normal form\ + +type_synonym forks_state = "nat list" + +definition forks_transitions:: "nat \ forks_state \ dining_event set" ("Tr\<^sub>F") + where "Tr\<^sub>F n fs = (\if i (fs!i))" + +lemma forks_transitions_take: "Tr\<^sub>F n fs = Tr\<^sub>F n (take n fs)" + by (simp add:forks_transitions_def) + +definition forks_state_update:: "forks_state \ dining_event \ forks_state" ("Up\<^sub>F") + where "Up\<^sub>F fs e = (let i=(fork e) in fs[i:=(Up\<^sub>f i (fs!i) e)])" + +lemma forks_update_take: "take n (Up\<^sub>F fs e) = Up\<^sub>F (take n fs) e" + unfolding forks_state_update_def + by (metis nat_less_le nat_neq_iff nth_take order_refl take_update_cancel take_update_swap) + +definition FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m:: "nat \ forks_state \ dining_event process" + where "FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n = P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr\<^sub>F n ,Up\<^sub>F\" + +lemma FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec: "FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n = (\ fs. \ e \ (Tr\<^sub>F n fs) \ FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (Up\<^sub>F fs e))" + using fix_eq[of "\ X. (\fs. Mprefix (Tr\<^sub>F n fs) (\e. X (Up\<^sub>F fs e)))"] FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def by simp + +lemma FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_0: "FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m 0 fs = STOP" + by (subst FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add:forks_transitions_def Mprefix_STOP) + +lemma FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1_dir1: "length fs > 0 \ FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m 1 fs \ (FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m 0 (fs!0))" +proof(unfold FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, + induct arbitrary:fs rule:fix_ind_k[where k=1 + and f="\ x. (\fs. Mprefix (Tr\<^sub>F 1 fs) (\e. x (Up\<^sub>F fs e)))"]) + case adm:1 + then show ?case by (simp add:cont_fun) +next + case base_k_steps:2 + then show ?case by simp +next + case step:(3 X) + hence "(\if i (fs ! i)) = Tr\<^sub>f 0 (fs ! 0)" by auto + with step show ?case + apply (subst FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add:forks_state_update_def forks_transitions_def) + apply (intro mono_mprefix_ref, safe) + by (metis ev_fork_idx step.prems list_update_nonempty nth_list_update_eq) +qed + +lemma FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1_dir2: "length fs > 0 \ (FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m 0 (fs!0)) \ FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m 1 fs" +proof(unfold FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, + induct arbitrary:fs rule:fix_ind_k[where k=1 + and f="\ x. (\s. Mprefix (Tr\<^sub>f 0 s) (\e. x (Up\<^sub>f 0 s e)))"]) + case adm:1 + then show ?case by (simp add:cont_fun) +next + case base_k_steps:2 + then show ?case by simp +next + case step:(3 X) + have "(\if i (fs ! i)) = Tr\<^sub>f 0 (fs ! 0)" by auto + with step show ?case + apply (subst FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add:forks_state_update_def forks_transitions_def) + apply (intro mono_mprefix_ref, safe) + by (metis ev_fork_idx step.prems list_update_nonempty nth_list_update_eq) +qed + +lemma FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1: "length fs > 0 \ (FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m 0 (fs!0)) = FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m 1 fs" + using FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1_dir1 FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1_dir2 below_antisym by blast + +lemma FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_unfold: +"0 < n \ length fs = Suc n \ + FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m (Suc n) fs = (FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (butlast fs)|||(FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (fs!n)))" +proof(rule below_antisym) + show "0 < n \ length fs = Suc n \ + FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m (Suc n) fs \ (FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (butlast fs)|||FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (fs!n))" + proof(subst FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, + induct arbitrary:fs + rule:fix_ind_k[where k=1 + and f="\ x. (\fs. Mprefix (Tr\<^sub>F (Suc n) fs) (\e. x (Up\<^sub>F fs e)))"]) + case adm:1 + then show ?case by (simp add:cont_fun) + next + case base_k_steps:2 + then show ?case by simp + next + case step:(3 X) + have indep:"\s\<^sub>1 s\<^sub>2. Tr\<^sub>F n s\<^sub>1 \ Tr\<^sub>f n s\<^sub>2 = {}" + by (auto simp add:forks_transitions_def fork_transitions_def) + from step show ?case + apply (auto simp add:indep dnorm_inter FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def) + apply (subst fix_eq, simp add:forks_transitions_def Un_commute lessThan_Suc nth_butlast) + proof(rule mono_mprefix_ref, safe, goal_cases) + case (1 e) + from 1(4) have a:"fork e = n" + by (auto simp add:fork_transitions_def split:if_splits) + show ?case + using 1(1)[rule_format, of "(Up\<^sub>F fs e)"] + apply (simp add: 1 a butlast_list_update forks_state_update_def) + by (metis 1(4) ev_fork_idx lessThan_iff less_not_refl) + next + case (2 e m) + hence a:"e \ Tr\<^sub>f n (fs ! n)" + using ev_fork_idx by fastforce + hence c:"Up\<^sub>F fs e ! n = fs ! n" + by (metis 2(4) ev_fork_idx forks_state_update_def nth_list_update_neq) + have d:"Up\<^sub>F (butlast fs) e = butlast (Up\<^sub>F fs e)" + apply(simp add:forks_state_update_def) + by (metis butlast_conv_take forks_state_update_def forks_update_take length_list_update) + from 2 a show ?case + using 2(1)[rule_format, of "(Up\<^sub>F fs e)"] c d forks_state_update_def by auto + qed + qed +next + have indep:"\s\<^sub>1 s\<^sub>2. Tr\<^sub>F n s\<^sub>1 \ Tr\<^sub>f n s\<^sub>2 = {}" + by (auto simp add:forks_transitions_def fork_transitions_def) + show "0 < n \ length fs = Suc n \ + (FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (butlast fs)|||FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (fs!n)) \ FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m (Suc n) fs" + apply (subst FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, auto simp add:indep dnorm_inter FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def) + proof(rule fix_ind[where + P="\a. 0 < n \ (\x. length x = Suc n \ a (butlast x, x ! n) \ FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m (Suc n) x)", + rule_format], simp_all, goal_cases) + case base:1 + then show ?case by (simp add:cont_fun) + next + case step:(2 X fs) + then show ?case + apply (unfold FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, subst fix_eq, simp add:forks_transitions_def + Un_commute lessThan_Suc nth_butlast) + proof(rule mono_mprefix_ref, safe, goal_cases) + case (1 e) + from 1(6) have a:"fork e = n" + by (auto simp add:fork_transitions_def split:if_splits) + show ?case + using 1(1)[rule_format, of "(Up\<^sub>F fs e)"] + apply (simp add: 1 a butlast_list_update forks_state_update_def) + using a ev_fork_idx by blast + next + case (2 e m) + have a:"Up\<^sub>F (butlast fs) e = butlast (Up\<^sub>F fs e)" + apply(simp add:forks_state_update_def) + by (metis butlast_conv_take forks_state_update_def forks_update_take length_list_update) + from 2 show ?case + using 2(1)[rule_format, of "(Up\<^sub>F fs e)"] a forks_state_update_def by auto + qed + qed +qed + +lemma ft: "0 < n \ FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (replicate n 0) = foldFORKs n" +proof (induct n, simp) + case (Suc n) + then show ?case + apply (auto simp add: FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_unfold FORK\<^sub>n\<^sub>o\<^sub>r\<^sub>m_is_FORK) + apply (metis Suc_le_D butlast_snoc replicate_Suc replicate_append_same) + by (metis FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1 One_nat_def leI length_replicate less_Suc0 nth_replicate replicate_Suc) +qed + +corollary FORKs_is_FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m: "FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m N (replicate N 0) = FORKs" + using ft N_pos by simp + +text \The one-philosopher process in normal form:\ + +type_synonym phil_id = nat +type_synonym phil_state = nat + +definition rphil_transitions:: "phil_id \ phil_state \ dining_event set" ("Tr\<^sub>r\<^sub>p") + where "Tr\<^sub>r\<^sub>p i s = ( if s = 0 then {picks i i} + else if s = 1 then {picks i (i-1)} + else if s = 2 then {putsdown i (i-1)} + else if s = 3 then {putsdown i i} + else {})" + +definition lphil0_transitions:: "phil_state \ dining_event set" ("Tr\<^sub>l\<^sub>p") + where "Tr\<^sub>l\<^sub>p s = ( if s = 0 then {picks 0 (N-1)} + else if s = 1 then {picks 0 0} + else if s = 2 then {putsdown 0 0} + else if s = 3 then {putsdown 0 (N-1)} + else {})" + +corollary rphil_phil: "e \ Tr\<^sub>r\<^sub>p i s \ phil e = i" + and lphil0_phil: "e \ Tr\<^sub>l\<^sub>p s \ phil e = 0" + by (simp_all add:rphil_transitions_def lphil0_transitions_def split:if_splits) + +definition rphil_state_update:: "fork_id \ fork_state \ dining_event \ fork_state" ("Up\<^sub>r\<^sub>p") + where "Up\<^sub>r\<^sub>p i s e = ( if e = (picks i i) then 1 + else if e = (picks i (i-1)) then 2 + else if e = (putsdown i (i-1)) then 3 + else 0 )" + +definition lphil0_state_update:: "fork_state \ dining_event \ fork_state" ("Up\<^sub>l\<^sub>p") + where "Up\<^sub>l\<^sub>p s e = ( if e = (picks 0 (N-1)) then 1 + else if e = (picks 0 0) then 2 + else if e = (putsdown 0 0) then 3 + else 0 )" + +definition RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m:: "fork_id \ fork_state \ dining_event process" + where "RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m i = P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr\<^sub>r\<^sub>p i,Up\<^sub>r\<^sub>p i\" + +definition LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m:: "fork_state \ dining_event process" + where "LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m = P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr\<^sub>l\<^sub>p,Up\<^sub>l\<^sub>p\" + +lemma RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec: "RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m i = (\ s. \ e \ (Tr\<^sub>r\<^sub>p i s) \ RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m i (Up\<^sub>r\<^sub>p i s e))" + using fix_eq[of "\ X. (\s. Mprefix (Tr\<^sub>r\<^sub>p i s) (\e. X (Up\<^sub>r\<^sub>p i s e)))"] RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def by simp + +lemma LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec: "LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m = (\ s. \ e \ (Tr\<^sub>l\<^sub>p s) \ LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m (Up\<^sub>l\<^sub>p s e))" + using fix_eq[of "\ X. (\s. Mprefix (Tr\<^sub>l\<^sub>p s) (\e. X (Up\<^sub>l\<^sub>p s e)))"] LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def by simp + +lemma RPHIL_refines_RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m: + assumes i_pos: "i > 0" + shows "RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m i 0 \ RPHIL i" +proof(unfold RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, + induct rule:fix_ind_k_skip[where k=4 + and f="\ x. (\s. Mprefix (Tr\<^sub>r\<^sub>p i s) (\e. x (Up\<^sub>r\<^sub>p i s e)))"]) + show "(1::nat) \ 4" by simp +next + show "adm (\a. a 0 \ RPHIL i)" by (simp add: cont_fun) +next + case base_k_steps:3 + show ?case (is "\j<4. (iterate j\?f\\) 0 \ RPHIL i") + proof - + have less_2:"\j. (j::nat) < 4 = (j = 0 \ j = 1 \ j = 2 \ j = 3)" by linarith + moreover have "(iterate 0\?f\\) 0 \ RPHIL i" by simp + moreover have "(iterate 1\?f\\) 0 \ RPHIL i" + by (subst RPHIL_rec) + (simp add: write0_def rphil_transitions_def mono_mprefix_ref) + moreover have "(iterate 2\?f\\) 0 \ RPHIL i" + by (subst RPHIL_rec) + (simp add: numeral_2_eq_2 write0_def rphil_transitions_def + rphil_state_update_def mono_mprefix_ref) + moreover have "(iterate 3\?f\\) 0 \ RPHIL i" + by (subst RPHIL_rec) + (simp add: numeral_3_eq_3 write0_def rphil_transitions_def + rphil_state_update_def mono_mprefix_ref minus_suc[OF i_pos]) + ultimately show ?thesis by simp + qed +next + case step:(4 x) + then show ?case (is "(iterate 4\?f\x) 0 \ RPHIL i") + apply (subst RPHIL_rec) + apply (simp add: write0_def numeral_4_eq_4 rphil_transitions_def rphil_state_update_def) + apply (rule mono_mprefix_ref, auto simp:minus_suc[OF i_pos])+ + using minus_suc[OF i_pos] by auto +qed + +lemma LPHIL0_refines_LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m: "LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m 0 \ LPHIL0" +proof(unfold LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, + induct rule:fix_ind_k_skip[where k=4 and f="\ x. (\s. Mprefix (Tr\<^sub>l\<^sub>p s) (\e. x (Up\<^sub>l\<^sub>p s e)))"]) + show "(1::nat) \ 4" by simp +next + show "adm (\a. a 0 \ LPHIL0)" by (simp add: cont_fun) +next + case base_k_steps:3 + show ?case (is "\j<4. (iterate j\?f\\) 0 \ LPHIL0") + proof - + have less_2:"\j. (j::nat) < 4 = (j = 0 \ j = 1 \ j = 2 \ j = 3)" by linarith + moreover have "(iterate 0\?f\\) 0 \ LPHIL0" by simp + moreover have "(iterate 1\?f\\) 0 \ LPHIL0" + by (subst LPHIL0_rec) + (simp add: write0_def lphil0_transitions_def mono_mprefix_ref) + moreover have "(iterate 2\?f\\) 0 \ LPHIL0" + by (subst LPHIL0_rec) + (simp add: numeral_2_eq_2 write0_def lphil0_transitions_def + lphil0_state_update_def mono_mprefix_ref) + moreover have "(iterate 3\?f\\) 0 \ LPHIL0" + by (subst LPHIL0_rec) + (simp add: numeral_3_eq_3 write0_def lphil0_transitions_def + lphil0_state_update_def mono_mprefix_ref) + ultimately show ?thesis by simp + qed +next + case step:(4 x) + then show ?case (is "(iterate 4\?f\x) 0 \ LPHIL0") + by (subst LPHIL0_rec) + (simp add: write0_def numeral_4_eq_4 lphil0_transitions_def + lphil0_state_update_def mono_mprefix_ref) +qed + +lemma RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_refines_RPHIL: + assumes i_pos: "i > 0" + shows "RPHIL i \ RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m i 0" +proof(unfold RPHIL_def, + induct rule:fix_ind_k_skip[where k=1]) + show "(1::nat) \ 1" by simp +next + show "adm (\a. a \ RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m i 0)" by simp +next + case base_k_steps:3 + show ?case by simp +next + case step:(4 x) + then show ?case + apply (subst RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add: write0_def rphil_transitions_def rphil_state_update_def) + apply (rule mono_mprefix_ref, simp) + apply (subst RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add: write0_def rphil_transitions_def rphil_state_update_def) + apply (rule mono_mprefix_ref, simp add:minus_suc[OF i_pos]) + apply (subst RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add: write0_def rphil_transitions_def rphil_state_update_def) + apply (rule mono_mprefix_ref, simp add:minus_suc[OF i_pos]) + apply (subst RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add: write0_def rphil_transitions_def rphil_state_update_def) + apply (rule mono_mprefix_ref, simp add:minus_suc[OF i_pos]) + using minus_suc[OF i_pos] by auto +qed + +lemma LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_refines_LPHIL0: "LPHIL0 \ LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m 0" +proof(unfold LPHIL0_def, + induct rule:fix_ind_k_skip[where k=1]) + show "(1::nat) \ 1" by simp +next + show "adm (\a. a \ LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m 0)" by simp +next + case base_k_steps:3 + show ?case by simp +next + case step:(4 x) + then show ?case (is "iterate 1\?f\x \ LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m 0") + apply (subst LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add: write0_def lphil0_transitions_def lphil0_state_update_def) + apply (rule mono_mprefix_ref, simp) + apply (subst LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add: write0_def lphil0_transitions_def lphil0_state_update_def) + apply (rule mono_mprefix_ref, simp) + apply (subst LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add: write0_def lphil0_transitions_def lphil0_state_update_def) + apply (rule mono_mprefix_ref, simp) + apply (subst LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add: write0_def lphil0_transitions_def lphil0_state_update_def) + apply (rule mono_mprefix_ref, simp) + done +qed + +lemma RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_is_RPHIL: "i > 0 \ RPHIL i = RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m i 0" + using RPHIL_refines_RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_refines_RPHIL below_antisym by blast + +lemma LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_is_LPHIL0: "LPHIL0 = LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m 0" + using LPHIL0_refines_LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_refines_LPHIL0 below_antisym by blast + +subsection \The normal form for the global philosopher network\ + +type_synonym phils_state = "nat list" + +definition phils_transitions:: "nat \ phils_state \ dining_event set" ("Tr\<^sub>P") + where "Tr\<^sub>P n ps = Tr\<^sub>l\<^sub>p (ps!0) \ (\i\{1..< n}. Tr\<^sub>r\<^sub>p i (ps!i))" + +corollary phils_phil: "0 < n \ e \ Tr\<^sub>P n s \ phil e < n" + by (auto simp add:phils_transitions_def lphil0_phil rphil_phil) + +lemma phils_transitions_take: "0 < n \ Tr\<^sub>P n ps = Tr\<^sub>P n (take n ps)" + by (auto simp add:phils_transitions_def) + +definition phils_state_update:: "phils_state \ dining_event \ phils_state" ("Up\<^sub>P") + where "Up\<^sub>P ps e = (let i=(phil e) in if i = 0 then ps[i:=(Up\<^sub>l\<^sub>p (ps!i) e)] + else ps[i:=(Up\<^sub>r\<^sub>p i (ps!i) e)])" + +lemma phils_update_take: "take n (Up\<^sub>P ps e) = Up\<^sub>P (take n ps) e" + by (cases e) (simp_all add: phils_state_update_def lphil0_state_update_def + rphil_state_update_def take_update_swap) + +definition PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m:: "nat \ phils_state \ dining_event process" + where "PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n = P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr\<^sub>P n,Up\<^sub>P\" + +lemma PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec: "PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n = (\ ps. \ e \ (Tr\<^sub>P n ps) \ PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (Up\<^sub>P ps e))" + using fix_eq[of "\ X. (\ps. Mprefix (Tr\<^sub>P n ps) (\e. X (Up\<^sub>P ps e)))"] PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def by simp + +lemma PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1_dir1: "length ps > 0 \ PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m 1 ps \ (LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m (ps!0))" +proof(unfold PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, + induct arbitrary:ps + rule:fix_ind_k[where k=1 + and f="\ x. (\ps. Mprefix (Tr\<^sub>P 1 ps) (\e. x (Up\<^sub>P ps e)))"]) + case adm:1 + then show ?case by (simp add:cont_fun) +next + case base_k_steps:2 + then show ?case by simp +next + case step:(3 X) + then show ?case + apply (subst LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add:phils_state_update_def phils_transitions_def) + proof (intro mono_mprefix_ref, safe, goal_cases) + case (1 e) + with 1(2) show ?case + using 1(1)[rule_format, of "ps[0 := Up\<^sub>l\<^sub>p (ps ! 0) e]"] + by (simp add:lphil0_transitions_def split:if_splits) + qed +qed + +lemma PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1_dir2: "length ps > 0 \ (LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m (ps!0)) \ PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m 1 ps" +proof(unfold LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, + induct arbitrary:ps rule:fix_ind_k[where k=1 + and f="\ x. (\s. Mprefix (Tr\<^sub>l\<^sub>p s) (\e. x (Up\<^sub>l\<^sub>p s e)))"]) + case adm:1 + then show ?case by (simp add:cont_fun) +next + case base_k_steps:2 + then show ?case by simp +next + case step:(3 X) + then show ?case + apply (subst PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_rec, simp add:phils_state_update_def phils_transitions_def) + proof (intro mono_mprefix_ref, safe, goal_cases) + case (1 e) + with 1(2) show ?case + using 1(1)[rule_format, of "ps[0 := Up\<^sub>l\<^sub>p (ps ! 0) e]"] + by (simp add:lphil0_transitions_def split:if_splits) + qed +qed + +lemma PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1: "length ps > 0 \ PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m 1 ps = (LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m (ps!0))" + using PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1_dir1 PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1_dir2 below_antisym by blast + +lemma PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_unfold: + assumes n_pos:"0 < n" + shows "length ps = Suc n \ + PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m (Suc n) ps = (PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (butlast ps)|||(RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (ps!n)))" +proof(rule below_antisym) + show "length ps = Suc n \ PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m (Suc n) ps \ (PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (butlast ps)|||RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (ps!n))" + proof(subst PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, + induct arbitrary:ps + rule:fix_ind_k[where k=1 + and f="\ x. (\ps. Mprefix (Tr\<^sub>P (Suc n) ps) (\e. x (Up\<^sub>P ps e)))"]) + case adm:1 + then show ?case by (simp add:cont_fun) + next + case base_k_steps:2 + then show ?case by simp + next + case step:(3 X) + have indep:"\s\<^sub>1 s\<^sub>2. Tr\<^sub>P n s\<^sub>1 \ Tr\<^sub>r\<^sub>p n s\<^sub>2 = {}" + using phils_phil rphil_phil n_pos by blast + from step have tra:"(Tr\<^sub>P (Suc n) ps) =(Tr\<^sub>P n (butlast ps) \ Tr\<^sub>r\<^sub>p n (ps ! n))" + by (auto simp add:n_pos phils_transitions_def nth_butlast Suc_leI + atLeastLessThanSuc Un_commute Un_assoc) + from step show ?case + apply (auto simp add:indep dnorm_inter PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def) + apply (subst fix_eq, auto simp add:tra) + proof(rule mono_mprefix_ref, safe, goal_cases) + case (1 e) + hence c:"Up\<^sub>P ps e ! n = ps ! n" + using 1(3) phils_phil phils_state_update_def step n_pos + by (cases "phil e", auto) (metis exists_least_iff nth_list_update_neq) + have d:"Up\<^sub>P (butlast ps) e = butlast (Up\<^sub>P ps e)" + by (cases "phil e", auto simp add:phils_state_update_def butlast_list_update + lphil0_state_update_def rphil_state_update_def) + have e:"length (Up\<^sub>P ps e) = Suc n" + by (metis (full_types) step(2) length_list_update phils_state_update_def) + from 1 show ?case + using 1(1)[rule_format, of "(Up\<^sub>P ps e)"] c d e by auto + next + case (2 e) + have e:"length (Up\<^sub>P ps e) = Suc n" + by (metis (full_types) step(2) length_list_update phils_state_update_def) + from 2 show ?case + using 2(1)[rule_format, of "(Up\<^sub>P ps e)", OF e] n_pos + apply(auto simp add: butlast_list_update rphil_phil phils_state_update_def) + by (meson disjoint_iff_not_equal indep) + qed + qed +next + have indep:"\s\<^sub>1 s\<^sub>2. Tr\<^sub>P n s\<^sub>1 \ Tr\<^sub>r\<^sub>p n s\<^sub>2 = {}" + using phils_phil rphil_phil using n_pos by blast + + show "length ps = Suc n \ (PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (butlast ps)|||RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (ps!n)) \ PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m (Suc n) ps" + apply (subst PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def, auto simp add:indep dnorm_inter RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def) + proof(rule fix_ind[where + P="\a. \x. length x = Suc n \ a (butlast x, x ! n) \ PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m (Suc n) x", rule_format], + simp_all, goal_cases base step) + case base + then show ?case by (simp add:cont_fun) + next + case (step X ps) + hence tra:"(Tr\<^sub>P (Suc n) ps) =(Tr\<^sub>P n (butlast ps) \ Tr\<^sub>r\<^sub>p n (ps ! n))" + by (auto simp add:n_pos phils_transitions_def nth_butlast + Suc_leI atLeastLessThanSuc Un_commute Un_assoc) + from step show ?case + apply (auto simp add:indep dnorm_inter PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def) + apply (subst fix_eq, auto simp add:tra) + proof(rule mono_mprefix_ref, safe, goal_cases) + case (1 e) + hence c:"Up\<^sub>P ps e ! n = ps ! n" + using 1(3) phils_phil phils_state_update_def step n_pos + by (cases "phil e", auto) (metis exists_least_iff nth_list_update_neq) + have d:"Up\<^sub>P (butlast ps) e = butlast (Up\<^sub>P ps e)" + by (cases "phil e", auto simp add:phils_state_update_def butlast_list_update + lphil0_state_update_def rphil_state_update_def) + have e:"length (Up\<^sub>P ps e) = Suc n" + by (metis (full_types) step(3) length_list_update phils_state_update_def) + from 1 show ?case + using 1(2)[rule_format, of "(Up\<^sub>P ps e)", OF e] c d by auto + next + case (2 e) + have e:"length (Up\<^sub>P ps e) = Suc n" + by (metis (full_types) 2(3) length_list_update phils_state_update_def) + from 2 show ?case + using 2(2)[rule_format, of "(Up\<^sub>P ps e)", OF e] n_pos + apply(auto simp add: butlast_list_update rphil_phil phils_state_update_def) + by (meson disjoint_iff_not_equal indep) + qed + qed +qed + +lemma pt: "0 < n \ PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (replicate n 0) = foldPHILs n" +proof (induct n, simp) + case (Suc n) + then show ?case + apply (auto simp add: PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_unfold LPHIL0\<^sub>n\<^sub>o\<^sub>r\<^sub>m_is_LPHIL0) + apply (metis Suc_le_eq butlast.simps(2) butlast_snoc RPHIL\<^sub>n\<^sub>o\<^sub>r\<^sub>m_is_RPHIL + nat_neq_iff replicate_append_same replicate_empty) + by (metis One_nat_def leI length_replicate less_Suc0 PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_1 nth_Cons_0 replicate_Suc) +qed + +corollary PHILs_is_PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m: "PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m N (replicate N 0) = PHILs" + using pt N_pos by simp + +subsection \The complete process system under normal form\ + +definition dining_transitions:: "nat \ phils_state \ forks_state \ dining_event set" ("Tr\<^sub>D") + where "Tr\<^sub>D n = (\(ps,fs). (Tr\<^sub>P n ps) \ (Tr\<^sub>F n fs))" + +definition dining_state_update:: + "phils_state \ forks_state \ dining_event \ phils_state \ forks_state" ("Up\<^sub>D") + where "Up\<^sub>D = (\(ps,fs) e. (Up\<^sub>P ps e, Up\<^sub>F fs e))" + +definition DINING\<^sub>n\<^sub>o\<^sub>r\<^sub>m:: "nat \ phils_state \ forks_state \ dining_event process" + where "DINING\<^sub>n\<^sub>o\<^sub>r\<^sub>m n = P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr\<^sub>D n, Up\<^sub>D\" + +lemma ltsDining_rec: "DINING\<^sub>n\<^sub>o\<^sub>r\<^sub>m n = (\ s. \ e \ (Tr\<^sub>D n s) \ DINING\<^sub>n\<^sub>o\<^sub>r\<^sub>m n (Up\<^sub>D s e))" + using fix_eq[of "\ X. (\s. Mprefix (Tr\<^sub>D n s) (\e. X (Up\<^sub>D s e)))"] DINING\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def by simp + +lemma DINING_is_DINING\<^sub>n\<^sub>o\<^sub>r\<^sub>m: "DINING = DINING\<^sub>n\<^sub>o\<^sub>r\<^sub>m N (replicate N 0, replicate N 0)" +proof - + have "DINING\<^sub>n\<^sub>o\<^sub>r\<^sub>m N (replicate N 0, replicate N 0) = + (PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m N (replicate N 0) || FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m N (replicate N 0))" + unfolding DINING\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def dining_transitions_def + dining_state_update_def dnorm_par by simp + thus ?thesis + using PHILs_is_PHILs\<^sub>n\<^sub>o\<^sub>r\<^sub>m FORKs_is_FORKs\<^sub>n\<^sub>o\<^sub>r\<^sub>m DINING_def + by (simp add: par_comm) +qed + +subsection \And finally: Philosophers may dine ! Always !\ + +corollary lphil_states:"Up\<^sub>l\<^sub>p r e = 0 \ Up\<^sub>l\<^sub>p r e = 1 \ Up\<^sub>l\<^sub>p r e = 2 \ Up\<^sub>l\<^sub>p r e = 3" + and rphil_states:"Up\<^sub>r\<^sub>p i r e = 0 \ Up\<^sub>r\<^sub>p i r e = 1 \ Up\<^sub>r\<^sub>p i r e = 2 \ Up\<^sub>r\<^sub>p i r e = 3" + unfolding lphil0_state_update_def rphil_state_update_def by auto + +lemma dining_events: +"e \ Tr\<^sub>D N s \ + (\i\{1.. e = picks i (i-1) \ e = putsdown i i \ e = putsdown i (i-1)) + \ (e = picks 0 0 \ e = picks 0 (N-1) \ e = putsdown 0 0 \ e = putsdown 0 (N-1))" + by (auto simp add:dining_transitions_def phils_transitions_def rphil_transitions_def + lphil0_transitions_def split:prod.splits if_splits) + +definition "inv_dining ps fs \ + (\i. Suc i < N \ ((fs!(Suc i) = 1) \ ps!Suc i \ 0)) \ (fs!(N-1) = 2 \ ps!0 \ 0) + \ (\i < N - 1. fs!i = 2 \ ps!Suc i = 2) \ (fs!0 = 1 \ ps!0 = 2) + \ (\i < N. fs!i = 0 \ fs!i = 1 \ fs!i = 2) + \ (\i < N. ps!i = 0 \ ps!i = 1 \ ps!i = 2 \ ps!i = 3) + \ length fs = N \ length ps = N" + +lemma inv_DINING: "s \ \ (Tr\<^sub>D N) Up\<^sub>D (replicate N 0, replicate N 0) \ inv_dining (fst s) (snd s)" +proof(induct rule:\.induct) + case rbase + show ?case + unfolding inv_dining_def + by (simp add:dining_transitions_def phils_transitions_def forks_transitions_def + lphil0_transitions_def rphil_transitions_def fork_transitions_def) +next + case (rstep s e) + from rstep(2,3) show ?case + apply(auto simp add:dining_transitions_def phils_transitions_def forks_transitions_def + lphil0_transitions_def rphil_transitions_def fork_transitions_def + lphil0_state_update_def rphil_state_update_def fork_state_update_def + dining_state_update_def phils_state_update_def forks_state_update_def + split:if_splits prod.split) + unfolding inv_dining_def + proof(goal_cases) + case (1 ps fs) + then show ?case + by (simp add:nth_list_update) force + next + case (2 ps fs) + then show ?case + by (auto simp add:nth_list_update) + next + case (3 ps fs) + then show ?case + using N_g1 by linarith + next + case (4 ps fs) + then show ?case + by (simp add:nth_list_update) force + next + case (5 ps fs) + then show ?case + using N_g1 by linarith + next + case (6 ps fs) + then show ?case + by (auto simp add:nth_list_update) + next + case (7 ps fs i) + then show ?case + apply (simp add:nth_list_update, intro impI conjI, simp_all) + by auto[1] (metis N_pos Suc_pred less_antisym, metis zero_neq_numeral) + next + case (8 ps fs i) + then show ?case + apply (simp add:nth_list_update, intro impI conjI allI, simp_all) + by (metis "8"(1) zero_neq_one)+ + next + case (9 ps fs i) + then show ?case + apply (simp add:nth_list_update, intro impI conjI allI, simp_all) + by (metis N_pos Suc_pred less_antisym) (metis n_not_Suc_n numeral_2_eq_2) + next + case (10 ps fs i) + then show ?case + apply (simp add:nth_list_update, intro impI conjI allI, simp_all) + by (metis "10"(1) "10"(5) One_nat_def n_not_Suc_n numeral_2_eq_2)+ + qed +qed + +lemma inv_implies_DF:"inv_dining ps fs \ Tr\<^sub>D N (ps, fs) \ {}" + unfolding inv_dining_def + apply(simp add:dining_transitions_def phils_transitions_def forks_transitions_def + lphil0_transitions_def + split: if_splits prod.splits) +proof(elim conjE, intro conjI impI, goal_cases) + case 1 + hence "putsdown 0 (N - Suc 0) \ (\if i (fs ! i))" + by (auto simp add:fork_transitions_def) + then show ?case + by blast +next + case 2 + hence "putsdown 0 0 \ (\if i (fs ! i))" + by (auto simp add:fork_transitions_def) + then show ?case + by (simp add:fork_transitions_def) force +next + case 3 + hence a:"fs!0 = 0 \ picks 0 0 \ (\if i (fs ! i))" + by (auto simp add:fork_transitions_def) + from 3 have b1:"ps!1 = 2 \ putsdown 1 0 \ (\x\{Suc 0..r\<^sub>p x (ps ! x))" + using N_g1 by (auto simp add:rphil_transitions_def) + from 3 have b2:"fs!0 = 2 \ putsdown 1 0 \ Tr\<^sub>f 0 (fs ! 0)" + using N_g1 by (auto simp add:fork_transitions_def) fastforce + from 3 have c:"fs!0 \ 0 \ ps!1 = 2" + by (metis N_pos N_pos_simps(3) One_nat_def diff_is_0_eq neq0_conv) + from 3 have d:"fs!0 \ 0 \ fs!0 = 2" + using N_pos by meson + then show ?case + apply(cases "fs!0 = 0") + using a apply (simp add: fork_transitions_def Un_insert_left) + using b1[OF c] b2[OF d] N_pos by blast +next + case 4 + then show ?case + using 4(5)[rule_format, of 0, OF N_pos] apply(elim disjE) + proof(goal_cases) + case 41:1 (* fs!0 = 0 *) + then show ?case + using 4(5)[rule_format, of 1, OF N_g1] apply(elim disjE) + proof(goal_cases) + case 411:1 (* fs!1 = 0 *) + from 411 have a0: "ps!1 = 0" + by (metis N_g1 One_nat_def neq0_conv) + from 411 have a1: "picks 1 1 \ (\if i (fs ! i))" + apply (auto simp add:fork_transitions_def) + by (metis (mono_tags, lifting) N_g1 Int_Collect One_nat_def lessThan_iff) + from 411 have a2: "ps!1 = 0 \ picks 1 1 \ (\i\{Suc 0..r\<^sub>p i (ps ! i))" + apply (auto simp add:rphil_transitions_def) + using N_g1 by linarith + from 411 show ?case + using a0 a1 a2 by blast + next + case 412:2 (* fs!1 = 1 *) + hence "ps!1 = 1 \ ps!1 = 3" + by (metis N_g1 One_nat_def less_numeral_extra(3) zero_less_diff) + with 412 show ?case + proof(elim disjE, goal_cases) + case 4121:1 (* ps!1 = 1 *) + from 4121 have b1: "picks 1 0 \ (\if i (fs ! i))" + apply (auto simp add:fork_transitions_def) + by (metis (full_types) Int_Collect N_g1 N_pos One_nat_def lessThan_iff mod_less) + from 4121 have b2: "picks 1 0 \ (\i\{Suc 0..r\<^sub>p i (ps ! i))" + apply (auto simp add:rphil_transitions_def) + using N_g1 by linarith + from 4121 show ?case + using b1 b2 by blast + next + case 4122:2 (* ps!1 = 3 *) + from 4122 have b3: "putsdown 1 1 \ (\if i (fs ! i))" + apply (auto simp add:fork_transitions_def) + using N_g1 by linarith + from 4122 have b4: "putsdown 1 1 \ (\i\{Suc 0..r\<^sub>p i (ps ! i))" + apply (auto simp add:rphil_transitions_def) + using N_g1 by linarith + then show ?case + using b3 b4 by blast + qed + next + case 413:3 (* fs!1 = 2 *) + then show ?case + proof(cases "N = 2") + case True + with 413 show ?thesis by simp + next + case False + from False 413 have c0: "ps!2 = 2" + by (metis N_g1 Suc_1 Suc_diff_1 nat_neq_iff not_gr0 zero_less_diff) + from False 413 have c1: "putsdown 2 1 \ (\if i (fs ! i))" + apply (auto simp add:fork_transitions_def) + using N_g1 apply linarith + using N_g1 by auto + from False 413 have c2: "ps!2 = 2 \ putsdown 2 1 \ (\i\{Suc 0..r\<^sub>p i (ps ! i))" + apply (auto simp add:rphil_transitions_def) + using N_g1 by linarith + from 413 False show ?thesis + using c0 c1 c2 by blast + qed + qed + next + case 42:2 (* fs!0 = 1 *) + then show ?case by blast + next + case 43:3 (* fs!0 = 2*) + from 43 have d0: "ps!1 = 2" + by (metis One_nat_def gr0I) + from 43 have d1: "putsdown 1 0 \ (\if i (fs ! i))" + by (auto simp add:fork_transitions_def) + from 43 have d2: "ps!1 = 2 \ putsdown 1 0 \ (\i\{Suc 0..r\<^sub>p i (ps ! i))" + apply (auto simp add:rphil_transitions_def) + using N_g1 by linarith + from 43 show ?case + using d0 d1 d2 by blast + qed +next + case 5 + then show ?case + using 5(6)[rule_format, of 0] by simp +qed + +corollary DF_DINING: "deadlock_free_v2 DINING" + unfolding DINING_is_DINING\<^sub>n\<^sub>o\<^sub>r\<^sub>m DINING\<^sub>n\<^sub>o\<^sub>r\<^sub>m_def + using inv_DINING inv_implies_DF by (subst deadlock_free_dnorm) auto + +end + +end + + + + + + + + + + + + + + + + + + + + + + + + diff --git a/thys/CSP_RefTK/Fix_ind_ext.thy b/thys/CSP_RefTK/Fix_ind_ext.thy new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/Fix_ind_ext.thy @@ -0,0 +1,142 @@ +(*<*) +\\ ******************************************************************** + * Project : CSP-RefTK - A Refinement Toolkit for HOL-CSP + * Version : 1.0 + * + * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * + * This file : More Fixpoint and k-Induction Schemata + * + * Copyright (c) 2020 Université Paris-Saclay, 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. + ******************************************************************************\ +(*>*) + +chapter \ Advanced Induction Schemata \ + +theory Fix_ind_ext + +imports HOLCF + +begin + +section\k-fixpoint-induction\ + +lemma nat_k_induct: + fixes k::nat + assumes "\in\<^sub>0. (\i0+i)) \ P (n\<^sub>0+k)" + shows "P (n::nat)" +proof(induct rule:nat_less_induct) + case (1 n) + then show ?case + apply(cases "n < k") + using assms(1) apply blast + using assms(2)[rule_format, of "n-k"] by auto +qed + +thm fix_ind fix_ind2 + +lemma fix_ind_k: + fixes k::nat + assumes adm: "adm P" + assumes base_k_steps: "\if\\)" + assumes step: "\x. (\if\x)) \ P (iterate k\f\x)" + shows "P (fix\f)" + unfolding fix_def2 apply (rule admD [OF adm chain_iterate]) + apply(rule nat_k_induct[of k], simp add: base_k_steps) + using step by (subst (1 2) add.commute, unfold iterate_iterate[symmetric]) blast + +lemma nat_k_skip_induct: + fixes k::nat + assumes "k \ 1" and "\in\<^sub>0. P (n\<^sub>0) \ P (n\<^sub>0+k)" + shows "P (n::nat)" +proof(induct rule:nat_less_induct) + case (1 n) + then show ?case + apply(cases "n < k") + using assms(2) apply blast + using assms(3)[rule_format, of "n-k"] assms(1) by auto +qed + +lemma fix_ind_k_skip: + fixes k::nat + assumes k_1: "k \ 1" + assumes adm: "adm P" + assumes base_k_steps: "\if\\)" + assumes step: "\x. P x \ P (iterate k\f\x)" + shows "P (fix\f)" + unfolding fix_def2 apply (rule admD [OF adm chain_iterate]) + apply(rule nat_k_skip_induct[of k]) + using k_1 base_k_steps apply auto + using step by (subst add.commute, unfold iterate_iterate[symmetric]) blast + +thm parallel_fix_ind + +section\Parallel fixpoint-induction\ + +lemma parallel_fix_ind_inc: + assumes adm: "adm (\x. P (fst x) (snd x))" + assumes base_fst: "\y. P \ y" and base_snd: "\x. P x \" + assumes step: "\x y. P x y \ P (G\x) y \ P x (H\y) \ P (G\x) (H\y)" + shows "P (fix\G) (fix\H)" +proof - + from adm have adm': "adm (case_prod P)" + unfolding split_def . + have "P (iterate i\G\\) (iterate j\H\\)" for i j + proof(induct "i+j" arbitrary:i j rule:nat_less_induct) + case 1 + { fix i' j' + assume i:"i = Suc i'" and j:"j = Suc j'" + have "P (iterate i'\G\\) (iterate j'\H\\)" + and "P (iterate i'\G\\) (iterate j\H\\)" + and "P (iterate i\G\\) (iterate j'\H\\)" + using "1.hyps" add_strict_mono i j apply blast + using "1.hyps" i apply auto[1] + using "1.hyps" j by auto + hence ?case by (simp add: i j step) + } + thus ?case + apply(cases i, simp add:base_fst) + apply(cases j, simp add:base_snd) + by assumption + qed + then have "\i. case_prod P (iterate i\G\\, iterate i\H\\)" + by simp + then have "case_prod P (\i. (iterate i\G\\, iterate i\H\\))" + by - (rule admD [OF adm'], simp, assumption) + then have "P (\i. iterate i\G\\) (\i. iterate i\H\\)" + by (simp add: lub_Pair) + then show "P (fix\G) (fix\H)" + by (simp add: fix_def2) +qed + +end \ No newline at end of file diff --git a/thys/CSP_RefTK/Introduction.thy b/thys/CSP_RefTK/Introduction.thy new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/Introduction.thy @@ -0,0 +1,123 @@ +(*<*) +\\ ******************************************************************** + * Project : CSP-RefTK - A Refinement Toolkit for HOL-CSP + * Version : 1.0 + * + * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * + * This file : An Introduction + * + * Copyright (c) 2020 Université Paris-Saclay, 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. + ******************************************************************************\ +(*>*) + +chapter\Context\ +(*<*) +theory Introduction + imports HOLCF +begin +(*>*) + +section\Introduction\ + +text\ +Communicating Sequential Processes CSP is a language +to specify and verify patterns of interaction of concurrent systems. +Together with CCS and LOTOS, it belongs to the family of \<^emph>\process algebras\. +CSP's rich theory comprises denotational, operational and algebraic semantic facets +and has influenced programming languages such as Limbo, Crystal, Clojure and +most notably Golang @{cite "donovan2015go"}. CSP has been applied in +industry as a tool for specifying and verifying the concurrent aspects of hardware +systems, such as the T9000 transputer @{cite "Barret95"}. + +The theory of CSP, in particular the denotational Failure/Divergence Denotational Semantics, +has been initially proposed in the book by Tony Hoare @{cite "Hoare:1985:CSP:3921"}, but evolved +substantially since @{cite "BrookesHR84" and "brookes-roscoe85" and "roscoe:csp:1998"}. + +Verification of CSP properties has been centered around the notion of \<^emph>\process refinement orderings\, +most notably \_\\<^sub>F\<^sub>D_\ and \_\_\. The latter turns the denotational domain of CSP into a Scott cpo +@{cite "scott:cpo:1972"}, which yields semantics for the fixed point operator \\x. f(x)\ provided +that \f\ is continuous with respect to \_\_\. Since it is possible to express deadlock-freeness and +livelock-freeness as a refinement problem, the verification of properties has been reduced +traditionally to a model-checking problem for a finite set of events \A\. + +We are interested in verification techniques for arbitrary event sets \A\ or arbitrarily +parameterized processes. Such processes can be used to model dense-timed processes, processes +with dynamic thread creation, and processes with unbounded thread-local variables and buffers. +Events may even be higher-order objects such as functions or again processes, paving the way +for the modeling of re-programmable compute servers or dynamic distributed computing architectures. +However, this adds substantial complexity to the process theory: when it comes to study the +interplay of different denotational models, refinement-orderings, and side-conditions for +continuity, paper-and-pencil proofs easily reach their limits of precision. + +Several attempts have been undertaken to develop the formal theory of CSP in an interactive proof system, +mostly in Isabelle/HOL @{cite "Camilleri91" and "tej.ea:corrected:1997" and "IsobeRoggenbach2010"}. +This work is based on the most recent instance in this line, HOL-CSP 2.0, which has been published +as AFP submission @{cite "HOL-CSP-AFP"} and whose development is hosted at +\<^url>\https://gitlri.lri.fr/burkhart.wolff/hol-csp2.0\. + +The present AFP Module is an add-on on this work and develops some support for +\<^enum> advanced induction schemes (mutual fixed-point Induction, K-induction), +\<^enum> bridge-Lemmas between the classical refinement relations in the FD-semantics, + which allow for reduced refinement proof complexity in certain cases, and +\<^enum> a theory of explicit state normalisation which allows for proofs over certain + communicating networks of arbitrary size. + +\newpage +\ + + + +section\The Global Architecture of CSP\_RefTk\ +text\ +\begin{figure}[ht] + \centering + \includegraphics[width=0.60\textwidth]{figures/session_graph.pdf} + \caption{The overall architecture: HOLCF, HOL-CSP, and CSP\_RefTk} + \label{fig:fig1} +\end{figure} +\ + +text\The global architecture of CSP\_RefTk is shown in \autoref{fig:fig1}. +The entire package resides on: +\<^enum> \<^session>\HOL-Eisbach\ from the Isabelle/HOL distribution, +\<^enum> \<^session>\HOLCF\ from the Isabelle/HOL distribution, and +\<^enum> \<^session>\HOL-CSP\ 2.0 from the Isabelle Archive of Formal Proofs. + +\<^noindent> The theories \<^verbatim>\Assertion_ext\ and \<^verbatim>\Fixind_ext\ are extensions of the +corresponding theories in \<^session>\HOL-CSP\.\ + + +(*<*) +end +(*>*) diff --git a/thys/CSP_RefTK/Process_norm.thy b/thys/CSP_RefTK/Process_norm.thy new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/Process_norm.thy @@ -0,0 +1,241 @@ +(*<*) +\\ ******************************************************************** + * Project : CSP-RefTK - A Refinement Toolkit for HOL-CSP + * Version : 1.0 + * + * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * + * This file : A Normalization Theory + * + * Copyright (c) 2020 Université Paris-Saclay, 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. + ******************************************************************************\ +(*>*) + +chapter\ Normalisation of Deterministic CSP Processes \ + +theory Process_norm + +imports "Properties" "Fix_ind_ext" + +begin + +section\Deterministic normal-forms with explicit state\ + +abbreviation "P_dnorm Tr Up \ (\ X. (\ s. \ e \ (Tr s) \ X (Up s e)))" + +notation P_dnorm ("P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\_,_\" 60) + +lemma dnorm_cont[simp]: + fixes Tr::"'state::type \ 'event::type set" and Up::"'state \ 'event \ 'state" + shows "cont (\X. (\s. \ e \ (Tr s) \ X (Ur s e)))" (is "cont ?f") +proof - + have "cont (\X. ?f X s)" for s by (simp add:cont_fun) + then show ?thesis by simp +qed + +section\Interleaving product lemma\ + +lemma dnorm_inter: + fixes Tr\<^sub>1 ::"'state\<^sub>1::type \ 'event::type set" and Tr\<^sub>2::"'state\<^sub>2::type \ 'event set" + and Up\<^sub>1 ::"'state\<^sub>1 \ 'event \ 'state\<^sub>1" and Up\<^sub>2::"'state\<^sub>2 \ 'event \ 'state\<^sub>2" + defines P: "P \ P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr\<^sub>1,Up\<^sub>1\" (is "P \ fix\(\ X. ?P X)") + defines Q: "Q \ P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr\<^sub>2,Up\<^sub>2\" (is "Q \ fix\(\ X. ?Q X)") + + assumes indep: \\s\<^sub>1 s\<^sub>2. Tr\<^sub>1 s\<^sub>1 \ Tr\<^sub>2 s\<^sub>2 = {}\ + + defines Tr: "Tr \ (\(s\<^sub>1,s\<^sub>2). Tr\<^sub>1 s\<^sub>1 \ Tr\<^sub>2 s\<^sub>2)" + defines Up: "Up \ (\(s\<^sub>1,s\<^sub>2) e. if e \ Tr\<^sub>1 s\<^sub>1 then (Up\<^sub>1 s\<^sub>1 e,s\<^sub>2) + else if e \ Tr\<^sub>2 s\<^sub>2 then (s\<^sub>1, Up\<^sub>2 s\<^sub>2 e) + else (s\<^sub>1,s\<^sub>2))" + defines S: "S \ P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr,Up\" (is "S \ fix\(\ X. ?S X)") + + shows "(P s\<^sub>1 ||| Q s\<^sub>2) = S (s\<^sub>1,s\<^sub>2)" +proof - + have P_rec: "P = ?P P" using fix_eq[of "(\ X. ?P X)"] P by simp + have Q_rec: "Q = ?Q Q" using fix_eq[of "(\ X. ?Q X)"] Q by simp + have S_rec: "S = ?S S" using fix_eq[of "(\ X. ?S X)"] S by simp + have dir1: "\ s\<^sub>1 s\<^sub>2. (P s\<^sub>1 ||| Q s\<^sub>2) \ S (s\<^sub>1, s\<^sub>2)" + proof(subst P, subst Q, + induct rule:parallel_fix_ind_inc[of "\x y. \ s\<^sub>1 s\<^sub>2. (x s\<^sub>1 ||| y s\<^sub>2) \ S (s\<^sub>1,s\<^sub>2)"]) + case adm:1 + then show ?case + by (intro adm_all adm_below) (simp_all add: cont2cont_fun) + next + case base_fst:(2 y) + then show ?case by (metis Inter_commute app_strict minimal par_Int_bot) + next + case base_snd:(3 x) + then show ?case by (simp add: par_Int_bot) + next + case step:(4 x y) + then show ?case (is "\ s\<^sub>1 s\<^sub>2. ?C s\<^sub>1 s\<^sub>2") + proof(intro allI) + fix s\<^sub>1 s\<^sub>2 + show "?C s\<^sub>1 s\<^sub>2" + apply(simp) + apply(subst mprefix_Par_Int[where C="{}", simplified]) + apply(subst S_rec, simp add: Tr Up mprefix_Un_distr) + apply(intro mono_det_ref mono_mprefix_ref) + using step(3)[simplified] indep apply simp + using step(2)[simplified] indep by fastforce + qed + qed + have dir2: "\ s\<^sub>1 s\<^sub>2. S (s\<^sub>1, s\<^sub>2) \ (P s\<^sub>1 ||| Q s\<^sub>2)" + proof(subst S, induct rule:fix_ind_k[of "\x. \ s\<^sub>1 s\<^sub>2. x (s\<^sub>1,s\<^sub>2) \ (P s\<^sub>1 ||| Q s\<^sub>2)" 1]) + case adm:1 + show ?case by (intro adm_all adm_below) (simp_all add: cont_fun) + next + case base_k_steps:2 + then show ?case by simp + next + case step:(3 x) + then show ?case (is "\ s\<^sub>1 s\<^sub>2. ?C s\<^sub>1 s\<^sub>2") + proof(intro allI) + fix s\<^sub>1 s\<^sub>2 + have P_rec_sym:"Mprefix (Tr\<^sub>1 s\<^sub>1) (\e. P (Up\<^sub>1 s\<^sub>1 e)) = P s\<^sub>1" using P_rec by metis + have Q_rec_sym:"Mprefix (Tr\<^sub>2 s\<^sub>2) (\e. Q (Up\<^sub>2 s\<^sub>2 e)) = Q s\<^sub>2" using Q_rec by metis + show "?C s\<^sub>1 s\<^sub>2" + apply(simp add: Tr Up mprefix_Un_distr) + apply(subst P_rec, subst Q_rec, subst mprefix_Par_Int[where C="{}", simplified]) + apply(intro mono_det_ref mono_mprefix_ref) + apply(subst Q_rec_sym, simp add:step[simplified]) + apply(subst P_rec_sym) using step[simplified] indep by fastforce + qed + qed + from dir1 dir2 show ?thesis using below_antisym by blast +qed + +section \Synchronous product lemma\ +lemma dnorm_par: + fixes Tr\<^sub>1 ::"'state\<^sub>1::type \ 'event::type set" and Tr\<^sub>2::"'state\<^sub>2::type \ 'event set" + and Up\<^sub>1 ::"'state\<^sub>1 \ 'event \ 'state\<^sub>1" and Up\<^sub>2::"'state\<^sub>2 \ 'event \ 'state\<^sub>2" + defines P: "P \ P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr\<^sub>1,Up\<^sub>1\" (is "P \ fix\(\ X. ?P X)") + defines Q: "Q \ P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr\<^sub>2,Up\<^sub>2\" (is "Q \ fix\(\ X. ?Q X)") + + defines Tr: "Tr \ (\(s\<^sub>1,s\<^sub>2). Tr\<^sub>1 s\<^sub>1 \ Tr\<^sub>2 s\<^sub>2)" + defines Up: "Up \ (\(s\<^sub>1,s\<^sub>2) e. (Up\<^sub>1 s\<^sub>1 e, Up\<^sub>2 s\<^sub>2 e))" + defines S: "S \ P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr,Up\" (is "S \ fix\(\ X. ?S X)") + + shows "(P s\<^sub>1 || Q s\<^sub>2) = S (s\<^sub>1,s\<^sub>2)" +proof - + have P_rec: "P = ?P P" using fix_eq[of "(\ X. ?P X)"] P by simp + have Q_rec: "Q = ?Q Q" using fix_eq[of "(\ X. ?Q X)"] Q by simp + have S_rec: "S = ?S S" using fix_eq[of "(\ X. ?S X)"] S by simp + have dir1: "\ s\<^sub>1 s\<^sub>2. (P s\<^sub>1 || Q s\<^sub>2) \ S (s\<^sub>1, s\<^sub>2)" + proof(subst P, subst Q, + induct rule:parallel_fix_ind[of "\x y. \ s\<^sub>1 s\<^sub>2. (x s\<^sub>1 || y s\<^sub>2) \ S (s\<^sub>1,s\<^sub>2)"]) + case adm:1 + then show ?case + by (intro adm_all adm_below) (simp_all add: cont2cont_fun) + next + case base:2 + then show ?case by (simp add: par_Int_bot) + next + case step:(3 x y) + then show ?case (is "\ s\<^sub>1 s\<^sub>2. ?C s\<^sub>1 s\<^sub>2") + proof(intro allI) + fix s\<^sub>1 s\<^sub>2 + show "?C s\<^sub>1 s\<^sub>2" + apply(simp) + apply(subst mprefix_Par_distr[where C="UNIV", simplified]) + apply(subst S_rec, simp add: Tr Up mprefix_Un_distr) + by (simp add:step mono_mprefix_ref) + qed + qed + have dir2: "\ s\<^sub>1 s\<^sub>2. S (s\<^sub>1, s\<^sub>2) \ (P s\<^sub>1 || Q s\<^sub>2)" + proof(subst S, induct rule:fix_ind_k[of "\x. \ s\<^sub>1 s\<^sub>2. x (s\<^sub>1,s\<^sub>2) \ (P s\<^sub>1 || Q s\<^sub>2)" 1]) + case adm:1 + show ?case by (intro adm_all adm_below) (simp_all add: cont_fun) + next + case base_k_steps:2 + then show ?case by simp + next + case step:(3 x) + then show ?case (is "\ s\<^sub>1 s\<^sub>2. ?C s\<^sub>1 s\<^sub>2") + proof(intro allI) + fix s\<^sub>1 s\<^sub>2 + have P_rec_sym:"Mprefix (Tr\<^sub>1 s\<^sub>1) (\e. P (Up\<^sub>1 s\<^sub>1 e)) = P s\<^sub>1" using P_rec by metis + have Q_rec_sym:"Mprefix (Tr\<^sub>2 s\<^sub>2) (\e. Q (Up\<^sub>2 s\<^sub>2 e)) = Q s\<^sub>2" using Q_rec by metis + show "?C s\<^sub>1 s\<^sub>2" + apply(simp add: Tr Up) + apply(subst P_rec, subst Q_rec, subst mprefix_Par_distr[where C="UNIV", simplified]) + apply(rule mono_mprefix_ref) + using step by auto + qed + qed + from dir1 dir2 show ?thesis using below_antisym by blast +qed + +section\Consequences\ +\\reachable states from one starting state\ +inductive_set \ for Tr ::"'state::type \ 'event::type set" + and Up::"'state \ 'event \ 'state" + and s\<^sub>0::'state + where rbase: "s\<^sub>0 \ \ Tr Up s\<^sub>0" + | rstep: "s \ \ Tr Up s\<^sub>0 \ e \ Tr s \ Up s e \ \ Tr Up s\<^sub>0" + + + +\\Deadlock freeness\ +lemma deadlock_free_dnorm_ : + fixes Tr ::"'state::type \ 'event::type set" + and Up ::"'state \ 'event \ 'state" + and s\<^sub>0::'state + assumes non_reachable_sink: "\s \ \ Tr Up s\<^sub>0. Tr s \ {}" + defines P: "P \ P\<^sub>n\<^sub>o\<^sub>r\<^sub>m\Tr,Up\" (is "P \ fix\(\ X. ?P X)") + shows "s \ \ Tr Up s\<^sub>0 \ deadlock_free_v2 (P s)" +proof(unfold deadlock_free_v2_FD DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, induct arbitrary:s rule:fix_ind) + show "adm (\a. \x. x \ \ Tr Up s\<^sub>0 \ a \\<^sub>F\<^sub>D P x)" by (simp add: le_FD_adm monofun_def) +next + fix s :: "'state" + show "s \ \ Tr Up s\<^sub>0 \ \ \\<^sub>F\<^sub>D P s" by simp +next + fix s :: "'state" and x :: "'event process" + have P_rec: "P = ?P P" using fix_eq[of "(\ X. ?P X)"] P by simp + assume 1 : "\s. s \ \ Tr Up s\<^sub>0 \ x \\<^sub>F\<^sub>D P s" + and 2 : "s \ \ Tr Up s\<^sub>0 " + from 1 2 show "(\ x. (\xa\UNIV \ x) \ SKIP)\x \\<^sub>F\<^sub>D P s" + apply(simp add:failure_divergence_refine_def) + apply(subst P_rec, rule_tac trans_FD[simplified failure_divergence_refine_def, + rotated, OF Mprefix_refines_Mndet]) + apply(rule_tac CSP.mono_ndet_FD_left) + by (metis CSP.mono_ndet_FD_right rstep empty_not_UNIV mndet_distrib mono_mndet_FD + non_reachable_sink sup_top_left) +qed + + + +lemmas deadlock_free_dnorm = deadlock_free_dnorm_[rotated, OF rbase, rule_format] + +end + diff --git a/thys/CSP_RefTK/Properties.thy b/thys/CSP_RefTK/Properties.thy new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/Properties.thy @@ -0,0 +1,243 @@ +\\ ******************************************************************** + * Project : CSP-RefTK - A Refinement Toolkit for HOL-CSP + * Version : 1.0 + * + * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * + * This file : Theorems on DF and LF + * + * Copyright (c) 2020 Université Paris-Saclay, 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. + ******************************************************************************\ + +theory Properties + imports "Assertions_ext" +begin + +section \Deadlock Free\ + +thm deadlock_free_def + +lemma deadlock_free_implies_div_free: "deadlock_free P \ D P = {}" + by (simp add: deadlock_free_def div_free_DF failure_divergence_refine_def le_ref_def) + +lemma deadlock_free_implies_non_terminating: "deadlock_free (P::'a process) \ \s\T P. tickFree s" + unfolding deadlock_free_def apply(drule FD_F, drule le_F_T) unfolding trace_refine_def + using DF_all_tickfree_traces1[of "(UNIV::'a set)"] tickFree_def by fastforce + +definition deadlock_free_v2 :: "'a process \ bool" + where "deadlock_free_v2 P \ DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>F P" + +lemma deadlock_free_v2_is_right: + "deadlock_free_v2 (P::'a process) \ (\s\T P. tickFree s \ (s, UNIV::'a event set) \ F P)" +proof + assume a:"deadlock_free_v2 P" + have "tickFree s \ (s, UNIV) \ F (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV)" for s::"'a event list" + proof(induct s) + case Nil + then show ?case by (subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, simp add:F_mndet write0_def F_Mprefix F_ndet F_SKIP) + next + case (Cons a s) + then show ?case + by (subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, simp add:F_mndet write0_def F_Mprefix F_ndet F_SKIP) + qed + with a show "\s\T P. tickFree s \ (s, UNIV) \ F P" + using deadlock_free_v2_def failure_refine_def by blast +next + assume as1:"\s\T P. tickFree s \ (s, UNIV) \ F P" + have as2:"front_tickFree s \ (\aa \ UNIV. ev aa \ b) \ (s, b) \ F (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P (UNIV::'a set))" + for s b + proof(induct s) + case Nil + then show ?case + by (subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, auto simp add:F_mndet write0_def F_Mprefix F_ndet F_SKIP) + next + case (Cons hda tla) + then show ?case + proof(simp add:DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def fix_def) + define Y where "Y \ \i. iterate i\(\ x. (\xa\(UNIV::'a set) \ x) \ SKIP)\\" + assume a:"front_tickFree (hda # tla)" and b:"front_tickFree tla \ (tla, b) \ F (\i. Y i)" + and c:"\aa. ev aa \ b" + from Y_def have cc:"chain Y" by simp + from b have d:"front_tickFree tla \ \aa\UNIV. ev aa \ b \(tla, b) \ F (Y i)" for i + using F_LUB[OF cc] limproc_is_thelub[OF cc] by simp + from Y_def have e:"F(mndet UNIV (\x. Y i) \ SKIP) \ F (Y (Suc i))" for i by(simp) + from a have f:"tla \ [] \ hda \ tick" "front_tickFree tla" + apply (metis butlast.simps(2) butlast_snoc front_tickFree_charn + list.distinct(1) tickFree_Cons) + by (metis a append_Cons append_Nil front_tickFree_Nil front_tickFree_mono) + have g:"(hda#tla, b) \ F (Y (Suc i))" for i + using f c e[of i] d[of i] + by (auto simp: F_mndet write0_def F_Mprefix Y_def F_ndet F_SKIP) (metis event.exhaust)+ + have h:"(hda#tla, b) \ F (Y 0)" + using NF_ND cc g po_class.chainE proc_ord2a by blast + from a b c show "(hda#tla, b) \ F (\i. Y i)" + using F_LUB[OF cc] is_ub_thelub[OF cc] + by (metis D_LUB_2 cc g limproc_is_thelub po_class.chainE proc_ord2a process_charn) + qed + qed + show "deadlock_free_v2 P" + proof(auto simp add:deadlock_free_v2_def failure_refine_def) + fix s b + assume as3:"(s, b) \ F P" + hence a1:"s \ T P" "front_tickFree s" + using F_T apply blast + using as3 is_processT2 by blast + show "(s, b) \ F (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV)" + proof(cases "tickFree s") + case FT_True:True + hence a2:"(s, UNIV) \ F P" + using a1 as1 by blast + then show ?thesis + by (metis FT_True UNIV_I UNIV_eq_I a1(2) as2 as3 emptyE event.exhaust + is_processT6_S1 tickFree_implies_front_tickFree_single) + next + case FT_False: False + then show ?thesis + by (meson T_F_spec UNIV_witness a1(2) append_single_T_imp_tickFree + as2 emptyE is_processT5_S7) + qed + qed +qed + +lemma deadlock_free_v2_implies_div_free: "deadlock_free_v2 P \ D P = {}" + by (metis F_T append_single_T_imp_tickFree deadlock_free_v2_is_right ex_in_conv + nonTickFree_n_frontTickFree process_charn) + +corollary deadlock_free_v2_FD: "deadlock_free_v2 P = DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>F\<^sub>D P" + unfolding deadlock_free_v2_def + using deadlock_free_v2_implies_div_free FD_F F_D_FD deadlock_free_v2_def divergence_refine_def + by fastforce + +lemma all_events_refusal: + "(s, {tick} \ ev ` (events_of P)) \ F P \ (s, UNIV::'a event set) \ F P" +proof - + assume a1:"(s, {tick} \ ev ` events_of P) \ F P" + { assume "(s, UNIV) \ F P" + then obtain c where "c \ {tick} \ ev ` events_of P \ s @ [c] \ T P" + using is_processT5_S1[of s "{tick} \ ev ` events_of P" P + "UNIV - ({tick} \ ev ` events_of P)", simplified] F_T a1 by auto + hence False by (simp add:events_of_def, cases c) fastforce+ + } + with a1 show "(s, UNIV) \ F P" by blast +qed + +corollary deadlock_free_v2_is_right_wrt_events: + "deadlock_free_v2 (P::'a process) \ + (\s\T P. tickFree s \ (s, {tick} \ ev ` (events_of P)) \ F P)" + unfolding deadlock_free_v2_is_right using all_events_refusal apply auto + using is_processT4 by blast + +lemma deadlock_free_is_deadlock_free_v2: + "deadlock_free P \ deadlock_free_v2 P" + using DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF_refine_FD deadlock_free_def deadlock_free_v2_FD trans_FD by blast + + +section \Non-terminating Runs\ + +definition non_terminating :: "'a process \ bool" + where "non_terminating P \ RUN UNIV \\<^sub>T P" + +corollary non_terminating_refine_DF: "non_terminating P = DF UNIV \\<^sub>T P" + and non_terminating_refine_CHAOS: "non_terminating P = CHAOS UNIV \\<^sub>T P" + by (simp_all add: DF_all_tickfree_traces1 RUN_all_tickfree_traces1 CHAOS_all_tickfree_traces1 + non_terminating_def trace_refine_def) + +lemma non_terminating_is_right: "non_terminating (P::'a process) \ (\s\T P. tickFree s)" + apply (rule iffI) + by (auto simp add:non_terminating_def trace_refine_def tickFree_def RUN_all_tickfree_traces1)[1] + (auto simp add:non_terminating_def trace_refine_def RUN_all_tickfree_traces2) + +lemma nonterminating_implies_div_free: "non_terminating P \ D P = {}" + unfolding non_terminating_is_right + by (metis NT_ND equals0I front_tickFree_charn process_charn tickFree_Cons tickFree_append) + +lemma non_terminating_implies_F: "non_terminating P \ CHAOS UNIV \\<^sub>F P" + apply(auto simp add: non_terminating_is_right failure_refine_def) + using CHAOS_has_all_tickFree_failures F_T by blast + +corollary non_terminating_F: "non_terminating P = CHAOS UNIV \\<^sub>F P" + by (auto simp add:non_terminating_implies_F non_terminating_refine_CHAOS le_F_T) + +corollary non_terminating_FD: "non_terminating P = CHAOS UNIV \\<^sub>F\<^sub>D P" + unfolding non_terminating_F + using div_free_CHAOS nonterminating_implies_div_free FD_F F_D_FD divergence_refine_def + non_terminating_F by fastforce + + +section \Lifelock Freeness\ + +thm lifelock_free_def + +corollary lifelock_free_is_non_terminating: "lifelock_free P = non_terminating P" + unfolding lifelock_free_def non_terminating_FD by rule + +lemma div_free_divergence_refine: + "D P = {} \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>D P" + "D P = {} \ CHAOS UNIV \\<^sub>D P" + "D P = {} \ RUN UNIV \\<^sub>D P" + "D P = {} \ DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>D P" + "D P = {} \ DF UNIV \\<^sub>D P" + by (simp_all add: div_free_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P div_free_CHAOS div_free_RUN div_free_DF div_free_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P + divergence_refine_def) + +definition lifelock_free_v2 :: "'a process \ bool" + where "lifelock_free_v2 P \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>F\<^sub>D P" + +lemma div_free_is_lifelock_free_v2: "lifelock_free_v2 P \ D P = {}" + using CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures_Un FD_D F_D_FD div_free_divergence_refine(1) lifelock_free_v2_def + by blast + +lemma lifelock_free_is_lifelock_free_v2: "lifelock_free P \ lifelock_free_v2 P" + by (simp add: FD_D div_free_divergence_refine(2) div_free_is_lifelock_free_v2 lifelock_free_def) + +corollary deadlock_free_v2_is_lifelock_free_v2: "deadlock_free_v2 P \ lifelock_free_v2 P" + by (simp add: deadlock_free_v2_implies_div_free div_free_is_lifelock_free_v2) + + +section \New laws\ + +lemma non_terminating_seq: "non_terminating P \ (P `;` Q) = P" + apply(auto simp add: non_terminating_is_right Process_eq_spec D_seq F_seq F_T is_processT7) + using process_charn apply blast + using process_charn apply blast + using F_T is_processT5_S2a apply fastforce + using D_T front_tickFree_Nil by blast + +lemma non_terminating_inter: "non_terminating P \ lifelock_free_v2 Q \ non_terminating (P \C\ Q)" + apply(auto simp add: non_terminating_is_right div_free_is_lifelock_free_v2 T_sync) + apply (metis equals0D ftf_syn1 ftf_syn21 insertI1 tickFree_def) + apply (meson NT_ND is_processT7_S tickFree_append) + by (metis D_T empty_iff ftf_syn1 ftf_syn21 insertI1 tickFree_def) + + +end diff --git a/thys/CSP_RefTK/ROOT b/thys/CSP_RefTK/ROOT new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/ROOT @@ -0,0 +1,67 @@ +chapter AFP + +(***************************************************************************** + * + * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL + * Version : 2.0 + * + * Author : Burkhart Wolff, Safouan Taha & Lina Ye + * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) + * + * This file : The Process Refinement Toolkit CSP_RefTK + * + * Copyright (c) 2020 Université Paris-Saclay, 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:$ *) + +session "CSP_RefTK" (AFP) = HOLCF + + description \An Environment for Advanced Refinement Proofs based on Normalization. + Together with a proof of the generalized dining philosophers problem.\ + options [timeout = 1200] + sessions + "HOLCF-Library" + "HOL-CSP" + theories + Introduction + Assertions_ext Properties + Fix_ind_ext Process_norm + CopyBuffer_props + DiningPhilosophers + Conclusion + document_files + "root.tex" + "adb-long.bib" + "root.bib" + "figures/session_graph.pdf" + diff --git a/thys/CSP_RefTK/document/adb-long.bib b/thys/CSP_RefTK/document/adb-long.bib new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/document/adb-long.bib @@ -0,0 +1,80 @@ +% $Id: adb-long.bib 6518 2010-01-24 14:18:10Z brucker $ +@PREAMBLE{ {\providecommand{\ac}[1]{\textsc{#1}} } + # {\providecommand{\acs}[1]{\textsc{#1}} } + # {\providecommand{\acf}[1]{\textsc{#1}} } + # {\providecommand{\TAP}{T\kern-.1em\lower-.5ex\hbox{A}\kern-.1em P} } + # {\providecommand{\leanTAP}{\mbox{\sf lean\it\TAP}} } + # {\providecommand{\holz}{\textsc{hol-z}} } + # {\providecommand{\holocl}{\textsc{hol-ocl}} } + # {\providecommand{\isbn}{\textsc{isbn}} } + # {\providecommand{\Cpp}{C++} } + # {\providecommand{\Specsharp}{Spec\#} } + # {\providecommand{\doi}[1]{\href{http://dx.doi.org/#1}{doi: + {\urlstyle{rm}\nolinkurl{#1}}}}} } +@STRING{conf-tphols="\acs{tphols}" } +@STRING{iso = {International Organization for Standardization} } +@STRING{j-ar = "Journal of Automated Reasoning" } +@STRING{j-cacm = "Communications of the \acs{acm}" } +@STRING{j-acta-informatica = "Acta Informatica" } +@STRING{j-sosym = "Software and Systems Modeling" } +@STRING{j-sttt = "International Journal on Software Tools for Technology" } +@STRING{j-ist = "Information and Software Technology" } +@STRING{j-toplas= "\acs{acm} Transactions on Programming Languages and + Systems" } +@STRING{j-tosem = "\acs{acm} Transactions on Software Engineering and + Methodology" } +@STRING{j-eceasst="Electronic Communications of the \acs{easst}" } +@STRING{j-fac = "Formal Aspects of Computing" } +@STRING{j-ucs = "Journal of Universal Computer Science" } +@STRING{j-sl = "Journal of Symbolic Logic" } +@STRING{j-fp = "Journal of Functional Programming" } +@STRING{j-tkde = {\acs{ieee} Transaction on Knowledge and Data Engineering} } +@STRING{j-tse = {\acs{ieee} Transaction on Software Engineering} } +@STRING{j-entcs = {Electronic Notes in Theoretical Computer Science} } +@STRING{s-lnai = "Lecture Notes in Computer Science" } +@STRING{s-lncs = "Lecture Notes in Computer Science" } +@STRING{s-lnbip = "Lecture Notes in Business Information Processing" } +@String{j-computer = "Computer"} +@String{j-tissec = "\acs{acm} Transactions on Information and System Security"} +@STRING{omg = {Object Management Group} } +@STRING{j-ipl = {Information Processing Letters} } +@STRING{j-login = ";login: the USENIX Association newsletter" } + +@STRING{PROC = "Proceedings of the " } + + +% Publisher: +% ========== +@STRING{pub-awl = {Addison-Wesley Longman, Inc.} } +@STRING{pub-awl:adr={Reading, MA, \acs{usa}} } +@STRING{pub-springer={Springer-Verlag} } +@STRING{pub-springer:adr={Heidelberg} } +@STRING{pub-cup = {Cambridge University Press} } +@STRING{pub-cup:adr={New York, \acs{ny}, \acs{usa}} } +@STRING{pub-mit = {\acs{mit} Press} } +@STRING{pub-mit:adr={Cambridge, Massachusetts} } +@STRING{pub-springer-ny={Springer-Verlag} } +, +@STRING{pub-springer-netherlands={Springer Netherlands} } +@STRING{pub-springer-netherlands:adr={} } +@STRING{pub-springer-ny:adr={New York, \acs{ny}, \acs{usa}} } +@STRING{pub-springer-london={Springer-Verlag} } +@STRING{pub-springer-london:adr={London} } +@STRING{pub-ieee= {\acs{ieee} Computer Society} } +@STRING{pub-ieee:adr={Los Alamitos, \acs{ca}, \acs{usa}} } +@STRING{pub-prentice={Prentice Hall, Inc.} } +@STRING{pub-prentice:adr={Upper Saddle River, \acs{nj}, \acs{usa}} } +@STRING{pub-acm = {\acs{acm} Press} } +@STRING{pub-acm:adr={New York, \acs{ny} \acs{usa}} } +@STRING{pub-oxford={Oxford University Press, Inc.} } +@STRING{pub-oxford:adr={New York, \acs{ny}, \acs{usa}} } +@STRING{pub-kluwer={Kluwer Academic Publishers} } +@STRING{pub-kluwer:adr={Dordrecht} } +@STRING{pub-elsevier={Elsevier Science Publishers} } +@STRING{pub-elsevier:adr={Amsterdam} } +@STRING{pub-north={North-Holland Publishing Co.} } +@STRING{pub-north:adr={Nijmegen, The Netherlands} } +@STRING{pub-ios = {\textsc{ios} Press} } +@STRING{pub-ios:adr={Amsterdam, The Netherlands} } +@STRING{pub-heise={Heise Zeitschriften Verlag} } +@STRING{pub-heise:adr={Hannover, Germany} } diff --git a/thys/CSP_RefTK/document/figures/session_graph.pdf b/thys/CSP_RefTK/document/figures/session_graph.pdf new file mode 100644 index e69de29bb2d1d6434b8b29ae775ad8c2e48c5391..7de827cd2fc6a5d5483562abf78f4d5e59300405 GIT binary patch literal 11999 zc$}SjWmufcvMvNCXatu5LW0i>FoQ#IcY+TDhQVQQ*FbO!P6!eR1PC6S;1+_rySo!4 zz$I(#z1O;DoqL{p@Au>DudA!N-l^`YsUI(+ii8w9n1dUaad~B61eXf{0w9eoaRmi| zQWn-u2nV2)wUHA-0s%*wAb<)8TQetf05_kou%nX$!pH{K?E~wo1MU(hR2{^{sg5nX zJ{&Fdfv~@z* zIynMFfYUINmF%I3VHb2q%CZ@E$q9g#{|ylO7r;LQ00wY^csY6hhY0aP!H~bjXwS>tD}HS3u66(#hc5P9fS?C2t`zYo zk8XF7`yr}78Ie6b6ZxYDCGx~Xk|E_x<+9lwPl?@_nIzj^gouay8%bSyc)71~22em)bZ_=?UpM@|5k;KU1Jat`IOHmmY z!soIIX`d<(_&i@$h1_T&Vpb;Z_^i-yGha|=217n(OqM8~8_5l15tzSglgi?#hAS^v zB+K+ExKqP>RXZyL_TR^W)>=5oXrIY>@1AcxrF*a8MLOgcN+#kgZQh^OUtOUWT6|6N zD_y`$srDr4fNE1h<9)<+_}1N62K7b$XAg!HEGdTzZWT12@ez+AeUVN_UA}h^gFAk& z;qQa%ZLkc8P&JSmo{_LVv#c6O69pd1mTwf_kJG^v42jFRzYwVkI7pdp%FU{2V`)gDYu} zcc}>JQDDYC_BOQ?6%B0J%Rf!_lvga;QFNY@oOzQKMnA%Y>!AqwYP}l>NS^Bnp!r(y zL><*G^(JN8P;))MvX)NORY6jU<||LhPt*Q}98to;GQm|aTs~ke`RGuWP2D?XTW_BO z$ja*2rfZt|1~;w~Q<0*(*C1k!p-)M{nS$I>CKGqve9*!{M+cGD`$PwUAjbb8m@wFt zWSD5P?L+u%kW+lHTUCXEo4V0!WUD6_t29S&hWk^k5sosV2=34rOsWH`W*5GL&Loh& zO7>6&k~|db5$dtmvOx+xCZTxD26PVYOvLa;%yUMQ*9eC|cxD*s<*K$0Q^6QI&Yvc2 zU5qz-e$wfCUDBzi=p({#v5omXOjE%=gz}VxmE7-lmmc@NbSM8BPby~tzi|d1X{f|6 zjLbaYwZDCcmLKUI_WcbuhDyV%dc`cM@n<=BTq?D($1D*i3#0zd!>7|)6%D;ETxj** z=L&Uu6pwuJRlH}QpGcNSst5GOAB)cR8-8JuG)3frW)uzs##X>HPxX7S{QQD@(P<(( zm_)C7-UdASq$*e{ieU?dVXUREtlwE(g~V!2}3GXrL5BOJ?B04 z@EqRj@QVGQd_JLO+=AaS-QNB{J@8mji{vwv7;94%^_^TZQ2kRYsimn!X<{Fi_eoM) zTMhK%=ZRn&)oGiO6A}~7fq2n4QRQKEitR#orxV{pZv^TFYrx zrL31cjTkX<>)gAkWLf$0tC9j<-K1$Y+hS-1+TSd32JNqwICJjrmP8tTSop%cZ}wCke?g7gm)AxhI8#nMrZW8=p?BpJ9~GGCp@*u&mg^W2WXLQ zc1>d4keR&{=}Z}nw%6M1s-ZlE&r)B3_fc=ieg^tPIQ~j+;dJo4w72o7#TrwNCzV;C zr*!0X^Stge&@lx*TOWvd+>zc>Ge)|7u!4J+O=UZNyPa2JwGv5BAar6TWaa!$?TLi5 z3A>|S@`i(;_pfw8W_KE;ukK+SjgK!qWp&)r)eI9OC6-hS@1}+6-#YI-^nB?r-xFPk zb2b3IzH}VMG)ukl4SPy#fO?Ko96Y;GQNcckRmc{g>V(ORzRAV0FFc6VwZPE*BTi_7_2vGzEwnC;-2{;}=> zjwrVvdZW={86gke%4gQ9m-BBQi1JrD!n zMNPK`=rmZhlQgz~Mq+XqxipIjPoAFS>~?%qd~SlUmQUs8NM0e2NJ`6zR`Ad+4q;Vb zc$`hlo=P?oniO^(Qtd@`JW&zz39w!4uIhJ zr2G3!hG8lHILxEP@|N-E?PS!0CO*(k7YEyEYCw6SI5Y~(L+d+9BcL1mKeeGBrnVWUagOuXxf;kb2`1n;yvTbB~%6i9w^>!Eib$wndU z4bw&NN(GdEYYc(WPjjcdMpuY2{wy^8$iz)8YKm0HXAiM1z0CXWgULCX&<`X53PqnL z+nZXUw!;U~MX#1O{GieS-}()ux$z{Jv0K{7eA8mM;G*ACc2Yi*wezkUEn}`584I3( z?6aF_Doar`h~i^3Qko1V9eMoAH=ctsUR%Fgpbv8^xtOY%GM@d5i*miSIgBV*D3z!7~NCj$lZ@|0G~a) zmyQ^%v?vmMQ;+IJ`A!bKbd2Ge+H{DI9%B1eNt|A99qVo?PeZ&^OPLzXo#qiu)lKba z9qT%471t|Fx*?jbgXxyBIoskRFm_!ljKI)CUL!oeu|uJ z0gDZ}Ti-eyo=wq~%bFS*$MrTo5n#aUv5qBAyf-gAoapE#nc$vvexF6Mw1a~vWfSRf zQUZ@f4D$_j@uy~JizY0?5sk+v`SCO|ErHFXOoPwZLU;6WR#!n0m!o;j36(i1;0Rs~0bNK@c zmchkwsvt#!?Dn5U2CXA8e|kVKBph?bheZ$ zU)l_8ANMjSlY4S|HD%^&vqz#6Ep)XB*Zh!`#$TU#b)-0+;uz`siU8tMW75$Up1@+Y zd78qt?`4bJE649nXHcg8^?1fc^WxLIqKNn&^NX0bWz!^CYJxub*^>Y|wXub$AlY~# z<{!2><GcJ~QekFseBO4u}_{ztO%?TTaEy_oyyV4{tR|BgT0{l2%Z6c0i~L)OG9- zU%m>yEmjtWZjtQxoi#eo^4s-%at$UV-D%eVD;BaQ3j6OlaT(X!`s3m2?Rq{eve_@*KQNqYQ&j9rCd2@~Nca?0_5-6?qvZGSA3xs7+qh(%Z>E zC`;4`xCvAmMMN}2syZmDa*kF>1fZAJwEa_m5!7>NNN(ky#ZC=C)e#%99SiWH;$`c8Aq+WP}e06nZ&b*uhDuCTgVe z1yehYr#C`h@hazhdijPumS{#&_L+@?`-}4rePNpptZvShv@(&$ENf4?6VV|b8@@dz z;Vb;+>hq|IdgehIUw;%s|1wi#Ihqhe=1ZW{vlKiIa?|>;*eV-z#@gvmm5Pp=PhKoX z1dg%|z4dzC`$kWvOuWtuZ8I zX_;k-q>V@Sabe{+ztx*brspeWanx4Ra^zOBtAwYEC?hgy1=S&ZTNGZ(&!HF}jIi|8 zSuAD56fW9Xtk$;2Sc`6#7x@x5KR9b2j9nNu=^hW3qzhcO1ZO~O6C3p&3Fq7RQP)YH zuwy_oO=jIh8=1qfTR7?^Wi?GT9UGm&YO(494^uAZIbii`&X1~|H4qKEdviLJD2)XT zcvubRUU{acUQcvUDq z-hFZ{aSaP%Z@42JOB?z6v$6DU07hqa!DfAeVZ8!y1;C#6wNS8!{~r>oe=uo>Vz$yQHWZ zLEX*Z`00w{)Jb;trFm}q89s5NxW97gYD`D1(8PB^35#^jIb5H!gm=Ggq#;3?ZRZ3T zRs5FSQ+oWHfyC3RQ zK?zGFp7#;7dj{G%k-Vpgemttmpr~a*L~lv(jhjMI@*PObgv@g{D}VYRL}0Q03D&r7 zkgpCvQbj5$YX5Y!J+R6q2Ca;KoY~Qjnox_7_$Kp3m`zf#yZ6<)2OoAlS=6_WApshQ zfi68#C0`FS@is_(lC99Nn-bH2h)$#TY`hSUx1&^d(R&{E@uMWq2uuB%9F_`X+ZWHr z8ztYnr-qAW+&60;^4kVBtmL63*eD|d9O&umYUvIMIGT=eCgBfzZT;G9&F3wjGCO`$ z0vCUZh=}0Oj|_Y=p($C~D`SFD=d{A4BI0|i)#Q=>@hAz+YNJ7N=}_;)5^H}p|CHZh zdz9FZvx?JXud1I zk;!;KUR>5z zA5>y8Iw(J@)%-oSJziFCIlf+zak@^{acfX%=D=Dkbk%`F{N+1m@GF_4JsA}qZf?B7 zBH>55hwrwhJJCwl*ZWHI>R`2N!6T`x;#n%ZN9ttdYV=YziQdB2$qKdpIrcI!+WGqvuu-8g=QgzW=5l_+P1EYy& z+fNTqL?TjFiGs?+&&`b=CG7+a4RxX@g#Nl<{OJCsNX^bs+DMDp(12qj; zwa?vWQ$o?LoTT(6k8eReYZ z>^4Z=*gIJevYT`@6v*}LE=LT!jNmjxaH$7oozT27ZsOK+madCnP2g ze$l#Q;a{GoJ)E0EnlZ0hpNgrA7+W$l^NF)17Jgxsa1d?m{O zLm*<6R%xb=kH0k(Xdm8nqG+1oBHo4LB}2$_mT!4twA)+IF& zVdu@k{3aT~!w!m_Oy}t<0l7954L_5ig<{3}m}p8mX?56GWuEM;;@gUk`9AyWZqt(G z#Xl}(=;j=WyD*#~U%UC=NpHRBvK;SE>qEv4i(IMM*v^`^+gcp4gf>a`6r5&zu&=rd zq3cc8@Wvh#$TLmUq1nGI_@FA4Fonl)8RalQ)Q82>Xu6iSCeTS-E7YAZDk*GyPC33i z!+kd0&+yACJ-EG|9y3GW$HbLx{eJdpT$y&9O)7^;%S!?6`Khg;-*^7fFX#!zlv@? zfb$^MVw-h#UNYn08FgowReYLX&4^7s8+%<~cP{sAk-aI)vf4vL2*2d>&%cDlaVoiT zwsa0ys_Y#aCkItj1Uwf+j|`2Ij`iW@b9F+L3GN%?$f5$?Y+F-W0BI@UIdyPwJ!& zg=&8k#C8>E%?VrT7f(=}*40SP9$!@iIBF64nt#xAxCNiN?C;6e=Claxj@ZUc7bfG+k#NC_TOs(upgJwTaC5YqJQ~e#qUd|7 z*pJNeCJuzWaltQoBPf6B$t>Cw zO|Ydr$d-*|CrCZ4IYP%k3XddUeoWq~ygBjO2fD88zaeF@=tU$joCVGwb z%LtY7P}O>rl(fYbj8`Lfm5os`1MX zMY6Nre^qx}Fi7xM5BrI@X>y8{#rP$YuuI3>93}4_TB{g!e11cCCRn z*f*4{*x*4o=Ro6r_Y_wuczvp);kqF!RkwZgt;L0`jxTT0fw*E@S4AVuZO9dcFjNzI z^nLP;2(91#*UCFl8-4s#TKs_|;cpoqqM}t_+yWdKHSd(Jj9%>0^c(|9wogy;5X5@? za^@`k^)pA$ju~+_&fA2#E4b7KOk1G|oCgX!Rx=#&G?mdubPXYn3fT+^n@GPq%FQcb zWB!5-UR}cEyrOIivVj=M z_3y$jnux#dxr%ioWkcR) zgfY1LNZsK-SD$fwPkRTnB_HyqqaJBHk}7P8hlX{K{%a(t{QyWHQmWG+zJwo}^> z4S>I2YadR?*4mNI52~G!CD1p*>0Bu=w4xb&;1-!4WVe)xsB_||mDJncBpvBa<~)2^27`vI3pXsa z;0q$=pmUqRz3>X8In_Fr)lBID#%TIXp3DMzT`a#PhaOFXz}(XJk1a6=sJlH{w0^?x zst0PIUYggV`|6V=E{0;5Gv8~II{98q3tmA5_A8XXzqp~=L!x}6t@_Zq!8f7iIv0Ms z;q^?ZjPTb=P1BP1F+RpRjr2HudyL@8v$-Dc$RW~`vD4U-Bwj1vg|B!3_1>;$v&h_~ zl?UmqF3ZAb$)=JG*8Cgckf9tW%IooW(rpzPBgR6cjwpL%cG47xb`-*}Q=R~my~Kuh zK_>{q6PkW8!|Z|fUmPB-;6xX`ME?W#3wONri*&im+x?OGkM&tdp2ljAqnmG?Wu#kQ z8WZfNGo<)DRw8^B`AFUcFPP09B%=w8jy%eK;XK@v9|z|9tg_AjJ+j7JsyEH4+7U(Q zX#>v{fjPgg*@x@Ls@Z5N$Tq^@u|-#3jerHLNUF&j%~?-cccXyfVWhRdh5mFySXgCR z2>js+T}AlKfbb$&jE~=4&VJwILqYUoLmNS5Ap~1wrgwi-UI)dYWAj8C0p_y3q$84k zCD40p89%;i)qu>$aW=eqg=3h{Wzv@&JR#ThNN#)~$3?Z*e*W1GM6GHfLU8-bf&!mv ziV0iRT}b*`4X$v_`uXix|N`rM7+fTsm1}|W184!+FOMixZo6RY8 z)5_OoPkE+1s-zkbT`uMYJz59}vRN$SCGy9m%<@IQqqf00vCHVUDqT=@da_q`@qA=k ztzg-GF4W+H0k6z;Y&*lu}?rq9NX&}?lS@I5nIZ__cK4RP@P@(PqLD(;1<8mgdsc&>* zYvOyK)?(Yt7-1U_1jRLf_X|^3F1#XR2)7f@L)in0yUpN~fxc@L;vPUEJ82n^9Y$j` zc;2>v!bm2(M&O>L@Ob^=7Yz98wk{S<3wCV)MA93>(?5t(x z;OmOVDf-~ipQ{Z78x97c2!FL7@sii8Dt|Ha&c=;7#s!1dV~~CW*i6{#67?V&h_M^| z5aqT8JKe8IyNM7O&$SK2JYQn;YQE5H79hP?%*zWLIP_-q_b!XPY0>VnMs3P8lp|9f z>n}QODr>Vyi4_enl5T<;Fj;aGVH{H0QM_DSx-~9@<6TC6{FI56{ACZDcF92e&OdC_ z%w%E6(AUqWts#s7yV0xQ!3CATq0&LuOlO0PuiS*wfHauLudY)`ya^=t!(L=n zr^^+!1S_4Ee4)>g7=`+lVFWBP80Tkpn$#VBq>p(zV8<@oDH48oY2#I#Xyd75fH%-drq_39UF0l%8u(`@*yKW|RoS zYvC1FyVF~CSM`8jjWi4B(cD8#*%}6M&;|xnN`}x9o-$%$)Fq_0{og z#)HOcI?I9a6A?@=WS|dk7`Qrc>GM9~?ngS5(`9r-Ym@p*T0-m870qN(9b@zJt-(FW zB3b)jTjo!@-usd-HhHJDZC{4%ZBxS3sV{S!aHC>XKv(pI zJEcm&eDrKx#_@X-9jOI++Ak|5*{&$QD7{#M_}UM(aV9^I<+?s%zYIKrWTbb6w)3H{E&%U51#@ltE8eC>;c&EzL<hm^=O2|i7hVTYIIN_O zmAFc#`$&vLe9mtkqWeaAbL!OwXkrNM|LD6FHP4p0Kyv%op{89Y02xi?^^hu;KAUXe z`AuFmoYfLRyL$TNl)@|B_7{!wy3Oj;$HO~6AX*h`l&o~g9AS+3R#YK&|GNv~8ylAK z47`I4F5^c>Hr!+<0e}TJ-Mqd}6PZ81PLplt^?wYGceUS__%TTHX@RL))od;5(4ug| z-)CSl#3vu~m>{-xO=`f6XZ+iC+2QpjJLdBBGh#@Lw+5==r{!KbNFegZO+4Oh>bo9s zVr2$@=?8?bphMWbgdH_xU|zAQCqC&`WFw{S-~s9JU}K+8x}i7K!`L6@)Qg3}Xt-&Q z1#xv=o}|=-mu^TkkN9mJ8wh8o6h|$l9*t{Q%%KvhCZvl#zq!6Uz2s?YbGyDq8`NV$ zZKi0ppd4*!zY8+qDnw#BC4926#cX(T_&vH=07|w+oNOL|=DNI-_JQzGhhLOv#$$mn zCIZe58}wh5-I9igjVOfPg0+;&?d;5?M;2FP^Tv!*{Z?F4YV%OiXFL^?gzMHIeEP zmDF3D_Qq#G)5G{^SMNncmdmJmzh~pft?7r+dFRIcLRVV{!R!*den%x9u_$o=O9jIQ zn04UZhhxIM_+x+j_KTvjSB53UoHQU)tC3Sqcpz9N!8-p^#1%hBl-uuk(=1+5h{TD3 z);X6Tz7gOqDwertFiU~esQ^-2EEjo*(p&SGLe>xOFvl0UxMoS?*z@~wn< z4PX)*SX;7MJ!L7JV8v%)`*era`eTXcpUf1Vzla|a2uJvD!U#JjHwY+hWG92LFf(@o zK>7H9>P`q7O#mvV(=46Vl<2sg#8^f)fD!gK8pWZDi&M;NtvaLF_l@ zgdGgx1OeE&_&@+oUI;JXehPv8v8-fdg8-_?NvX+*vC29c86&K%5fTVXBTZ*@BU?wH ztdo(o1zgnD%o+gz0Y(4B#{=Q}i@3u3H~d=n8-qYVb!TIz-!R2xC4Rg34Vm-zMBUx- z{%6SAnj-I&)t&9^tPwW9_X7dd5M~zlVR2^?z26hT3{-Y7K{#00nlZ_m+$&f(xibq3 z|Mk;>Vn&XLKR*7?!2kCU0eJohj21u-%%{f%1;M!Z^q{=l04Oi?KD@jD9v%O0 zKp+?oSPu%mKZkoA2sa<#PhALr3-YH7%EblX4WXXjPn?f8z|n;uOPY zKtw?Yw@>+9wpv1kVjp2UPUXaii9MCGQNNW_=k4e?ZJW$-OWkLd80qW*k<9#X64#zT2wa-(l}sVbps8jwh=z+i4Gv)T z-3YE`6zWYnxjfrh&u8+@<;XSqW+Fb8(KPj8|Kt}*iM@Vr?rSJ)_e9wnxlY>$d9Oy` z^Wy{_l^c7_{v2DK6zTZb*s@cejUd(QXi2ZJCItR-VP$$nV`gWf8p8S|O6!%o<}it6 zms~;SdEqCD0wYFZjUKsj75{|QOnO!etg)}L=ULlf2~RTFSo4D&3M(j?$RAaZRVLDj zhO0VieV~X#eU;TL7u=W!&pAisFrWtVm953^ zTLj?ylic=iB;bD{kwQAy{CD^74{ZNp$^jLPoE$9N?z0i_n+^BJ8}^6gZ!G}$|DTHY z8}{$}f1khqq2Y=BSJvIXbAb=S^S4|`(^5pvLW$N=CT>yZV=G#MsC+`U+41$elvc8; zVm;Eh`JTyyzvPH9xcYf*%A}?XuWscCteo%1a;4cgYFns)Fu$wqnt87p%u&fRBYl%} z3%e^yyVQ!Sl^E69D|9bfxH0N8yTK>$xxM@v$z?WPwtHF~^%LXbQK(@Dyzi{mrI>4# zZw|vsUk>}*&^*7mcTmDrO=ytAq=`6efVUkV6|4LD6EmW!u9tYoCe}VgTtbGEhW^bQ z`U<^0x_gveI)gs6UsqywGwcB45`Y+u-FGHfo`Z~2Rh|2=)I5q+p<%($eo;*Z=4;Wz zVc&URPJI$$xm!EJC`@G?%}XEHWN-_)wN zm2L{=^H3N14{rkMJUL5`iJYOm$IirDf#fSG2k4P+m|XRZBOYpe>s^m8gott2hqtzW z*Ga<7{;3?;`L<5;;H8wzUXB>e50~=9tV~rEj*d##BTzWQ<8f>^v&`sxIxw)n_lt1B z&uGiHeWOxt z8+Vi9S(VtTT*%&M@opuDBmd{3AQpO~g9y`iX|a`35o*;$DQ{N;C)HIh{fa@AJ$Epm z@T)Z}1+z;XshlV0{pPt2a^nTFb!#4FDbG0unwOMDopCx~qu#HZ`qk1CQ54OeJR5eE z2bH{Eanu|ODsnd9`-b!I-Lq#LgAqm6#%PUm1*ta@Y(~jS=UZ&%Y)s{g9RdtTxCxKT zHM=vt30_FeVJ#D+R@+$lIKM7x2QhbCvUo^I(%L~mUv@U*Ap}gm9}E40aNllNW@pi# zQ6E?eku^;AprrkT#b+01W?|C0Kb}T~%{4|T=?GC@&Svm{c4#;+P}wI7lD4TH)bqb= z7Y(Hg8D_?&c<8{{v74r9zOxco_)67XU)Wq&zwTEIq5puq%nnsXG}Z-?iGFU%w1!$L z?5IEQ?EbLEY{*=Mk%KT9O{Uxi8D z!o*Pz@~=YnZ?dl=P#o!e-+F-mR?9$t+5Dj=Zgk)LAkF^N|C7c0H|@WbK1HO7#C@S- zlHljOZ?nMQ-^~^$7nqY31bzVmyiMGdtW~}aY6ri08j{o3j#0&{3-kG0=z+OF|B!L={Z}6#UOwJ`*z$pS|6zMygt`CWmz$ICpE91m z;d62@x^H+Len%Q^fGlKIW zz{bLY##~T14+sK=BEVoiBVJBUV+4qci`N)#%z58zBY1^jKs6-t{&s-=*j3%aHbTV>q`XlOKFcid#%g89HEQR~O0Ea#<%K!iX diff --git a/thys/CSP_RefTK/document/root.bib b/thys/CSP_RefTK/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/document/root.bib @@ -0,0 +1,11107 @@ + +% $Id: fmde.bib 6539 2010-01-29 10:33:20Z brucker $ +@InProceedings{ zhang.ea:dynamic:2003, + author = {Guangsen Zhang and Manish Parashar}, + title = {Dynamic Context-aware Access Control for Grid + Applications}, + booktitle = {GRID '03: Proceedings of the Fourth International Workshop + on Grid Computing}, + year = 2003, + pages = 101, + address = {Washington, DC, USA}, + publisher = pub-ieee, + isbn = {0-7695-2026-X}, + tags = {ReadingList, SoKNOS, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ thomas:role-based:1996, + author = {Roshan Thomas}, + title = {Role-based access control and distributed object-based + enterprise computing}, + booktitle = {RBAC '95: Proceedings of the first ACM Workshop on + Role-based access control}, + year = 1996, + pages = 21, + address = pub-acm:adr, + publisher = pub-acm, + isbn = {0-89791-759-6}, + location = {Gaithersburg, Maryland, United States}, + doi = {10.1145/270152.270194}, + tags = {ReadingList, SoKNOS, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Article{ harrison.ea:protection:1976, + author = {Michael A. Harrison and Walter L. Ruzzo and Jeffrey D. + Ullman}, + title = {Protection in operating systems}, + journal = j-cacm, + year = 1976, + volume = 19, + number = 8, + pages = {461--471}, + issn = {0001-0782}, + doi = {10.1145/360303.360333}, + publisher = pub-acm, + address = pub-acm:adr, + tags = {ReadingList, SoKNOS, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ hafner.ea:modeling:2007, + author = {Michael Hafner and Mukhtiar Memon and Muhammad Alam}, + title = {Modeling and Enforcing Advanced Access Control Policies in + Healthcare Systems with \textsc{Sectet}}, + booktitle = {MoDELS Workshops}, + year = 2007, + pages = {132--144}, + doi = {10.1007/978-3-540-69073-3_15}, + crossref = {giese:models:2007} +} + +@Proceedings{ giese:models:2007, + editor = {Holger Giese}, + title = {Models in Software Engineering, Workshops and Symposia at + MoDELS 2007, Nashville, TN, USA, September 30 - October 5, + 2007, Reports and Revised Selected Papers}, + booktitle = {MoDELS Workshops}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 5002, + year = 2008, + isbn = {978-3-540-69069-6} +} + +@InProceedings{ hu.ea:dynamic:2004, + author = {Junzhe Hu and Alfred C. Weaver}, + title = {Dynamic, Context-Aware Access Control for Distributed + Healthcare Applications}, + booktitle = PROC # { the First Workshop on Pervasive Security, + Privacy and Trust (\acs{pspt})}, + year = 2004, + tags = {SoKNOS, AccessControl}, + abstract = {The rapid worldwide deployment of the Internet and Web is + the enabler of a new generation of e-healthcare + applications, but the provision of a security architecture + that can ensure the privacy and security of sensitive + healthcare data is still an open question. Current + solutions to this problem (mostly built on static RBAC + models) are application-dependent and do not address the + intricate security requirements of healthcare applications. + The healthcare industry requires flexible, on-demand + authentication, extensible context-aware access control, + and dynamic authorization enforcement. With on-demand + authentication, users are authenticated according to their + task-specific situations. Extensible context-aware access + control enables administrators to specify more precise and + fine-grain authorization polices for any application. + Dynamic authorization enforcement makes authorization + decisions based upon runtime parameters rather than simply + the role of the user. In this paper we describe a dynamic, + context-aware security infrastructure that can fulfill the + security requirements of healthcare applications and that + can also be easily adapted to offer security support for + similar enterprise applications.}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@TechReport{ spc:break-glass:2004, + title = {Break-Glass: An Approach to Granting Emergency Access to + Healthcare Systems}, + year = 2004, + abstract = {This white paper discusses a simple yet effective + emergency-access solution, sometimes called + ``break-glass.'' The purpose of break-glass is to allow + operators emergency access to the system in cases where the + normal authentication cannot be successfully completed or + is not working properly. The systems include medical data + acquisition devices as well as information systems which + are collectively referred to as Medical Information Systems + (MedIS).}, + institution = {Joint \acs{nema}/\acs{cocir}/\acs{jira} Security and + Privacy Committee (\acs{spc})}, + type = {White paper}, + tags = {ReadingList, SoKNOS, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ ferreira.ea:how:2006, + author = {A. Ferreira and R. Cruz-Correia and L. Antunes and P. + Farinha and E. Oliveira-Palhares and D.W. Chadwick and A. + Costa-Pereira}, + title = {How to Break Access Control in a Controlled Manner}, + booktitle = PROC # {\acs{ieee} International Symposium + on Computer-Based Medical Systems (\acs{cbms})}, + year = 2006, + pages = {847--854}, + abstract = {The electronic medical record (EMR) integrates + heterogeneous information within a healthcare institution + stressing the need for security and access control. The + Biostatistics and Medical Informatics Department from Porto + Faculty of Melsdicine has recently implemented a virtual + EMR (VEMR) in order to integrate patient information and + clinical reports within a university hospital. With more + than 500 medical doctors using the system on a daily basis, + an access control policy and model were implemented. + However, the healthcare environment has unanticipated + situations (i.e. emergency situations) where access to + information is essential. Most traditional policies do not + allow for overriding. A policy that allows for + "Break-The-Glass (BTG)" was implemented in order to + override access control whilst providing for + non-repudiation mechanisms for its usage. The policy was + easily integrated within the model confirming its + modularity and the fact that user intervention in defining + security procedures is crucial to its successful + implementation and use}, + keywords = {access control, medical administrative data processing, + security of dataaccess control, clinical reports, + electronic medical record, healthcare institution, patient + information, university hospital}, + doi = {10.1109/CBMS.2006.95}, + issn = {1063-7125}, + tags = {ReadingList, SoKNOS, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Article{ basin.ea:model:2006, + author = {David A. Basin and J\"{u}rgen Doser and Torsten + Lodderstedt}, + title = {Model driven security: From {UML} models to access control + infrastructures}, + journal = j-tosem, + year = 2006, + volume = 15, + number = 1, + pages = {39--91}, + issn = {1049-331X}, + doi = {10.1145/1125808.1125810}, + publisher = pub-acm, + address = pub-acm:adr, + tags = {ReadingList, SoKNOS, AccessControl, SecureUML, MDS}, + clearance = {unclassified}, + abstract = {We present a new approach to building secure systems. In + our approach, which we call Model Driven Security, + designers specify system models along with their security + requirements and use tools to automatically generate system + architectures from the models including complete, + configured access control infrastructures. Rather than + fixing one particular modeling language for this process, + we propose a general schema for constructing such languages + that combines languages for modeling systems with languages + for modeling security. We present several instances of this + schema thatcombine (both syntactically and semantically) + different UML modeling languages with a security modeling + language for formalizing access control requirements. From + models in the combined languages, we automatically generate + access control infrastructures for server-based + applications, built from declarative and programmatic + access control mechanisms. The modeling languages and + generation process are semantically well-founded and are + based on an extension of Role-Based Access Control. We have + implemented this approach ina UML-based CASE-tool and + report on experiments.}, + timestap = {2008-05-26} +} + +@InProceedings{ schaad.ea:role-based:2001, + author = {Andreas Schaad and Jonathan Moffett and Jeremy Jacob}, + title = {The role-based access control system of a European bank: a + case study and discussion}, + booktitle = PROC # { the sixth \acs{acm} symposium on Access + control models and technologies (\acs{sacmat})}, + year = 2001, + pages = {3--9}, + address = pub-acm:adr, + publisher = pub-acm, + isbn = {1-58113-350-2}, + location = {Chantilly, Virginia, United States}, + doi = {10.1145/373256.373257}, + tags = {ReadingList, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ beznosov:requirements:1998, + author = {Konstantin Beznosov}, + title = {Requirements for access control: \acs{us} Healthcare + domain}, + booktitle = PROC # { the third \acs{acm} workshop on Role-based + access control (\acs{rbac})}, + year = 1998, + pages = 43, + address = pub-acm:adr, + publisher = pub-acm, + isbn = {1-58113-113-5}, + location = {Fairfax, Virginia, United States}, + doi = {10.1145/286884.286892}, + tags = {ReadingList, SoKNOS, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ etalle.ea:posteriori:2007, + author = {Sandro Etalle and William H. Winsborough}, + title = {A posteriori compliance control}, + booktitle = PROC # { the 12th \acs{acm} symposium on Access + control models and technologies (\acs{sacmat})}, + year = 2007, + pages = {11--20}, + address = pub-acm:adr, + publisher = pub-acm, + isbn = {978-1-59593-745-2}, + location = {Sophia Antipolis, France}, + doi = {10.1145/1266840.1266843}, + tags = {ReadingList, SoKNOS, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ lupu.ea:policy:1996, + author = {Emil C. Lupu and Damian A. Marriott and Morris S. Sloman + and Nicholas Yialelis}, + title = {A policy based role framework for access control}, + booktitle = {RBAC '95: Proceedings of the first ACM Workshop on + Role-based access control}, + year = 1996, + pages = 11, + address = pub-acm:adr, + publisher = pub-acm, + isbn = {0-89791-759-6}, + location = {Gaithersburg, Maryland, United States}, + doi = {10.1145/270152.270171}, + tags = {ReadingList, SoKNOS, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Article{ wainer.ea:dw-rbac:2007, + author = {Jacques Wainer and Akhil Kumar and Paulo Barthelmess}, + title = {DW-RBAC: A formal security model of delegation and + revocation in workflow systems}, + journal = {Inf. Syst.}, + year = 2007, + volume = 32, + number = 3, + pages = {365--384}, + abstract = {One reason workflow systems have been criticized as being + inflexible is that they lack support for delegation. This + paper shows how delegation can be introduced in a workflow + system by extending the role-based access control (RBAC) + model. The current RBAC model is a security mechanism to + implement access control in organizations by allowing users + to be assigned to roles and privileges to be associated + with the roles. Thus, users can perform tasks based on the + privileges possessed by their own role or roles they + inherit by virtue of their organizational position. + However, there is no easy way to handle delegations within + this model. This paper tries to treat the issues + surrounding delegation in workflow systems in a + comprehensive way. We show how delegations can be + incorporated into the RBAC model in a simple and + straightforward manner. The new extended model is called + RBAC with delegation in a workflow context (DW-RBAC). It + allows for delegations to be specified from a user to + another user, and later revoked when the delegation is no + longer required. The implications of such specifications + and their subsequent revocations are examined. Several + formal definitions for assertion, acceptance, execution and + 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}, + publisher = pub-elsevier, + address = {Oxford, UK, UK}, + tags = {ReadingList, SoKNOS}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ bertino.ea:flexible:1997, + author = {Elisa Bertino and Elena Ferrari and Vijayalakshmi Atluri}, + title = {A flexible model supporting the specification and + enforcement of role-based authorization in workflow + management systems}, + booktitle = {RBAC '97: Proceedings of the second ACM workshop on + Role-based access control}, + year = 1997, + pages = {1--12}, + address = pub-acm:adr, + publisher = pub-acm, + isbn = {0-89791-985-8}, + location = {Fairfax, Virginia, United States}, + doi = {10.1145/266741.266746}, + tags = {ReadingList, SoKNOS, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ bracha.ea:mixin-based:1990, + author = {Gilad Bracha and William Cook}, + title = {Mixin-based inheritance}, + booktitle = {OOPSLA/ECOOP '90: Proceedings of the European conference + on object-oriented programming on Object-oriented + programming systems, languages, and applications}, + year = 1990, + pages = {303--311}, + address = pub-acm:adr, + publisher = pub-acm, + isbn = {0-201-52430-X}, + location = {Ottawa, Canada}, + doi = {10.1145/97945.97982}, + tags = {ReadingList, OOP}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ park.ea:towards:2002, + author = {Jaehong Park and Ravi Sandhu}, + title = {Towards usage control models: beyond traditional access + control}, + booktitle = {SACMAT '02: Proceedings of the seventh ACM symposium on + Access control models and technologies}, + year = 2002, + pages = {57--64}, + address = pub-acm:adr, + publisher = pub-acm, + isbn = {1-58113-496-7}, + location = {Monterey, California, USA}, + doi = {10.1145/507711.507722}, + tags = {ReadingList, AccessControl, SoKNOS}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Article{ graham.ea:associations:1997, + author = {Ian Graham and Julia Bischof and Brian Henderson-Sellers}, + title = {Associations Considered a Bad Thing}, + journal = {JOOP}, + year = 1997, + volume = 9, + number = 9, + pages = {41--48}, + tags = {ReadingList, OOP}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ scharli.ea:traits:2003, + author = {Nathanael Sch{\"a}rli and St{\'e}phane Ducasse and Oscar + Nierstrasz and Andrew P. Black}, + title = {Traits: Composable Units of Behaviour}, + booktitle = {ECOOP}, + year = 2003, + pages = {248--274}, + ee = {http://springerlink.metapress.com/openurl.asp?genre=article{\&}issn=0302-9743{\&}volume=2743{\&}spage=248} + , + crossref = {cardelli:ecoop:2003}, + tags = {ReadingList, OOP}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ cardelli:ecoop:2003, + title = {ECOOP 2003 - Object-Oriented Programming, 17th European + Conference, Darmstadt, Germany, July 21-25, 2003, + Proceedings}, + year = 2003, + editor = {Luca Cardelli}, + volume = 2743, + series = s-lncs, + publisher = pub-springer, + booktitle = {ECOOP}, + isbn = {3-540-40531-3}, + tags = {ReadingList, OOP}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ wolter.ea:modelling:2008, + author = {Christian Wolter and Michael Menzel and Christoph Meinel}, + title = {Modelling Security Goals in Business Processes}, + booktitle = {Modellierung}, + year = 2008, + pages = {197--212}, + crossref = {kuhne.ea:modellierung:2008}, + tags = {ReadingList, MDS, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ kuhne.ea:modellierung:2008, + title = {Modellierung 2008, 12.-14. M{\"a}rz 2008, Berlin}, + year = 2008, + editor = {Thomas K{\"u}hne and Wolfgang Reisig and Friedrich + Steimann}, + volume = 127, + series = {LNI}, + publisher = {GI}, + booktitle = {Modellierung}, + isbn = {978-3-88579-221-5}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ wolter.ea:deriving:2007, + author = {Christian Wolter and Andreas Schaad and Christoph Meinel}, + title = {Deriving XACML Policies from Business Process Models}, + booktitle = {WISE Workshops}, + year = 2007, + pages = {142--153}, + ee = {http://dx.doi.org/10.1007/978-3-540-77010-7_15}, + crossref = {weske.ea:web:2007}, + tags = {ReadingList, SoKNOS, MDS}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ weske.ea:web:2007, + title = {Web Information Systems Engineering - WISE 2007 Workshops, + WISE 2007 International Workshops, Nancy, France, December + 3, 2007, Proceedings}, + year = 2007, + editor = {Mathias Weske and Mohand-Said Hacid and Claude Godart}, + volume = 4832, + series = s-lncs, + publisher = pub-springer, + booktitle = {WISE Workshops}, + isbn = {978-3-540-77009-1}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ sandhu.ea:nist:2000, + author = {Ravi S. Sandhu and David F. Ferraiolo and D. Richard + Kuhn}, + title = {The NIST model for role-based access control: towards a + unified standard}, + booktitle = {ACM Workshop on Role-Based Access Control}, + year = 2000, + pages = {47--63}, + doi = {10.1145/344287.344301}, + tags = {ReadingList, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Article{ ferraiolo.ea:proposed:2001, + author = {David F. Ferraiolo and Ravi S. Sandhu and Serban I. + Gavrila and D. Richard Kuhn and Ramaswamy Chandramouli}, + title = {Proposed \acs{nist} standard for role-based access + control}, + journal = j-tissec, + year = 2001, + pub = pub-acm, + address = pub-acm:adr, + volume = 4, + number = 3, + pages = {224--274}, + doi = {10.1145/501978.501980}, + tags = {ReadingList, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ drouineaud.ea:first:2004, + author = {Michael Drouineaud and Maksym Bortin and Paolo Torrini and + Karsten Sohr}, + title = {A First Step Towards Formal Verification of Security + Policy Properties for \acs{rbac}}, + booktitle = {QSIC}, + year = 2004, + pages = {60--67}, + doi = {10.1109/QSIC.2004.1357945}, + crossref = {anonymous:qsic:2004}, + tags = {AccessControl, FormalMethods, TheoremProving}, + abstract = {Considering the current expansion of IT-infrastructure the + security of the data inside this infrastructure becomes + increasingly important. Therefore assuring certain security + properties of IT-systems by formal methods is desirable. So + far in security formal methods have mostly been used to + prove properties of security protocols. However, access + control is an indispensable part of security inside a given + IT-system, which has not yet been sufficiently examined + using formal methods. The paper presents an example of a + RBAC security policy having the dual control property. This + is proved in a first-order linear temporal logic (LTL) that + has been embedded in the theorem prover Isabelle/HOL by the + authors. Thus the correctness of the proof is assured by + Isabelle/HOL. The authors consider first-order LTL a good + formalism for expressing RBAC authorisation constraints and + deriving properties from given RBAC security policies. + Furthermore it might also be applied to safety-related + issues in similar manner.}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ anonymous:qsic:2004, + title = {4th International Conference on Quality Software (QSIC + 2004), 8-10 September 2004, Braunschweig, Germany}, + year = 2004, + address = pub-ieee:adr, + publisher = pub-ieee, + booktitle = {QSIC}, + isbn = {0-7695-2207-6}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ sohr.ea:specification:2005, + author = {Karsten Sohr and Gail-Joon Ahn and Martin Gogolla and Lars + Migge}, + title = {Specification and Validation of Authorisation Constraints + Using UML and OCL}, + booktitle = {ESORICS}, + year = 2005, + pages = {64--79}, + ee = {http://dx.doi.org/10.1007/11555827_5}, + crossref = {vimercati.ea:computer:2005}, + abstracts = {Abstract. Authorisation constraints can help the policy + architect design and express higher-level security policies + for organisations such as financial institutes or + governmental agencies. Although the importance of + constraints has been addressed in the literature, there + does not exist a systematic way to validate and test + authorisation constraints. In this paper, we attempt to + specify non-temporal constraints and history-based + constraints in Object Constraint Language (OCL) which is a + constraint specification language of Unified Modeling + Language (UML) and describe how we can facilitate the USE + tool to validate and test such policies. We also discuss + the issues of identification of conflicting constraints and + missing constraints. }, + tags = {AccessControl, SecureUML}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ vimercati.ea:computer:2005, + title = {Computer Security - ESORICS 2005, 10th European Symposium + on Research in Computer Security, Milan, Italy, September + 12-14, 2005, Proceedings}, + year = 2005, + editor = {Sabrina De Capitani di Vimercati and Paul F. Syverson and + Dieter Gollmann}, + volume = 3679, + series = s-lncs, + publisher = pub-springer, + booktitle = {ESORICS}, + isbn = {3-540-28963-1}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ wolter.ea:modeling:2007, + author = {Christian Wolter and Andreas Schaad}, + title = {Modeling of Task-Based Authorization Constraints in BPMN}, + booktitle = {BPM}, + year = 2007, + pages = {64--79}, + ee = {http://dx.doi.org/10.1007/978-3-540-75183-0_5}, + crossref = {alonso.ea:business:2007}, + tags = {ReadingList, SoKNOS, MDS, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ alonso.ea:business:2007, + title = {Business Process Management, 5th International Conference, + BPM 2007, Brisbane, Australia, September 24-28, 2007, + Proceedings}, + year = 2007, + editor = {Gustavo Alonso and Peter Dadam and Michael Rosemann}, + volume = 4714, + series = s-lncs, + publisher = pub-springer, + booktitle = {BPM}, + isbn = {978-3-540-75182-3}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Article{ sohr.ea:analyzing:2008, + author = {Karsten Sohr and Michael Drouineaud and Gail-Joon Ahn and + Martin Gogolla}, + title = {Analyzing and Managing Role-Based Access Control + Policies}, + journal = j-tkde, + year = 2008, + doi = {10.1109/TKDE.2008.28}, + abstract = {Today more and more security-relevant data is stored on + computer systems; security-critical business processes are + mapped to their digital counterparts. This situation + applies to various domains such as health care industry, + digital government, and financial service institutes + requiring that different security requirements must be + fulfilled. Authorisation constraints can help the policy + architect design and express higher-level organisational + rules. Although the importance of authorisation constraints + has been addressed in the literature, there does not exist + a systematic way to verify and validate authorisation + constraints. In this paper, we specify both non-temporal + and history-based authorisation constraints in the Object + Constraint Language (OCL) and first-order linear temporal + logic (LTL). Based upon these specifications, we attempt to + formally verify role-based access control policies with the + help of a theorem prover and to validate policies with the + USE system, a validation tool for OCL constraints. We also + describe an authorisation engine, which supports the + enforcement of authorisation constraints.}, + tags = {ReadingList, AccessControl}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Article{ samuel.ea:context-aware:2008, + author = {Samuel, A. and Ghafoor, A. and Bertino, E.}, + title = {Context-Aware Adaptation of Access-Control Policies}, + journal = {Internet Computing, IEEE}, + year = 2008, + volume = 12, + number = 1, + pages = {51--54}, + abstract = {Today, public-service delivery mechanisms such as + hospitals, police, and fire departments rely on digital + generation, storage, and analysis of vital information. To + protect critical digital resources, these organizations + employ access-control mechanisms, which define rules under + which authorized users can access the resources they need + to perform organizational tasks. Natural or man-made + disasters pose a unique challenge, whereby previously + defined constraints can potentially debilitate an + organization's ability to act. Here, the authors propose + employing contextual parameters - specifically, activity + context in the form of emergency warnings - to adapt + access-control policies according to a priori + configuration.}, + keywords = {authorisation, disasters, organisational + aspectsaccess-control policy, context-aware adaptation, + digital resource protection, natural disaster, + organizational task, public-service delivery mechanism}, + doi = {10.1109/MIC.2008.6}, + issn = {1089-7801}, + tags = {ReadingList, AccessControl, SoKNOS}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Article{ adam.ea:secure:2007, + author = {Nabil Adam and Ahmet Kozanoglu and Aabhas Paliwal and + Basit Shafiq}, + title = {Secure Information Sharing in a Virtual Multi-Agency Team + Environment}, + journal = j-entcs, + year = 2007, + volume = 179, + pages = {97--109}, + issn = {1571-0661}, + abstract = {This paper proposes a two tier RBAC approach for secure + and selective information sharing among virtual + multi-agency response team (VMART) and allows expansion of + the VMART by admitting new collaborators (government + agencies or NGOs) as need arise. A coordinator Web Service + for each member agency is proposed. The coordinator Web + Service is responsible for authentication, information + dissemination, information acquisition, role creation and + enforcement of predefined access control policies. Secure, + selective and fine-grained information sharing is realized + through the encryption of XML documents according to RBAC + policies defined for the corresponding XML schema.}, + doi = {http://dx.doi.org/10.1016/j.entcs.2006.08.034}, + publisher = pub-elsevier, + address = pub-elsevier:adr, + tags = {ReadingList, SoKNOS}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ alam.ea:framework:2006, + author = {Muhammad Alam and Michael Hafner and Ruth Breu and Stefan + Unterthiner}, + title = {A Framework for Modeling Restricted Delegation in Service + Oriented Architecture}, + booktitle = {TrustBus}, + year = 2006, + pages = {142--151}, + ee = {http://dx.doi.org/10.1007/11824633_15}, + crossref = {fischer-hubner.ea:trust:2006}, + tags = {ReadingList, SoKNOS}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ fischer-hubner.ea:trust:2006, + title = {Trust and Privacy in Digital Business, Third International + Conference, TrustBus 2006, Krakow, Poland, September 4-8, + 2006, Proceedings}, + year = 2006, + editor = {Simone Fischer-H{\"u}bner and Steven Furnell and Costas + Lambrinoudakis}, + volume = 4083, + series = s-lncs, + publisher = pub-springer, + booktitle = {TrustBus}, + isbn = {3-540-37750-6}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ el-atawy.ea:policy:2005, + author = {Adel El-Atawy and K. Ibrahim and H. Hamed and Ehab + Al-Shaer}, + title = {Policy segmentation for intelligent firewall testing}, + booktitle = {NPSec 05}, + year = 2005, + pages = {67--72}, + month = nov, + publisher = pub-ieee, + day = 6, + keywords = {computer networks, intelligent networks, telecommunication + security, telecommunication traffic intelligent firewall + testing, network security, network traffic, packet + filtering algorithms, policy segmentation}, + acknowledgement={none}, + tags = {ReadingList, Testing, FWTesting}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ el-atawy.ea:automated:2007, + author = {Adel El-Atawy and Taghrid Samak and Zein Wali and Ehab + Al-Shaer and Frank Lin and Christopher Pham and Sheng Li}, + title = {An Automated Framework for Validating Firewall Policy + Enforcement}, + booktitle = {\acs{policy} '07}, + year = 2007, + pages = {151--160}, + publisher = pub-ieee, + acknowledgement={none}, + tags = {ReadingList, Testing, FWTesting}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ bishop.ea:engineering:2006, + author = {Steve Bishop and Matthew Fairbairn and Michael Norrish and + Peter Sewell and Michael Smith and Keith Wansbrough}, + title = {Engineering with logic: \acs{hol} specification and + symbolic-evaluation testing for \acs{tcp} implementations}, + booktitle = {\acs{popl}}, + year = 2006, + pages = {55--66}, + crossref = {morrisett.ea:proceedings:2006}, + acknowledgement={none}, + tags = {ReadingList, FWTesting, Testing}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ morrisett.ea:proceedings:2006, + title = {\acs{popl}}, + year = 2006, + editor = {J. Gregory Morrisett and Simon L. Peyton Jones}, + publisher = pub-acm, + booktitle = {\acs{popl}}, + adress = pub-acm:adr, + acknowledgement={none}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ marmorstein.ea:firewall:2006, + author = {Robert Marmorstein and Phil Kearns}, + title = {Firewall analysis with policy-based host classification}, + booktitle = {\acs{lisa}'06}, + year = 2006, + pages = {4--4}, + publisher = {\acs{usenix} Association}, + location = {Washington, \acs{dc}}, + acknowledgement={none}, + tags = {ReadingList, FWTesting, Testing}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ richters.ea:formalizing:1998, + abstract = {We present a formal semantics for the Object Constraint + Language (OCL) which is part of the Unified Modeling + Language (UML) - an emerging standard language and notation + for object-oriented analysis and design. In context of + information systems modeling, UML class diagrams can be + utilized for describing the overall structure, whereas + additional integrity constraints and queries are specified + with OCL expressions. By using OCL, constraints and queries + can be specified in a formal yet comprehensible way. + However, the OCL itself is currently defined only in a + semi-formal way. Thus the semantics of constraints is in + general not precisely defined. Our approach gives precise + meaning to OCL concepts and to some central aspects of UML + class models. A formal semantics facilitates verification, + validation and simulation of models and helps to improve + the quality of models and software designs.}, + bibkey = {richters.ea:formalizing:1998}, + author = {Mark Richters and Martin Gogolla}, + title = {On Formalizing the \acs{uml} Object Constraint Language + \acs{ocl}}, + pages = {449--464}, + doi = {10.1007/b68220}, + crossref = {ling.ea:conceptual:1998}, + acknowledgement={brucker, 2007-02-19}, + tags = {MDE}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ ling.ea:conceptual:1998, + language = {USenglish}, + editor = {Tok Wang Ling and Sudha Ram and Mong-Li Lee}, + booktitle = {Conceptual Modeling---{ER} '98}, + title = {Conceptual Modeling---{ER} '98}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 1507, + doi = {10.1007/b68220}, + year = 1998, + isbn = {978-3-540-65189-5}, + acknowledgement={brucker, 2007-02-19}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ cook.ea::amsterdam:2002, + abstract = {In November 1998 the authors participated in a two-day + workshop on the Object Constraint Language (OCL) in + Amsterdam. The focus was to clarify issues about the + semantics and the use of OCL, and to discuss useful and + necessary extensions of OCL. Various topics have been + raised and clarified. This manifesto contains the results + of that workshop and the following work on these topics. + Overview of OCL.}, + author = {Steve Cook and Anneke Kleppe and Richard Mitchell and + Bernhard Rumpe and Jos Warmer and Alan Wills}, + title = {The Amsterdam Manifesto on \acs{ocl}}, + pages = {115--149}, + crossref = {clark.ea:object:2002}, + acknowledgement={brucker, 2007-02-19}, + tags = {MDE}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ gogolla.ea:expressing:2001, + author = {Martin Gogolla and Mark Richters}, + bibkey = {gogolla.ea:expressing:2001}, + abstract = {The Unified Modeling Language \acs{uml} is a complex + language offering many modeling features. Especially the + description of static structures with class diagrams is + supported by a rich set of primitives. This paper shows how + to transfrom \acs{uml} class diagrams involving cardinality + constraints, qualifiers, association classes, aggregations, + compositions, and generalizations into equivalent \acs{uml} + class diagrams employing only binary associations and + \acs{ocl} constraints. Thus we provide a better + understanding of \acs{uml} features. By reducing more + complex features in terms of basic ones, we suggest an easy + way users can gradually extend the set of \acs{uml} + elements they commonly apply in the modeling process.}, + title = {Expressing \acs{uml} Class Diagrams Properties with + \acs{ocl}}, + pages = {85--114}, + crossref = {clark.ea:object:2002}, + acknowledgement={brucker, 2007-02-19}, + tags = {MDE}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ richters.ea:ocl:2001, + abstract = {{The Object Constraint Language \acs{ocl} allows to + formally specify constraints on a \acs{uml} model. We + present a formal syntax and semantics for \acs{ocl} based + on set theory including expressions, invariants and pre- + and postconditions. A formal foundation for \acs{ocl} makes + the meaning of constraints precise and helps to eliminate + ambiguities and inconsistencies. A precise language de + nition is also a prerequisite for implementing CASE tools + providing enhanced support for \acs{uml} models and + \acs{ocl} constraints. We give a survey of some \acs{ocl} + tools and discuss one of the tools in some more detail. The + design and implementation of the USE tool supporting the + validation of \acs{uml} models and \acs{ocl} constraints is + based on the formal approach presented in this paper.}}, + bibkey = {richters.ea:ocl:2001}, + author = {Mark Richters and Martin Gogolla}, + title = {\acs{ocl}: Syntax, Semantics, and Tools.}, + pages = {42--68}, + crossref = {clark.ea:object:2002}, + acknowledgement={brucker, 2007-02-19}, + tags = {MDE}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ hennicker.ea:precise:2002, + author = {Rolf Hennicker and Heinrich Hu{\ss}mann and Michel Bidoit}, + title = {On the Precise Meaning of \acs{ocl} Constraints}, + pages = {69--84}, + crossref = {clark.ea:object:2002}, + abstract = {When OCL is applied in concrete examples, many questions + arise about the precise meaning of OCL constraints. The + same kind of difficulties appears when automatic support + tools for OCL are designed. These questions are due to the + lack of a precise semantics of OCL constraints in the + context of a UML model. The aim of this paper is to + contribute to a clarification of several issues, like + interpretation of invariants and pre- and postconditions, + treatment of undefined values, inheritance of constraints, + transformation rules for OCL constraints and computation of + proof obligations. Our study is based on a formal, abstract + semantics of OCL.}, + bibkey = {hennicker.ea:precise:2002}, + acknowledgement={brucker, 2007-02-19}, + tags = {MDE}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ clark.ea:object:2002, + editor = {Tony Clark and Jos Warmer}, + booktitle = {Object Modeling with the \acs{ocl}: The Rationale behind + the Object Constraint Language}, + title = {Object Modeling with the \acs{ocl}: The Rationale behind + the Object Constraint Language}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 2263, + year = 2002, + isbn = {3-540-43169-1}, + acknowledgement={brucker, 2007-02-19}, + tags = {MDE}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ cengarle.ea:formal:2001, + author = {Mar\'{\i}a Victoria Cengarle and Alexander Knapp}, + title = {A Formal Semantics for \acs{ocl} 1.4}, + year = 2001, + abstract = {The OCL 1.4 specification introduces let-declarations for + adding auxiliary class features in static structures of the + UML. We provide a type inference system and a big-step + operational semantics for the OCL 1.4 that treat UML static + structures and UML object models and accommodate for + additional declarations; the operational semantics + satisfies a subject reduction property with respect to the + type inference system. We also discuss an alternative, + non-operational interpretation of let-declarations as + constraints.}, + pages = {118--133}, + crossref = {gogolla.ea:uml:2001}, + acknowledgement={brucker, 2007-02-19}, + tags = {MDE}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ gogolla.ea:uml:2001, + editor = {Martin Gogolla and Cris Kobryn}, + booktitle = {\acs{uml} 2001---The Unified Modeling Language. Modeling + Languages, Concepts, and Tools}, + title = {\acs{uml} 2001---The Unified Modeling Language. Modeling + Languages, Concepts, and Tools}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 2185, + year = 2001, + isbn = {3-540-42667-1}, + location = {Toronto, Canada}, + acknowledgement={brucker, 2007-02-19}, + tags = {MDE}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@TechReport{ balzer.ea:objects:2008, + author = {Stephanie Balzer and Alexandra Burns and Thomas R. Gross}, + title = {Objects in Context: An Empirical Study of Object + Relationships}, + institution = {\acs{eth} Zurich}, + year = 2008, + abstract = {Object collaborations are at the core of all + object-oriented programming, yet current class-based + objectoriented programming languages do not provide an + explicit construct to capture the relationships between + objects. This paper reports on an empirical study that + investigates the occurrence of object collaborations to + assess the need of intrinsic support for relationships in a + programming language. We introduce a categorization of + possible forms of object collaborations and their + corresponding implementation patterns when using a + traditional class-based object-oriented language (Java) and + analyze 25 Java programs (ranging from 4 to 6275 classes) + with the Relationship Detector for Java (RelDJ) to identify + occurrences of these patterns. The empirical results show + that object collaborations are indeed a frequent phenomenon + and reveal that collaborationrelated code does not remain + encapsulated in a single class. These observations strongly + support efforts to define language constructs to express + object relationships: relationships allow the encapsulation + of a frequently occurring phenomenon and increase program + expressiveness. }, + keywords = {Relationship-based Programming Languages, First-class + Relationships, Object Collaborations, Java, Bytecode + Analysis}, + tags = {ReadingList, OOP}, + clearance = {unclassified}, + timestap = {2008-05-28}, + number = 594 +} + +@InProceedings{ owre.ea:pvs:1996, + author = {Sam Owre and S. Rajan and John M. Rushby and Natarajan + Shankar and Mandayam K. Srivas}, + title = {\acs{pvs}: Combining Specification, Proof Checking, and + Model Checking}, + year = 1996, + bibkey = {owre.ea:pvs:1996}, + pages = {411--414}, + crossref = {alur.ea:computer:1996}, + doi = {10.1007/3-540-61474-5_91}, + acknowledgement={brucker, 2007-02-19}, + tags = {TheoremProving, FormalMethods}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ alur.ea:computer:1996, + editor = {Rajeev Alur and Thomas A. Henzinger}, + booktitle = {Computer Aided Verification (\acs{cav})}, + title = {Computer Aided Verification (\acs{cav})}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 1102, + year = 1996, + location = {New Brunswick, \acs{nj}, \acs{usa}}, + doi = {10.1007/3-540-61474-5}, + isbn = {3-540-61474-5}, + acknowledgement={brucker, 2007-02-19}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ naraschewski.ea:object-oriented:1998, + author = {Wolfgang Naraschewski and Markus Wenzel}, + title = {Object-Oriented Verification Based on Record Subtyping in + Higher-Order Logic.}, + pages = {349--366}, + doi = {10.1007/BFb0055146}, + crossref = {grundy.ea:theorem:1998}, + abstract = {We show how extensible records with structural subtyping + can be represented directly in Higher-Order Logic + (\acs{hol}). Exploiting some specific properties of + \acs{hol}, this encoding turns out to be extremely simple. + In particular, structural subtyping is subsumed by naive + parametric polymorphism, while overridable generic + functions may be based on overloading. Taking \acs{hol} + plus extensible records as a starting point, we then set + out to build an environment for object-oriented + specification and verification (HOOL). This framework + offers several well-known concepts like classes, objects, + methods and late-binding. All of this is achieved by very + simple means within \acs{hol}. }, + keywords = {Isabelle/\acs{hol}, extensible records, + object-orientation, verification}, + bibkey = {naraschewski.ea:object-oriented:1998}, + acknowledgement={brucker, 2007-02-19}, + tags = {TheoremProving, FormalMethods}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ grundy.ea:theorem:1998, + editor = {Jim Grundy and Malcolm C. Newey}, + title = {Theorem Proving in Higher Order Logics (\acs{tphols})}, + booktitle = {Theorem Proving in Higher Order Logics (\acs{tphols})}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 1479, + year = 1998, + doi = {10.1007/BFb0055125}, + isbn = {3-540-64987-5}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ oheimb.ea:hoare:2002, + author = {David von Oheimb and Tobias Nipkow}, + title = {Hoare Logic for {NanoJava}: Auxiliary Variables, Side + Effects, and Virtual Methods Revisited}, + pages = {89--105}, + doi = {10.1007/3-540-45614-7_6}, + crossref = {eriksson.ea:fme:2002}, + acknowledgement={brucker, 2007-02-19}, + abstract = {We define NanoJava, a kernel of Java tailored to the + investigation of Hoare logics. We then introduce a Hoare + logic for this language featuring an elegant new approach + for expressing auxiliary variables: by universal + quantification on the outer logical level. Furthermore, we + give simple means of handling side-effecting expressions + and dynamic binding within method calls. The logic is + proved sound and (relatively) complete using + Isabelle/\acs{hol}.}, + keywords = {Languages, Reliability, Theory, Verification}, + tags = {OOP, FormalMethods}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Proceedings{ eriksson.ea:fme:2002, + editor = {Lars-Henrik Eriksson and Peter Alexander Lindsay}, + booktitle = {\acs{fme} 2002: Formal Methods---Getting {IT} Right}, + title = {\acs{fme} 2002: Formal Methods---Getting {IT} Right}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 2391, + doi = {10.1007/3-540-45614-7}, + year = 2002, + isbn = {3-540-43928-5}, + acknowledgement={brucker, 2007-02-19}, + tags = {noTAG}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@Article{ rudnicki:obvious:1987, + author = {Piotr Rudnicki}, + key = {Rudnicki}, + journal = j-ar, + title = {Obvious Inferences}, + year = 1987, + month = dec, + volume = 3, + doi = {10.1007/BF00247436}, + number = 4, + abstract = {The notion of 'obvious' inference in predicate logic is + discussed from the viewpoint of proof- checker applications + in logic and mathematics education. A class of inferences + in predicate logic is defined and it is proposed to + identify it with the class of 'obvious' logical inferences. + The definition is compared with other approaches. The + algorithm for implementing the "obviousness' decision + procedure follows directly from the definition.}, + pages = {383--394}, + publisher = pub-springer-netherlands, + address = pub-springer-netherlands:adr, + tags = {noTAG}, + timestap = {2008-05-26} +} + +@InProceedings{ stroustrup:what:1987, + author = {Bjarne Stroustrup}, + title = {What is "Object-Oriented Programming?"}, + booktitle = {ECOOP}, + year = 1987, + pages = {51--70}, + abstract = {"Object-Oriented Programming" and "Data ion" have become + very common terms. Unfortunately, few people agree on what + they mean. I will offer informal definitions that appear to + make sense in the context of languages like Ada, C++, + Modula-2, Simula67, and Smalltalk. The general idea is to + equate "support for data abstraction" with the ability to + define and use new types and equate "support for + object-oriented programming" with the ability to express + type hierarchies. Features necessary to support these + programming styles in a general purpose programming + language will be discussed. The presentation centers around + C++ but is not limited to facilities provided by that + language.}, + crossref = {bezivin.ea:ecoop87:1987} +} + +@Proceedings{ bezivin.ea:ecoop87:1987, + editor = {Jean B{\'e}zivin and Jean-Marie Hullot and Pierre Cointe + and Henry Lieberman}, + title = {ECOOP'87 European Conference on Object-Oriented + Programming, Paris, France, June 15-17, 1987, Proceedings}, + booktitle = {ECOOP}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 276, + year = 1987, + isbn = {3-540-18353-1}, + location = {Paris, France} +} + +@Book{ nipkow.ea:isabelle:2002, + author = {Tobias Nipkow and Lawrence C. Paulson and Markus Wenzel}, + title = {Isabelle/\acs{hol}---A Proof Assistant for Higher-Order + Logic}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 2283, + doi = {10.1007/3-540-45949-9}, + abstract = {This book is a self-contained introduction to interactive + proof in higher-order logic (\acs{hol}), using the proof + assistant Isabelle2002. It is a tutorial for potential + users rather than a monograph for researchers. The book has + three parts. + + 1. Elementary Techniques shows how to model functional + programs in higher-order logic. Early examples involve + lists and the natural numbers. Most proofs are two steps + long, consisting of induction on a chosen variable followed + by the auto tactic. But even this elementary part covers + such advanced topics as nested and mutual recursion. 2. + Logic and Sets presents a collection of lower-level tactics + that you can use to apply rules selectively. It also + describes Isabelle/\acs{hol}'s treatment of sets, functions + and relations and explains how to define sets inductively. + One of the examples concerns the theory of model checking, + and another is drawn from a classic textbook on formal + languages. 3. Advanced Material describes a variety of + other topics. Among these are the real numbers, records and + overloading. Advanced techniques are described involving + induction and recursion. A whole chapter is devoted to an + extended example: the verification of a security protocol. }, + year = 2002, + acknowledgement={brucker, 2007-02-19}, + bibkey = {nipkow.ea:isabelle:2002}, + tags = {noTAG}, + clearance = {unclassified}, + timestap = {2008-05-26} +} + +@InProceedings{ kerber.ea:mechanization:1994, + author = {Manfred Kerber and Michael Kohlhase}, + title = {A Mechanization of Strong Kleene Logic for Partial + Functions}, + pages = {371--385}, + crossref = {bundy:automated:1994}, + abstract = {Even though it is not very often admitted, partial + functions do play a significant role in many practical + applications of deduction systems. Kleene has already given + a semantic account of partial functions using three-valued + logic decades ago, but there has not been a satisfactory + mechanization. Recent years have seen a thorough + investigation of the framework of many-valued + truth-functional logics. However, strong Kleene logic, + where quantification is restricted and therefore not + truth-functional, does not fit the framework directly. We + solve this problem by applying recent methods from sorted + logics. This paper presents a resolution calculus that + combines the proper treatment of partial functions with the + efficiency of sorted calculi.}, + doi = {10.1007/3-540-58156-1_26}, + acknowledgement={brucker, 2007-02-19}, + bibkey = {kerber.ea:mechanization:1994} +} + +@Proceedings{ bundy:automated:1994, + editor = {Alan Bundy}, + booktitle = {Automated Deduction---\acs{cade}-12}, + title = {Automated Deduction---\acs{cade}-12}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + location = {Nancy, France}, + volume = 814, + year = 1994, + isbn = {3-540-58156-1}, + acknowledgement={brucker, 2007-02-19}, + doi = {10.1007/3-540-58156-1} +} + +@InProceedings{ hahnle:towards:1991, + author = {Reiner H{\"a}hnle}, + title = {Towards an Efficient Tableau Proof Procedure for + Multiple-Valued Logics}, + pages = {248--260}, + crossref = {borger.ea:computer:1991}, + doi = {10.1007/3-540-54487-9_62}, + abstract = {One of the obstacles against the use of tableau-based + theorem provers for non-standard logics is the inefficiency + of tableau systems in practical applications, though they + are highly intuitive and extremely flexible from a proof + theoretical point of view. We present a method for + increasing the efficiency of tableau systems in the case of + multiple-valued logics by introducing a generalized notion + of signed formulas and give sound and complete tableau + systems for arbitrary propositional finite-valued logics.}, + acknowledgement={brucker, 2007-02-19}, + bibkey = {hahnle:towards:1991} +} + +@Proceedings{ borger.ea:computer:1991, + editor = {Egon B{\"o}rger and Hans Kleine B{\"u}ning and Michael M. + Richter and Wolfgang Sch{\"o}nfeld}, + title = {Computer Science Logic (\acs{csl})}, + booktitle = {Computer Science Logic (\acs{csl})}, + series = s-lncs, + volume = 533, + year = 1991, + doi = {10.1007/3-540-54487-9}, + isbn = {978-3-540-54487-6}, + acknowledgement={brucker, 2007-02-19}, + publisher = pub-springer, + address = pub-springer:adr +} + +@InProceedings{ berghofer.ea:inductive:1999, + author = {Stefan Berghofer and Markus Wenzel}, + title = {Inductive datatypes in \acs{hol}---lessons learned in + Formal-Logic Engineering}, + pages = {19--36}, + crossref = {bertot.ea:theorem:1999}, + bibkey = {berghofer.ea:inductive:1999}, + month = sep, + doi = {10.1007/3-540-48256-3_3}, + acknowledgement={brucker, 2007-02-19} +} + +@Proceedings{ bertot.ea:theorem:1999, + editor = {Yves Bertot and Gilles Dowek and Andr{\'e} Hirschowitz and + C. Paulin and Laurent Th{\'e}ry}, + title = {Theorem Proving in Higher Order Logics (\acs{tphols})}, + booktitle = {Theorem Proving in Higher Order Logics (\acs{tphols})}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 1690, + acknowledgement={brucker, 2007-02-19}, + location = {Nice, France}, + year = 1999, + isbn = {3-540-66463-7} +} + +@Proceedings{ geuvers.ea:types:2003, + editor = {Herman Geuvers and Freek Wiedijk}, + title = {Types for Proofs and Programs (\acs{types})}, + booktitle = {Types for Proofs and Programs (\acs{types})}, + publisher = pub-springer, + location = {Nijmegen}, + address = pub-springer:adr, + series = s-lncs, + volume = 2646, + language = {USenglish}, + year = 2003, + isbn = {3-540-14031-X}, + acknowledgement={brucker, 2007-02-19} +} + +@InProceedings{ angelo.ea:degrees:1994, + author = {Catia M. Angelo and Luc J. M. Claesen and Hugo De Man}, + title = {Degrees of Formality in Shallow Embedding Hardware + Description Languages in \acs{hol}}, + pages = {89--100}, + doi = {10.1007/3-540-57826-9_127}, + crossref = {joyce.ea:higher:1994}, + bibkey = {angelo.ea:degrees:1994}, + acknowledgement={brucker, 2007-02-19} +} + +@Proceedings{ joyce.ea:higher:1994, + editor = {Jeffrey J. Joyce and Carl-Johan H. Seger}, + title = {Higher Order Logic Theorem Proving and Its Applications + (\acs{hug})}, + booktitle = {Higher Order Logic Theorem Proving and Its Applications + (\acs{hug})}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + abstract = {Theorem proving based techniques for formal hardware + verification have been evolving constantly and researchers + are getting able to reason about more complex issues than + it was possible or practically feasible in the past. It is + often the case that a model of a system is built in a + formal logic and then reasoning about this model is carried + out in the logic. Concern is growing on how to consistently + interface a model built in a formal logic with an informal + CAD environment. Researchers have been investigating how to + define the formal semantics of hardware description + languages so that one can formally reason about models + informally dealt with in a CAD environment. At the + University of Cambridge, the embedding of hardware + description languages in a logic is classified in two + categories: deep embedding and shallow embedding. In this + paper we argue that there are degrees of formality in + shallow embedding a language in a logic. The choice of the + degree of formality is a trade-off between the security of + the embedding and the amount and complexity of the proof + effort in the logic. We also argue that the design of a + language could consider this verifiability issue. There are + choices in the design of a language that can make it easier + to improve the degree of formality, without implying + serious drawbacks for the CAD environment.}, + volume = 780, + year = 1994, + doi = {10.1007/3-540-57826-9}, + isbn = {3-540-57826-9}, + acknowledgement={brucker, 2007-02-19} +} + +@InProceedings{ huffman.ea:axiomatic:2005, + author = {Brian Huffman and John Matthews and Peter White}, + title = {Axiomatic Constructor Classes in {Isabelle}/\acs{holcf}.}, + pages = {147--162}, + doi = {10.1007/11541868_10}, + crossref = {hurd.ea:theorem:2005}, + acknowledgement={brucker, 2007-02-19}, + bibkey = {huffman.ea:axiomatic:2005}, + abstract = {We have definitionally extended Isabelle/HOLCF to support + axiomatic Haskell-style constructor classes. We have + subsequently defined the functor and monad classes, + together with their laws, and implemented state and + resumption monad transformers as generic constructor class + instances. This is a step towards our goal of giving + modular denotational semantics for concurrent lazy + functional programming languages, such as GHC Haskell.} +} + +@InProceedings{ marche.ea:reasoning:2005, + author = {Claude March{\'e} and Christine Paulin-Mohring}, + title = {Reasoning About {Java} Programs with Aliasing and Frame + Conditions}, + pages = {179--194}, + crossref = {hurd.ea:theorem:2005}, + abstract = {Several tools exist for reasoning about Java programs + annotated with JML specifications. A main issue is to deal + with possible aliasing between objects and to handle + correctly the frame conditions limiting the part of memory + that a method is allowed to modify. Tools designed for + automatic use (like ESC/Java) are not complete and even not + necessarily correct. On the other side, tools which offer a + full modeling of the program require a heavy user + interaction for discharging proof obligations. In this + paper, we present the modeling of Java programs used in the + Krakatoa tool, which generates proof obligations expressed + in a logic language suitable for both automatic and + interactive reasoning. Using the Simplify automatic theorem + prover, we are able to establish automatically more + properties than static analysis tools, with a method which + is guaranteed to be sound, assuming only the correctness of + our logical interpretation of programs and + specifications.}, + doi = {10.1007/11541868_12}, + acknowledgement={brucker, 2007-02-19}, + bibkey = {marche.ea:reasoning:2005} +} + +@Proceedings{ hurd.ea:theorem:2005, + editor = {Joe Hurd and Thomas F. Melham}, + title = {Theorem Proving in Higher Order Logics (\acs{tphols})}, + booktitle = {Theorem Proving in Higher Order Logics (\acs{tphols})}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 3603, + doi = {10.1007/11541868}, + year = 2005, + isbn = {978-3-540-28372-0}, + acknowledgement={brucker, 2007-02-19} +} + +@InProceedings{ leino.ea:modular:2005, + author = {K. Rustan M. Leino and Peter M{\"u}ller}, + title = {Modular Verification of Static Class Invariants.}, + pages = {26--42}, + doi = {10.1007/11526841_4}, + abstract = {Object invariants describe the consistency of + object-oriented data structures and are central to + reasoning about the correctness of object-oriented + software. But object invariants are not the only + consistency conditions on which a program may depend. The + data in object-oriented programs consists not just of + object fields, but also of static fields, which hold data + that is shared among objects. The consistency of static + fields is described by static class invariants, which are + enforced at the class level. Static class invariants can + also mention instance fields, describing the consistency of + dynamic data structures rooted in static fields. Sometimes + there are even consistency conditions that relate the + instance fields of many or all objects of a class; static + class invariants describe these relations, too, since they + cannot be enforced by any one object in isolation. This + paper presents a systematic way (a methodology) for + specifying and verifying static class invariants in + object-oriented programs. The methodology supports the + three major uses of static fields and invariants in the + Java library. The methodology is amenable to static, + modular verification and is sound.}, + crossref = {fitzgerald.ea:fm:2005}, + acknowledgement={brucker, 2007-02-19} +} + +@InProceedings{ basin.ea:verification:2005, + author = {David A. Basin and Hironobu Kuruma and Kazuo Takaragi and + Burkhart Wolff}, + abstract = {We report on a case study in using \holz, an embedding of + Z in higher-order logic, to specify and verify a security + architecture for administering digital signatures. We have + used \holz{} to formalize and combine both data-oriented + and process-oriented architectural views. Afterwards, we + formalized temporal requirements in Z and carried out + verification in higher-order logic. The same architecture + has been previously verified using the SPIN model checker. + Based on this, we provide a detailed comparison of these + two di erent approaches to formalization (infinite state + with rich data types versus finite state) and verification + (theorem proving versus model checking). Contrary to common + belief, our case study suggests that Z is well suited for + temporal reasoning about process models with rich data. + Moreover, our comparison highlights the advantages of this + approach and provides evidence that, in the hands of + experienced users, theorem proving is neither substantially + more time-consuming nor more complex than model checking.}, + title = {Verification of a Signature Architecture with \holz}, + pages = {269--285}, + crossref = {fitzgerald.ea:fm:2005}, + language = {USenglish}, + acknowledgement={brucker, 2007-02-19}, + doi = {10.1007/11526841_19}, + bibkey = {basin.ea:verification:2005} +} + +@Proceedings{ fitzgerald.ea:fm:2005, + editor = {John Fitzgerald and Ian J. Hayes and Andrzej Tarlecki}, + booktitle = {{FM} 2005: Formal Methods}, + title = {{FM} 2005: Formal Methods}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 3582, + year = 2005, + acknowledgement={brucker, 2007-02-19}, + doi = {10.1007/11526841}, + isbn = {978-3-540-27882-5}, + location = {Newcastle, UK} +} + +@InProceedings{ barnett.ea:spec:2004, + author = {Mike Barnett and K. Rustan M. Leino and Wolfram Schulte}, + abstract = "Spec# is the latest in a long line of work on programming + languages and systems aimed at improving the development of + correct software. This paper describes the goals and + architecture of the Spec# programming system, consisting of + the object-oriented Spec# programming language, the Spec# + compiler, and the Boogie static program verifier. The + language includes constructs for writing specifications + that capture programmer intentions about how methods and + data are to be used, the compiler emits run-time checks to + enforce these specifications, and the verifier can check + the consistency between a program and its specifications.", + language = {USenglish}, + title = {The {\Specsharp} programming system: An overview}, + pages = {49--69}, + crossref = {barthe.ea:construction:2005}, + bibkey = {barnett.ea:spec:2004}, + doi = {10.1007/b105030}, + acknowledgement={brucker, 2007-02-19}, + month = may # {~25} +} + +@Proceedings{ barthe.ea:construction:2005, + editor = {Gilles Barthe and Lilian Burdy and Marieke Huisman and + Jean-Louis Lanet and Traian Muntean}, + title = {Construction and Analysis of Safe, Secure, and + Interoperable Smart Devices (\acs{cassis})}, + booktitle = {Construction and Analysis of Safe, Secure, and + Interoperable Smart Devices (\acs{cassis})}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 3362, + year = 2005, + isbn = {978-3-540-24287-1}, + acknowledgement={brucker, 2007-02-19}, + doi = {10.1007/b105030} +} + +@InProceedings{ jacobs.ea:java:2004, + author = {Bart Jacobs and Erik Poll}, + title = {{Java} Program Verification at {Nijmegen}: Developments + and Perspective.}, + doi = {10.1007/b102118}, + pages = {134--153}, + acknowledgement={brucker, 2007-02-19}, + abstract = {This paper presents a historical overview of the work on + Java program verification at the University of Nijmegen + (the Netherlands) over the past six years (1997-2003). It + describes the development and use of the LOOP tool that is + central in this work. Also, it gives a perspective on the + field.}, + crossref = {futatsugi.ea:software:2004} +} + +@Proceedings{ futatsugi.ea:software:2004, + editor = {Kokichi Futatsugi and Fumio Mizoguchi and Naoki Yonezaki}, + title = {Software Security---Theories and Systems (\acs{isss})}, + booktitle = {Software Security---Theories and Systems (\acs{isss})}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + acknowledgement={brucker, 2007-02-19}, + volume = 3233, + year = 2004, + doi = {10.1007/b102118}, + isbn = {978-3-540-23635-1} +} + +@InProceedings{ meyer.ea:architecture:2000, + author = {J{\"o}rg Meyer and Arnd Poetzsch-Heffter}, + title = {An Architecture for Interactive Program Provers}, + abstract = {Formal specification and verification techniques can + improve the quality of programs by enabling the analysis + and proof of semantic program properties. This paper + describes the modular architecture of an interactive + program prover that we are currently developing for a Java + subset. In particular, it discusses the integration of a + programming language-specific prover component with a + general purpose theorem prover.}, + pages = {63--77}, + crossref = {graf.ea:tools:2000} +} + +@Proceedings{ graf.ea:tools:2000, + editor = {Susanne Graf and Michael I. Schwartzbach}, + booktitle = {Tools and Algorithms for the Construction and Analysis of + Systems (\acs{tacas})}, + title = {Tools and Algorithms for the Construction and Analysis of + Systems (\acs{tacas})}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 1785, + year = 2000, + isbn = {3-540-67282-6} +} + +@InProceedings{ markovic.ea:ocl:2006, + author = {Sali{\v s}a Markovi{\'c} and Thomas Baar}, + language = {USenglish}, + doi = {10.1007/11880240_46}, + acknowledgement={brucker, 2007-02-19}, + pages = {661--675}, + title = {An {\acs{ocl}} Semantics Specified with {\textsc{qvt}}}, + crossref = {nierstrasz.ea:model:2006}, + abstract = {Metamodeling became in the last decade a widely accepted + tool to describe the (abstract) syntax of modeling + languages in a concise, but yet precise way. For the + description of the language's semantics, the situation is + less satisfactory and formal semantics definitions are + still seen as a challenge. In this paper, we propose an + approach to specify the semantics of modeling languages in + a graphical way. As an example, we describe the evaluation + semantics of OCL by transformation rules written in the + graphical formalism QVT. We believe that the graphical + format of our OCL semantics has natural advantages with + respect to understandability compared to existing + formalizations of OCL's semantics. Our semantics can also + be seen as a reference implementation of an OCL evaluator, + because the transformation rules can be executed by any QVT + compliant transformation engine.} +} + +@InProceedings{ pons.ea:ocl-based:2006, + author = {Claudia Pons and Diego Garcia}, + title = {An {OCL}-Based Technique for Specifying and Verifying + Refinement-Oriented Transformations in {MDE}}, + booktitle = {MoDELS}, + year = 2006, + pages = {646--660}, + doi = {10.1007/11880240_45}, + crossref = {nierstrasz.ea:model:2006} +} + +@InProceedings{ kosiuczenko:specification:2006, + author = {Piotr Kosiuczenko}, + title = {Specification of Invariability in \acs{ocl}}, + pages = {676--691}, + doi = {10.1007/11880240_47}, + crossref = {nierstrasz.ea:model:2006}, + abstract = {The paradigm of contractual specification provides a + transparent way of specifying systems. It clearly + distinguishes between client and implementer obligations. + One of the best known languages used for this purpose is + OCL. Nevertheless, OCL does not provide primitives for a + compact specification of what remains unchanged when a + method is executed. In this paper, problems with specifying + invariability are listed and some weaknesses of existing + solutions are pointed out. The question of specifying + invariability in OCL is studied and a simple but expressive + and flexible extension is proposed. It is shown that this + extension has a simple OCL based semantics.} +} + +@Proceedings{ nierstrasz.ea:model:2006, + editor = {Oscar Nierstrasz and Jon Whittle and David Harel and + Gianna Reggio}, + title = {Model Driven Engineering Languages and Systems + (\acs{models})}, + booktitle = {Model Driven Engineering Languages and Systems + (\acs{models})}, + address = pub-springer:adr, + location = {Genova, Italy}, + publisher = pub-springer, + series = s-lncs, + acknowledgement={brucker, 2007-02-19}, + volume = 4199, + year = 2006, + doi = {10.1007/11880240}, + isbn = {978-3-540-45772-5} +} + +@InProceedings{ syme:proving:1999, + author = {Don Syme}, + title = {Proving {Java} Type Soundness}, + pages = {83--118}, + crossref = {alves-foss:formal:1999}, + acknowledgement={brucker, 2007-02-19}, + abstract = {This chapter describes a machine checked proof of the type + soundness of a subset of Java (we call this subset + Javatex2html_wrap_inline102). In Chapter 3, a formal + semantics for approximately the same subset was presented + by Drossopoulou and Eisenbach. The work presented here + serves two roles: it complements the written semantics by + correcting and clarifying some details; and it demonstrates + the utility of formal, machine checking when exploring a + large and detailed proof based on operational semantics.}, + bibkey = {syme:proving:1999} +} + +@InProceedings{ flatt.ea:programmers:1999, + author = {Matthew Flatt and Shriram Krishnamurthi and Matthias + Felleisen}, + title = {A Programmer's Reduction Semantics for Classes and + Mixins.}, + doi = {10.1007/3-540-48737-9_7}, + pages = {241--269}, + crossref = {alves-foss:formal:1999}, + acknowledgement={brucker, 2007-02-19}, + abstract = {While class-based object-oriented programming languages + provide a flexible mechanism for re-using and managing + related pieces of code, they typically lack linguistic + facilities for specifying a uniform extension of many + classes with one set of fields and methods. As a result, + programmers are unable to express certain abstractions over + classes. In this paper we develop a model of class-to-class + functions that we refer to as mixins. A mixin function maps + a class to an extended class by adding or overriding fields + and methods. Programming with mixins is similar to + programming with single inheritance classes, but mixins + more directly encourage programming to interfaces. The + paper develops these ideas within the context of Java. The + results are an intuitive model of an essential Java subset; + an extension that explains and models mixins; and type + soundness theorems for these languages.} +} + +@InProceedings{ drossopoulou.ea:describing:1999, + author = {Sophia Drossopoulou and Susan Eisenbach}, + title = {Describing the Semantics of {Java} and Proving Type + Soundness}, + pages = {41--82}, + doi = {10.1007/3-540-48737-9_2}, + crossref = {alves-foss:formal:1999}, + acknowledgement={brucker, 2007-02-19}, + abstract = {Java combines the experience from the development of + several object oriented languages, such as C++, Smalltalk + and CLOS. The philosophy of the language designers was to + include only features with already known semantics, and to + provide a small and simple language. + + Nevertheless, we feel that the introduction of some new + features in Java, as well as the specific combination of + features, justifies a study of the Java formal semantics. + The use of interfaces, reminiscent of [6,10] is a + simplification of the signatures extension for C++ [4] and + is - to the best of our knowledge - novel. The mechanism + for dynamic method binding is that of C++, but we know of + no formal definition. Java adopts the Smalltalk [15] + approach whereby all object variables are implicitly + pointers. + + Furthermore, although there are a large number of studies + of the semantics of isolated programming language features + or of minimal programming languages [1,31,34] there have + not been many studies of the formal semantics of actual + programming languages. In addition, the interplay of + features which are very well understood in isolation, might + introduce unexpected effects. } +} + +@InProceedings{ oheimb.ea:machine-checking:1999, + author = {David von Oheimb and Tobias Nipkow}, + title = {Machine-Checking the {Java} Specification: Proving + Type-Safety}, + pages = {119--156}, + crossref = {alves-foss:formal:1999}, + acknowledgement={brucker, 2007-02-19}, + bibkey = {oheimb.ea:machine-checking:1999} +} + +@Proceedings{ alves-foss:formal:1999, + editor = {Jim Alves-Foss}, + title = {Formal Syntax and Semantics of {Java}}, + booktitle = {Formal Syntax and Semantics of {Java}}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 1523, + year = 1999, + isbn = {3-540-66158-1}, + acknowledgement={brucker, 2007-02-19}, + bibkey = {alves-foss:formal:1999} +} + +@InProceedings{ smith.ea:encoding:2002, + author = {Graeme Smith and Florian Kamm{\"u}ller and Thomas Santen}, + title = {Encoding {Object-Z} in {Isabelle}/{\acs{hol}}.}, + pages = {82--99}, + doi = {10.1007/3-540-45648-1_5}, + crossref = {bert.ea:zb:2002}, + abstract = {In this paper, we present a formalization of the reference + semantics of Object-Z in the higher-order logic (HOL) + instantiation of the generic theorem prover Isabelle, + Isabelle/HOL. This formalization has the effect of both + clarifying the semantics and providing the basis for a + theorem prover for Object-Z. The work builds on an earlier + encoding of a value semantics for object-oriented Z in + Isabelle/HOL and a denotational semantics of Object-Z based + on separating the internal and external effects of class + methods.}, + keywords = {Object-Z, reference semantics, higher-order logic, + Isabelle}, + acknowledgement={brucker, 2007-02-19} +} + +@Proceedings{ bert.ea:zb:2002, + editor = {Didier Bert and Jonathan P. Bowen and Martin C. Henson and + Ken Robinson}, + title = {{ZB} 2002: Formal Specification and Development in {Z} and + {B}}, + booktitle = {{ZB} 2002: Formal Specification and Development in {Z} and + {B}}, + location = {Grenoble, France}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 2272, + year = 2002, + isbn = {3-540-43166-7}, + acknowledgement={brucker, 2007-02-19}, + bibkey = {bert.ea:zb:2002} +} + +@InProceedings{ paulson:formulation:1990, + author = {Lawrence C. Paulson}, + title = {A formulation of the simple theory of types (for + {Isabelle})}, + pages = {246--274}, + doi = {10.1007/3-540-52335-9_58}, + crossref = {martin-lof.ea:colog-88:1990}, + acknowledgement={brucker, 2007-02-19}, + abstract = {Simple type theory is formulated for use with the generic + theorem prover Isabelle. This requires explicit type + inference rules. There are function, product, and subset + types, which may be empty. Descriptions (the eta-operator) + introduce the Axiom of Choice. Higher-order logic is + obtained through reflection between formulae and terms of + type bool. Recursive types and functions can be formally + constructed. Isabelle proof procedures are described. The + logic appears suitable for general mathematics as well as + computational problems. } +} + +@Proceedings{ martin-lof.ea:colog-88:1990, + editor = {Per Martin-L{\"o}f and Grigori Mints}, + title = {\acs{colog}-88}, + booktitle = {\acs{colog}-88}, + location = {Tallinn, USSR}, + publisher = pub-springer, + acknowledgement={brucker, 2007-02-19}, + address = pub-springer:adr, + doi = {10.1007/3-540-52335-9}, + series = s-lncs, + volume = 417, + year = 1990, + isbn = {3-540-52335-9} +} + +@InProceedings{ beckert.ea:dynamic:2006, + author = {Bernhard Beckert and Andr{\'e} Platzer}, + title = {Dynamic Logic with Non-rigid Functions.}, + pages = {266--280}, + acknowledgement={brucker, 2007-02-19}, + doi = {10.1007/11814771_23}, + abstract = {We introduce a dynamic logic that is enriched by non-rigid + functions, i.e., functions that may change their value from + state to state (during program execution), and we present a + (relatively) complete sequent calculus for this logic. In + conjunction with dynamically typed object enumerators, + non-rigid functions allow to embed notions of + object-orientation in dynamic logic, thereby forming a + basis for verification of object-oriented programs. A + semantical generalisation of substitutions, called state + update, which we add to the logic, constitutes the central + technical device for dealing with object aliasing during + function modification. With these few extensions, our + dynamic logic captures the essential aspects of the complex + verification system KeY and, hence, constitutes a + foundation for object-oriented verification with the + principles of reasoning that underly the successful KeY + case studies.}, + crossref = {furbach.ea:automated:2006} +} + +@Proceedings{ furbach.ea:automated:2006, + editor = {Ulrich Furbach and Natarajan Shankar}, + doi = {10.1007/11814771}, + title = {Automated Reasoning (\acs{ijcar})}, + booktitle = {Automated Reasoning (\acs{ijcar})}, + location = {Seattle, WA}, + publisher = pub-springer, + address = pub-springer:adr, + acknowledgement={brucker, 2007-02-19}, + series = s-lncs, + volume = 4130, + year = 2006, + isbn = {978-3-540-37187-8} +} + +@InProceedings{ yatake.ea:implementing:2005, + author = {Kenro Yatake and Toshiaki Aoki and Takuya Katayama}, + title = {Implementing Application-Specific Object-Oriented Theories + in {\acs{hol}}}, + acknowledgement={brucker, 2007-02-19}, + doi = {10.1007/11560647_33}, + pages = {501--516}, + abstract = {This paper presents a theory of Object-Oriented concepts + embedded shallowly in HOL for the verification of OO + analysis models. The theory is application-specific in the + sense that it is automatically constructed depending on the + type information of the application. This allows objects to + have attributes of arbitrary types, making it possible to + verify models using not only basic types but also highly + abstracted types specific to the target domain. The theory + is constructed by definitional extension based on the + operational semantics of a heap memory model, which + guarantees the soundness of the theory. This paper mainly + focuses on the implementation details of the theory.}, + crossref = {hung.ea:theoretical:2005} +} + +@Proceedings{ hung.ea:theoretical:2005, + editor = {Dang Van Hung and Martin Wirsing}, + title = {Theoretical Aspects of Computing---\acs{ictac} 2005}, + booktitle = {Theoretical Aspects of Computing---\acs{ictac} 2005}, + location = {Hanoi, Vietnam}, + publisher = pub-springer, + address = pub-springer:adr, + doi = {10.1007/11560647}, + series = s-lncs, + acknowledgement={brucker, 2007-02-19}, + volume = 3722, + year = 2005, + isbn = {3-540-29107-5} +} + +@InProceedings{ aspinall:proof:2000, + author = {David Aspinall}, + title = {{P}roof {G}eneral: A Generic Tool for Proof Development}, + acknowledgement={brucker, 2007-02-19}, + pages = {38--42}, + crossref = {graf.ea:tools:2000-b}, + abstract = {This note describes Proof General, a tool for developing + machine proofs with an interactive proof assistant. + Interaction is based around a proof script, which is the + target of a proof development. Proof General provides a + powerful user-interface with relatively little effort, + alleviating the need for a proof assistant to provide its + own GUI, and providing a uniform appearance for diverse + proof assistants. + + Proof General has a growing user base and is currently used + for several interactive proof systems, including Coq, LEGO, + and Isabelle. Support for others is on the way. Here we + give a brief overview of what Proof General does and the + philosophy behind it; technical details are available + elsewhere. The program and user documentation are available + on the web at http://www.dcs.ed.ac.uk/home/proofgen.} +} + +@InProceedings{ beckert.ea:many-valued:1992, + author = {Bernhard Beckert and Stefan Gerberding and Reiner + H{\"a}hnle and Werner Kernig}, + title = {The Many-Valued Tableau-Based Theorem Prover {\threeTAP}}, + acknowledgement={brucker, 2007-02-19}, + pages = {758--760}, + bibkey = {beckert.ea:many-valued:1992}, + crossref = {kapur:automated:1992}, + doi = {10.1007/3-540-55602-8_219} +} + +@Proceedings{ kapur:automated:1992, + editor = {Deepak Kapur}, + title = {Automated Deduction---\acs{cade}-11}, + booktitle = {Automated Deduction---\acs{cade}-11}, + location = {Saratoga Springs, \acs{ny}, \acs{usa}}, + publisher = pub-springer, + address = pub-springer:adr, + doi = {10.1007/3-540-55602-8}, + series = s-lncs, + volume = 607, + year = 1992, + acknowledgement={brucker, 2007-02-19}, + isbn = {978-3-540-55602-2} +} + +@Article{ ahrendt.ea:key:2005, + bibkey = {ahrendt.ea:key:2005}, + author = {Wolfgang Ahrendt and Thomas Baar and Bernhard Beckert and + Richard Bubel and Martin Giese and Reiner H\"ahnle and + Wolfram Menzel and Wojciech Mostowski and Andreas Roth and + Steffen Schlager and Peter H. Schmitt}, + title = {The {\KeY} Tool}, + doi = {10.1007/s10270-004-0058-x}, + publisher = pub-springer, + address = pub-springer:adr, + journal = j-sosym, + volume = 4, + number = 1, + year = 2005, + pages = {32--54}, + acknowledgement={brucker, 2007-04-23} +} + +@Article{ cengarle.ea:ocl:2004, + journal = j-sosym, + volume = 3, + pages = {9--30}, + number = 1, + year = 2004, + publisher = pub-springer, + address = pub-springer:adr, + issn = {1619-1366}, + doi = {10.1007/s10270-003-0035-9}, + title = {{\acs{ocl}} 1.4/5 vs. 2.0 Expressions Formal semantics and + expressiveness}, + author = {Mar{\`\i}a Victoria Cengarle and Alexander Knapp}, + acknowledgement={brucker, 2007-04-23}, + bibkey = {cengarle.ea:ocl:2004} +} + +@Article{ toval.ea:emerging:2003, + journal = j-sosym, + pages = {248--261}, + volume = 2, + number = 4, + year = 2003, + publisher = pub-springer, + doi = {10.1007/s10270-003-0031-0}, + address = pub-springer:adr, + issn = {1619-1366}, + month = dec, + title = {Emerging {\acs{ocl}} tools}, + author = {Jos{\'e} Ambrosio Toval and V{\`\i}ctor Requena and + Jos{\'e} Luis Fern{\'a}ndez}, + acknowledgement={brucker, 2007-04-23}, + bibkey = {toval.ea:emerging:2003} +} + +@Article{ bubel.ea:formal:2005, + author = {Richard Bubel and Reiner H\"{a}hnle}, + title = {Integration of informal and formal development of + object-oriented safety-critical software.}, + year = 2005, + journal = j-sttt, + publisher = pub-springer, + address = pub-springer:adr, + issn = {1433-2779}, + volume = 7, + number = 3, + language = {USenglish}, + doi = {10.1007/s10009-004-0166-5}, + pages = {197--211}, + acknowledgement={brucker, 2007-04-23}, + bibkey = {bubel.ea:formal:2005} +} + +@Proceedings{ stepney.ea:object:1992, + abstract = {This collection of papers draws together a variety of + approaches for adding OO concepts and structuring + capability to the Z formal specification language. Each + approach is used to specify the same two problems, to allow + a comparison. }, + editor = {Susan Stepney and Rosalind Barden and David Cooper}, + isbn = {3-540-19778-8}, + language = {USenglish}, + public = {yes}, + publisher = pub-springer, + address = pub-springer:adr, + series = {Workshops in Computing}, + topic = {formalism}, + title = {Object Orientation in {Z}}, + year = 1992, + keywords = {Object Orientation, Z}, + acknowledgement={brucker, 2007-04-23}, + bibkey = {stepney.ea:object:1992} +} + +@InProceedings{ hamie.ea:reflections:1998, + bibkey = {hamie.ea:reflections:1998}, + author = {Ali Hamie and Franco Civello and John Howse and Stuart + Kent and Richard Mitchell}, + title = {{Reflections on the Object Constraint Language}}, + year = 1998, + doi = {10.1007/b72309}, + topic = {formalism}, + acknowledgement={brucker, 2007-04-23}, + pages = {162--172}, + crossref = {bezivin.ea:unified:1999}, + abstract = {The \acf{ocl}, which forms part of the \acs{uml} set of + modelling notations, is a precise, textual language for + expressing constraints that cannot be shown + diagrammatically in \acs{uml}. This paper reflects on a + number of aspects of the syntax and semantics of the + \acs{ocl}, and makes proposals for clarification or + extension. Specifically, the paper suggests that: the + concept of flattening collections of collections is + unnecessary, state models should be connectable to class + models, defining object creation should be made more + convenient, \acs{ocl} should be based on a 2-valued logic, + set subtraction should be covered more fully, and a "let" + feature should be introduced. } +} + +@Proceedings{ bezivin.ea:unified:1999, + editor = {Jean B{\'e}zivin and Pierre-Alain Muller}, + doi = {10.1007/b72309}, + booktitle = {The Unified Modeling Language. \guillemotleft + \acs{uml}\guillemotright'98: Beyond the Notation}, + title = {The Unified Modeling Language. \guillemotleft + \acs{uml}\guillemotright'98: Beyond the Notation}, + publisher = pub-springer, + address = pub-springer:adr, + acknowledgement={brucker, 2007-04-23}, + series = s-lncs, + volume = 1618, + year = 1999, + isbn = {3-540-66252-9} +} + +@Book{ guttag.ea:larch:1993, + author = {John V. Guttag and James J. Horning}, + title = {{Larch}: Languages and Tools for Formal Specification}, + publisher = pub-springer-ny, + address = pub-springer-ny:adr, + series = {Texts and Monographs in Computer Science}, + year = 1993, + isbn = {0-387-94006-5}, + acknowledgement={brucker, 2007-04-23} +} + +@Article{ beckert.ea:refinement:2005, + title = {Refinement and Retrenchment for Programming Language Data + Types}, + author = {Bernhard Beckert and Steffen Schlager}, + journal = j-fac, + volume = 17, + number = 4, + acknowledgement={brucker, 2007-04-23}, + pages = {423--442}, + year = 2005, + doi = {10.1007/s00165-005-0073-x}, + publisher = pub-springer, + address = pub-springer:adr +} + +@Article{ nipkow:winskel:1998, + author = {Tobias Nipkow}, + title = {Winskel is (almost) Right: Towards a Mechanized Semantics + Textbook}, + publisher = pub-springer, + journal = j-fac, + volume = 10, + number = 2, + doi = {10.1007/s001650050009}, + pages = {171--186}, + abstract = {We present a formalization of the first 100 pages of + Winskel's textbook `The Formal Semantics of Programming + Languages' in the theorem prover Isabelle/\acs{hol}: 2 + operational, 2 denotational, 2 axiomatic semantics, a + verification condition generator, and the necessary + soundness, completeness and equivalence proofs, all for a + simple imperative programming language.}, + acknowledgement={brucker, 2007-04-23}, + year = 1998, + bibkey = {nipkow:winskel:1998} +} + +@InCollection{ dupuy.ea:using:2000, + author = {Sophie Dupuy and Ang{\`e}s Front-Conte and Christophe + Saint-Marcel}, + chapter = 6, + title = {Using \acs{uml} with a Behaviour-Driven Method}, + page = {97--112}, + acknowledgement={brucker, 2007-04-23}, + crossref = {frappier.ea:software:2000} +} + +@Book{ frappier.ea:software:2000, + editor = {Marc Frappier and Henri Habrias}, + title = {Software Specification Methods: An Overview Using a Case + Study}, + acknowledgement={brucker, 2007-04-23}, + publisher = pub-springer-london, + address = pub-springer-london:adr, + year = 2000, + isbn = {1-85233-353-7}, + series = {Formal Approaches to Computing and Information + Technology} +} + +@InProceedings{ hamie.ea:interpreting:1998, + bibkey = {hamie.ea:interpreting:1998}, + author = {Ali Hamie and John Howse and Stuart Kent}, + title = {Interpreting the {Object Constraint Language}}, + abstract = {The \acf{ocl}, which forms part of the \acs{uml} 1.1. set + of modelling notations is a precise, textual language for + expressing constraints that cannot be shown in the standard + diagrammatic notation used in \acs{uml}. A semantics for + \acs{ocl} lays the foundation for building CASE tools that + support integrity checking of the whole \acs{uml} models, + not just the component expressed using \acs{ocl}. This + paper provides a semantics for \acs{ocl}, at the same time + providing a semantics for classes, associations, attributes + and states. }, + pages = {288--295}, + ee = {http://csdl.computer.org/comp/proceedings/apsec/1998/9183/00/91830288abs.htm} + , + doi = {10.1109/apsec.1998.733731}, + keywords = {OCL}, + acknowledgement={brucker, 2007-04-23}, + topic = {formalism}, + crossref = {ieee:apsec:1998} +} + +@Proceedings{ ieee:apsec:1998, + bibkey = {ieee:apsec:1998}, + booktitle = PROC # { Asia Pacific Conference in Software + Engineering (\acs{apsec})}, + title = PROC # { Asia Pacific Conference in Software + Engineering (\acs{apsec})}, + publisher = pub-ieee, + address = pub-ieee:adr, + acknowledgement={brucker, 2007-04-23}, + year = 1998, + isbn = {0-8186-9183-2} +} + +@InProceedings{ mandel.ea:ocl:1999, + author = {Luis Mandel and Mar{\`i}a Victoria Cengarle}, + bibkey = {mandel.ea:ocl:1999}, + language = {USenglish}, + topic = {formalism}, + public = {yes}, + title = {On the expressive power of {\acs{ocl}}}, + acknowledgement={brucker, 2007-04-23}, + timestamp = 962971498, + abstract = {This paper examines the expressive power of \acs{ocl} in + terms of navigability and computability. First the + expressive power of \acs{ocl} is compared with the + relational calculus; it is showed that \acs{ocl} is not + equivalent to the relational calculus. Then an algorithm + computing the transitive closure of a binary relation + operation that cannot be encoded in the relational calculus + is expressed in \acs{ocl}. Finally the equivalence of + \acs{ocl} with a Turing machine is pondered.}, + pages = {854--874}, + crossref = {wing.ea:world:1999}, + ee = {http://link.springer.de/link/service/series/0558/bibs/1708/17080854.htm} + +} + +@Proceedings{ wing.ea:world:1999, + editor = {Jeannette M. Wing and Jim Woodcock and Jim Davies}, + booktitle = {World Congress on Formal Methods in the Development of + Computing Systems (FM)}, + title = {World Congress on Formal Methods in the Development of + Computing Systems (FM)}, + publisher = pub-springer, + address = pub-springer:adr, + acknowledgement={brucker, 2007-04-23}, + series = s-lncs, + volume = 1708, + year = 1999, + isbn = {3-540-66587-0} +} + +@Book{ spivey:z-notation:1992, + bibkey = {spivey:z-notation:1992}, + author = {J. M. Spivey}, + title = {The {Z} Notation: A Reference Manual}, + publisher = pub-prentice, + address = pub-prentice:adr, + edition = {2nd}, + length = 150, + year = 1992, + isbn = {0-139-78529-9}, + acknowledgement={brucker, 2007-04-23}, + abstract = {This is a revised edition of the first widely available + reference manual on Z originally published in 1989. The + book provides a complete and definitive guide to the use of + Z in specifying information systems, writing specifications + and designing implementations. \par Contents: Tutorial + introduction; Background; The Z language; The mathematical + tool-kit; Sequential systems; Syntax summary; Changes from + the first edition; Glossary.} +} + +@Book{ jones:vdm:1990, + bibkey = {jones:vdm:1990}, + author = {Cliff B.\ Jones}, + title = {Systematic Software Development Using \acs{vdm}}, + publisher = pub-prentice, + address = pub-prentice:adr, + year = 1990, + size = 333, + edition = {2nd}, + note = {0-13-880733-7}, + abstract = {This book deals with the Vienna Development Method. The + approach explains formal (functional) specifications and + verified design with an emphasis on the study of proofs in + the development process.}, + acknowledgement={brucker, 2007-04-23} +} + +@Article{ liskov.ea:behavioral:1994, + bibkey = {liskov.ea:behavioral:1994}, + abstract = {The use of hierarchy is an important component of + object-oriented design.Hierarchy allows the use of type + families, in whichhigher level supertypes capture the + behavior that all of their subtypes havein common. For this + methodology to be effective,it is necessary to have a clear + understanding of how subtypes and supertypesare related. + This paper takes the position thatthe relationship should + ensure that any property proved about supertypeobjects also + holds for its subtype objects. It presentstwo ways of + defining the subtype relation, each of which meets this + criterion,and each of which is easy for programmers touse. + The subtype relation is based on the specifications of the + sub- and supertypes; the paper presents a way of + specifyingtypes that makes it convenient to define the + subtype relation. The paper alsodiscusses the ramifications + of this notion ofsubtyping on the design of type + families.}, + author = {Barbara H. Liskov and Jeannette M. Wing}, + journal = j-toplas, + month = nov, + pages = {1811--1841}, + issn = {0164-0925}, + keywords = {languages, verficiation}, + language = {USenglish}, + number = 6, + publisher = pub-acm, + address = pub-acm:adr, + doi = {10.1145/197320.197383}, + public = {yes}, + title = {A behavioral notion of subtyping}, + volume = 16, + year = 1994, + acknowledgement={brucker, 2007-04-23} +} + +@Book{ winskel:semantics:1993, + bibkey = {winskel:semantics:1993}, + author = {Glynn Winskel}, + title = {The Formal Semantics of Programming Languages}, + publisher = pub-mit, + address = pub-mit:adr, + isbn = {0-262-23169-7}, + pages = 384, + year = 1993, + acknowledgement={brucker, 2007-04-23} +} + +@Book{ andrews:introduction:2002, + author = {Peter B. Andrews}, + title = {Introduction to Mathematical Logic and Type Theory: To + Truth through Proof}, + year = 2002, + isbn = {1-402-00763-9}, + edition = {2nd}, + publisher = pub-kluwer, + address = pub-kluwer:adr, + acknowledgement={brucker, 2007-04-23}, + bibkey = {andrews:introduction:2002} +} + +@PhDThesis{ santen:mechanized:1999, + author = {Thomas Santen}, + title = {A Mechanized Logical Model of {Z} and Object-Oriented + Specification}, + school = {Technical University Berlin}, + year = 1999, + month = jun, + annote = {Also available as book: Shaker Verlag, Aachen. ISBN: + 3826576500}, + bibkey = {santen:mechanized:1999}, + acknowledgement={brucker, 2007-04-23} +} + +@Book{ kleene:mathematics:1971, + bibkey = {kleene:mathematics:1971}, + author = {Stephen C. Kleene}, + title = {Introduction to Meta Mathematics}, + publisher = {Wolters-Noord\-hoff Publishing}, + address = {Amsterdam}, + isbn = {0-7204-2103-9}, + year = 1971, + note = {Originally published by Van Nostrand, 1952}, + acknowledgement={brucker, 2007-04-23} +} + +@Book{ gordon.ea:hol:1993, + bibkey = {gordon.ea:hol:1993}, + author = {Mike J. C. Gordon and Tom F. Melham}, + title = {Introduction to \acs{hol}: a theorem proving environment + for higher order logic}, + publisher = pub-cup, + address = pub-cup:adr, + year = 1993, + pages = 472, + isbn = {0-521-44189-7}, + month = jul, + abstract = {Currently being applied to a wide variety of problems, + Higher-Order Logic (\acs{hol}) is a proof development + system intended for applications to both hardware and + software. This self-contained description contains a + tutorial introduction and most of the material needed for + day-to-day work.}, + acknowledgement={brucker, 2007-04-23} +} + +@PhDThesis{ richters:precise:2002, + author = {Mark Richters}, + title = {A Precise Approach to Validating {\acs{uml}} Models and + {\acs{ocl}} Constraints}, + school = {Universit{\"a}t Bremen}, + year = 2002, + address = {Logos Verlag, Berlin, \acs{biss} Monographs, No. 14}, + isbn = {3-89722-842-4}, + abstract = {We present a precise approach that allows an analysis and + validation of \acs{uml} models and OCL constraints. We + focus on models and constraints specified in the analysis + and early design stage of a software development process. + For this purpose, a suitable subset of \acs{uml} + corresponding to information that is usually represented in + class diagrams is identified and formally defined. This + basic modeling language provides a context for all OCL + constraints. We define a formal syntax and semantics of OCL + types, operations, expressions, invariants, and + pre-/postconditions. We also give solutions for problems + with the current OCL definition and discuss possible + extensions. A metamodel for OCL is introduced that defines + the abstract syntax of OCL expressions and the structure of + types and values. The metamodel approach allows a seamless + integration with the \acs{uml} metamodeling architecture + and makes the benefits of a precise OCL definition easier + accessible. The OCL metamodel also allows to define + context-sensitive conditions for well-formed OCL + expressions more precisely. These conditions can now be + specified with OCL whereas they previously were specified + only informally. In order to demonstrate the practical + applicability of our work, we have realized substantial + parts of it in a tool supporting the validation of models + and constraints. Design specifications can be ``executed'' + and animated thus providing early feedback in an iterative + development process. Our approach offers novel ways for + checking user data against specifications, for automating + test procedures, and for checking CASE tools for standards + conformance. Therefore, this work contributes to the goal + of improving the overall quality of software systems by + combining theoretical and practical techniques.}, + acknowledgement={brucker, 2007-04-23} +} + +@Booklet{ cc:cc-part3:2006, + bibkey = {cc:cc-part3:2006}, + key = {Common Criteria}, + institution = {Common Criteria}, + language = {USenglish}, + month = sep, + year = 2006, + public = {yes}, + title = {Common Criteria for Information Technology Security + Evaluation (Version 3.1), {Part} 3: Security assurance + components}, + note = {Available as document + \href{http://www.commoncriteriaportal.org/public/files/CCPART3V3.1R1.pdf} + {CCMB-2006-09-003}}, + number = {CCMB-2006-09-003}, + acknowledgement={brucker, 2007-04-24} +} + +@Booklet{ omg:ocl:1997, + bibkey = {omg:ocl:1997}, + key = omg, + abstract = {This document introduces and defines the Object Constraint + Language (\acs{ocl}), a formal language to express side + effect-free constraints. Users of the Unified Modeling + Language and other languages can use \acs{ocl} to specify + constraints and other expressions attached to their models. + \acs{ocl} was used in the \acs{uml} Semantics document to + specify the well-formedness rules of the \acs{uml} + metamodel. Each well-formedness rule in the static + semantics sections in the \acs{uml} Semantics document + contains an \acs{ocl} expression, which is an invariant for + the involved class. The grammar for \acs{ocl} is specified + at the end of this document. A parser generated from this + grammar has correctly parsed all the constraints in the + \acs{uml} Semantics document, a process which improved the + correctness of the specifications for \acs{ocl} and \acs{uml}.}, + institution = omg, + language = {USenglish}, + month = sep, + note = {Available as \acs{omg} document + \href{http://www.omg.org/cgi-bin/doc?ad/97-08-08} + {ad/97-08-08}}, + keywords = {\acs{uml}, OCL}, + topic = {formalism}, + public = {yes}, + title = {Object Constraint Language Specification (Version 1.1)}, + year = 1997, + acknowledgement={brucker, 2007-04-23} +} + +@Booklet{ omg:xmi:2000, + bibkey = {omg:xmi:2000}, + key = omg, + abstract = {The main purpose of XMI is to enable easy interchange of + metadata between modeling tools (based on the + \acs{omg}-\acs{uml}) and metadata repositories + (\acs{omg}-MOF based) in distributed heterogeneous + environments. XMI integrates three key industry standards: + XML, \acs{uml}, MOF.}, + publisher = omg, + language = {USenglish}, + month = nov, + year = 2000, + keywords = {\acs{uml}, XML, XMI}, + topic = {formalism}, + public = {yes}, + title = {\acs{omg} \acs{xml} Metadata Interchange (\acs{xmi}) + Specification (Version 1.1)}, + note = {Available as \acs{omg} document + \href{http://www.omg.org/cgi-bin/doc?formal/00-11-02} + {formal/00-11-02}}, + acknowledgement={brucker, 2007-04-23} +} + +@Booklet{ omg:ocl:2003, + bibkey = {omg:ocl:2003}, + key = omg, + abstract = {This document introduces and defines the Object Constraint + Language (OCL), a formal language to express side + effect-free constraints. Users of the Unified Modeling + Language and other languages can use OCL to specify + constraints and other expressions attached to their models. + OCL was used in the \acs{uml} Semantics document to specify + the well-formedness rules of the \acs{uml} metamodel. Each + well-formedness rule in the static semantics sections in + the \acs{uml} Semantics document contains an OCL + expression, which is an invariant for the involved class. + The grammar for OCL is specified at the end of this + document. A parser generated from this grammar has + correctly parsed all the constraints in the \acs{uml} + Semantics document, a process which improved the + correctness of the specifications for OCL and \acs{uml}.}, + publisher = omg, + language = {USenglish}, + month = oct, + keywords = {\acs{uml}, OCL}, + topic = {formalism}, + public = {yes}, + note = {Available as \acs{omg} document + \href{http://www.omg.org/cgi-bin/doc?ptc/03-10-14} + {ptc/03-10-14}}, + title = {\acs{uml} 2.0 \acs{ocl} Specification}, + year = 2003, + acknowledgement={brucker, 2007-04-23} +} + +@Booklet{ omg:uml:2003, + bibkey = {omg:uml:2003}, + key = omg, + abstract = {The Unified Modeling Language (\acs{uml}) provides system + architects working on object analysis and design with one + consistent language for specifying, visualizing, + constructing, and documenting the artifacts of software + systems, as well as for business modeling.This + specification represents the convergence of best practices + in the object-technology industry. \acs{uml} is the proper + successor to the object modeling languages of three + previouslyleading object-oriented methods (Booch, OMT, and + OOSE). The \acs{uml} is the union of thesemodeling + languages and more, since it includes additional + expressiveness to handle modelingproblems that these + methods did not fully address.One of the primary goals of + \acs{uml} is to advance the state of the industry by + enabling objectvisual modeling tool interoperability. + However, in order to enable meaningful exchange ofmodel + information between tools, agreement on semantics and + notation is required. \acs{uml} meets the following + requirements: \begin{enumerate} \item Formal definition of + a common object analysis and design (OA\&D) metamodel to + representthe semantics of OA\&D models, which include + static models, behavioral models, usagemodels, and + architectural models. \item IDL specifications for + mechanisms for model interchange between OA\&D tools. + Thisdocument includes a set of IDL interfaces that support + dynamic construction and traversal ofa user model. \item + readable notation for representing OA\&D models. + \end{enumerate} This document defines the \acs{uml} + notation, an elegant graphic syntax for consistently + expressing the \acs{uml}'s richsemantics. Notation is an + essential part of OA\&D modeling and the \acs{uml}.}, + publisher = omg, + language = {USenglish}, + month = mar, + year = 2003, + note = {Available as \acs{omg} document + \href{http://www.omg.org/cgi-bin/doc?formal/03-03-01} + {formal/03-03-01}}, + keywords = {\acs{uml}, OCL}, + topic = {formalism}, + public = {yes}, + title = {Unified Modeling Language Specification (Version 1.5)}, + acknowledgement={brucker, 2007-04-23} +} + +@Booklet{ omg:ocl:2006, + bibkey = {omg:ocl:2006}, + key = omg, + abstract = {This document introduces and defines the Object Constraint + Language (OCL), a formal language to express side + effect-free constraints. Users of the Unified Modeling + Language and other languages can use OCL to specify + constraints and other expressions attached to their models. + OCL was used in the \acs{uml} Semantics document to specify + the well-formedness rules of the \acs{uml} metamodel. Each + well-formedness rule in the static semantics sections in + the \acs{uml} Semantics document contains an OCL + expression, which is an invariant for the involved class. + The grammar for OCL is specified at the end of this + document. A parser generated from this grammar has + correctly parsed all the constraints in the \acs{uml} + Semantics document, a process which improved the + correctness of the specifications for OCL and \acs{uml}.}, + publisher = omg, + language = {USenglish}, + month = apr, + keywords = {\acs{uml}, OCL}, + topic = {formalism}, + note = {Available as \acs{omg} document + \href{http://www.omg.org/cgi-bin/doc?formal/06-05-01} + {formal/06-05-01}}, + public = {yes}, + title = {\acs{uml} 2.0 \acs{ocl} Specification}, + year = 2006, + acknowledgement={brucker, 2007-04-23} +} + +@Booklet{ omg:uml:2005, + bibkey = {omg:uml:2005}, + key = omg, + publisher = omg, + language = {USenglish}, + note = {Available as \acs{omg} document + \href{http://www.omg.org/cgi-bin/doc?formal/05-07-04} + {formal/05-07-04}}, + keywords = {\acs{uml}}, + topic = {formalism}, + public = {yes}, + title = {\acs{uml} 2.0 Superstructure Specification}, + year = 2005, + month = jul, + acknowledgement={brucker, 2007-04-23} +} + +@PhDThesis{ oheimb:analyzing:2001, + author = {David von Oheimb}, + title = {Analyzing {J}ava in {Isabelle/\acs{hol}}: Formalization, + Type Safety and {H}oare Logic}, + school = {Technische Universit\"{a}t M\"{u}nchen}, + year = 2001, + crclassification={D.2.4, D.3.1, F.3.1}, + crgenterms = {Languages, Verification, Theory}, + keywords = {Java, formalization, operational semantics, type + soundness, axiomatic semantics, Isabelle/HOL}, + abstract = {This thesis deals with machine-checking a large + sublanguage of sequential Java, covering nearly all + features, in particular the object-oriented ones. It shows + that embedding such a language in a theorem prover and + deducing practically important properties is meanwhile + possible and explains in detail how this can be achieved. + We formalize the abstract syntax, and the static semantics + including the type system and well-formedness conditions, + as well as an operational (evaluation) semantics of the + language. Based on these definitions, we can express + soundness of the type system, an important design goal + claimed to be reached by the designers of Java, and prove + that type safety holds indeed. Moreover, we give an + axiomatic semantics of partial correctness for both + statements and (side-effecting) expressions. We prove the + soundness of this semantics relative to the operational + semantics, and even prove completeness. We further give a + small but instructive application example. A direct outcome + of this work is the confirmation that the design and + specification of Java (or at least the subset considered) + is reasonable, yet some omissions in the language + specification and possibilities for generalizing the design + can be pointed out. The second main contribution is a sound + and complete Hoare logic, where the soundness proof for our + Hoare logic gives new insights into the role of type + safety. To our knowledge, this logic is the first one for + an object-oriented language that has been proved complete. + By-products of this work are a new general technique for + handling side-effecting expressions and their results, the + discovery of the strongest possible rule of consequence, + and a new rule for flexible handling of mutual recursion. + All definitions and proofs have been done fully formally + with the interactive theorem prover Isabelle/HOL, + representing one of its major applications. This approach + guarantees not only rigorous definitions, but also gives + maximal confidence in the results obtained. Thus this + thesis demonstrates that machine-checking the design of an + important non-trivial programming language and conducting + meta-theory on it entirely within a theorem proving system + has become a reality. }, + acknowledgement={brucker, 2007-04-23}, + bibkey = {oheimb:analyzing:2001} +} + +@Article{ chiarada.ea:improving:2006, + bibkey = {chiarada.ea:improving:2006}, + language = {USenglish}, + public = {yes}, + title = {Improving the {\acs{ocl}} Semantics Definition by Applying + Dynamic Meta Modeling and Design Patterns}, + author = {Juan Mart{\'\i}n Chiarad{\'\i}a and Claudia Pons}, + editor = {Birgith Demuth and Dan Chiorean and Martin Gogolla and Jos + Warmer}, + issn = {1863-2122}, + volume = 5, + year = 2006, + acknowledgement={brucker, 2007-04-23}, + journal = j-eceasst +} + +@PhDThesis{ schirmer:verification:2006, + author = {Norbert Schirmer}, + title = {Verification of Sequential Imperative Programs in + {I}sabelle/{\acs{hol}}}, + school = {Technische Universit\"at M\"unchen}, + year = 2006, + acknowledgement={brucker, 2007-04-23}, + abstract = {The purpose of this thesis is to create a verification + environment for sequential imperative programs. First a + general language model is proposed, which is independent of + a concrete programming language but expressive enough to + cover all common language features: mutually recursive + procedures, abrupt termination and exceptions, runtime + faults, local and global variables, pointers and heap, + expressions with side effects, pointers to procedures, + partial application and closures, dynamic method invocation + and also unbounded nondeterminism. + + For this language a Hoare logic for both partial and total + correctness is developed and on top of it a verification + condition generator is implemented. The Hoare logic is + designed to allow the integration of program analysis or + software model checking into the verification. + + To demonstrate the continuity to a real programming + language a subset of C is embedded into the verification + environment. + + The whole work is developed in the theorem prover Isabelle. + Therefore the correctness is machine-checked and in + addition the rich infrastructure of the general purpose + theorem prover Isabelle can be employed for the + verification of imperative programs. + + } +} + +@Article{ harel.ea:meaningful:2004, + author = {David Harel and Bernhard Rumpe}, + title = {Meaningful Modeling: What's the Semantics of + ``Semantics''?}, + journal = {\acs{ieee} Computer}, + year = 2004, + pages = {64--72}, + volume = 37, + issn = {0018-9162}, + number = 10, + month = oct, + publisher = pub-ieee, + address = pub-ieee:adr, + doi = {10.1109/MC.2004.172}, + acknowledgement={brucker, 2007-04-23} +} + +@Book{ hahnle:automated:1994, + author = {Reiner H{\"a}hnle}, + title = {Automated Deduction in Multiple-valued Logics}, + publisher = pub-oxford, + address = pub-oxford:adr, + disvolume = 10, + disseries = {International Series of Monographs on Computer Science}, + year = 1994, + acknowledgement={brucker, 2007-04-23}, + isbn = {0-19-853989-4}, + bibkey = {hahnle:automated:1994} +} + +@Book{ vigano:labelled:2000, + author = {Luca Vigan{\`o}}, + title = {Labelled Non-Classical Logics}, + year = 2000, + language = {USenglish}, + publisher = pub-kluwer, + address = pub-kluwer:adr, + isbn = {0-7923-7749-4}, + cover = {2000/lncl.png}, + abstract = {The subject of the book is the development and + investigation of a framework for the modular and uniform + presentation and implementation of non-classical logics, in + particular modal and relevance logics. Logics are presented + as labelled deduction systems, which are proved to be sound + and complete with respect to the corresponding Kripke-style + semantics. We investigate the proof theory of our systems, + and show them to possess structural properties such as + normalization and the subformula property, which we exploit + not only to establish advantages and limitations of our + approach with respect to related ones, but also to give, by + means of a substructural analysis, a new proof-theoretic + method for investigating decidability and complexity of + (some of) the logics we consider. All of our deduction + systems have been implemented in the generic theorem prover + Isabelle, thus providing a simple and natural environment + for interactive proof development.}, + acknowledgement={brucker, 2007-04-23}, + bibkey = {vigano:labelled:2000} +} + +@Book{ gabbay:labelled:1997, + author = {Dov M. Gabbay}, + title = {Labelled Deductive Systems}, + publisher = pub-oxford, + address = pub-oxford:adr, + series = {Oxford Logic Guides}, + year = 1997, + isbn = {978-0-198-53833-2}, + volume = 1, + acknowledgement={brucker, 2007-04-23}, + bibkey = {gabbay:labelled:1997} +} + +@Book{ warmer.ea:ocl2:2003, + bibkey = {warmer.ea:ocl2:2003}, + author = {Jos Warmer and Anneke Kleppe}, + abstract = {This book covers \acs{ocl} 2.0}, + keywords = {OCL, \acs{uml}}, + isbn = {0-321-17936-6}, + edition = {2nd}, + language = {USenglish}, + public = {yes}, + topic = {formalism}, + publisher = pub-awl, + address = pub-awl:adr, + title = {The Object Constraint Language: Getting Your Models Ready + for \acs{mda}}, + year = 2003, + month = aug, + num_pages = 240, + price = {39.99}, + cover = {2003/ocl2.png}, + currency = {USD}, + acknowledgement={brucker, 2007-04-23} +} + +@Book{ smith:object:2000, + author = {Graeme Smith}, + title = {The Object {Z} Specification Language}, + publisher = pub-kluwer, + address = pub-kluwer:adr, + year = 2000, + isbn = {0-7923-8684-1}, + pages = 160, + abstract = {bject-Z is an object-oriented extension of the formal + specification language Z. It adds, to Z, notions of classes + and objects, and inheritance and polymorphism. By extending + Z's semantic basis, it enables the specification of systems + as collections of independent objects in which self and + mutual referencing are possible. + + The Object-Z Specification Language presents a + comprehensive description of Object-Z including discussions + of semantic issues, definitions of all language constructs, + type rules and other rules of usage, specification + guidelines, and a full concrete syntax. It will enable you + to confidently construct Object-Z specifications and is + intended as a reference manual to keep by your side as you + use and learn to use Object-Z. + + The Object-Z Specification Language is suitable as a + textbook or as a secondary text for a graduate level + course, and as a reference for researchers and + practitioners in industry.}, + series = {Advances in Formal Methods Series}, + acknowledgement={brucker, 2007-04-23} +} + +@Book{ meyer:object-oriented:1988, + author = {Bertrand Meyer}, + title = {Object-Oriented Software Construction}, + acknowledgement={brucker, 2007-04-23}, + publisher = pub-prentice, + address = pub-prentice:adr, + year = 1988, + isbn = {0-13-629031-0}, + descriptor = {Eiffel, Objekt-orientiert, Software} +} + +@Article{ meyer.ea:interactive:1999, + title = {Interactive Verification Environments for Object-Oriented + Languages }, + author = {J{\"o}rg Meyer and Arnd Poetzsch-Heffter }, + journal = j-ucs, + volume = 5, + url = {http://www.jucs.org/jucs_5_3/interactive_verification_environments_for} + , + number = 3, + pages = {208--225 }, + year = 1999, + doi = {10.1007/3-540-46419-0_6}, + acknowledgement={brucker, 2007-04-23}, + abstract = {Formal specification and verification techniques can + improve the quality of object-oriented software by enabling + semantic checks and certification of properties. To be + applicable to object-oriented programs, they have to cope + with subtyping, aliasing via object references, as well as + abstract and recursive methods. For mastering the resulting + complexity, mechanical aid is needed. The article outlines + the specific technical requirements for the specification + and verification of object-oriented programs. Based on + these requirements, it argues that verification of + OO-programs should be done interactively and develops an + modular architecture for this task. In particular, it shows + how to integrate interactive program verification with + existing universal interactive theorem provers, and + explains the new developed parts of the architecture. To + underline the general approach, we describe interesting + features of our prototype implementation.} +} + +@Article{ church:types:1940, + author = {Church, Alonzo}, + title = {A formulation of the simple theory of types}, + journal = j-sl, + year = 1940, + volume = 5, + number = 2, + month = jun, + pages = {56--68}, + acknowledgement={brucker, 2007-04-23}, + bibkey = {church:types:1940} +} + +@Article{ muller.ea:holcf:1999, + author = {Olaf M\"uller and Tobias Nipkow and David von Oheimb and + Oskar Slotosch}, + title = {\acs{holcf} = \acs{hol} + \acs{lcf}}, + journal = j-fp, + number = 2, + doi = {10.1017/S095679689900341X}, + volume = 9, + pages = {191--223}, + year = 1999, + abstract = {HOLCF is the definitional extension of Church's + Higher-Order Logic with Scott's Logic for Computable + Functions that has been implemented in the theorem prover + Isabelle. This results in a flexible setup for reasoning + about functional programs. HOLCF supports standard domain + theory (in particular fixedpoint reasoning and recursive + domain equations) but also coinductive arguments about lazy + datatypes. This paper describes in detail how domain theory + is embedded in HOL and presents applications from + functional programming, concurrency and denotational + semantics. }, + acknowledgement={brucker, 2007-04-23}, + bibkey = {muller.ea:holcf:1999} +} + +@Article{ huet:programtransformations:1978, + author = {G{\'e}rard Huet and Bernard Lang}, + title = {Proving and Applying Program Transformations Expressed + with Second Order Patterns}, + journal = {Acta Informatica}, + volume = 11, + year = 1978, + pages = {31--55}, + number = 1, + doi = {10.1007/BF00264598}, + acknowledgement={brucker, 2007-04-23}, + bibkey = {huet:programtransformations:1978} +} + +@PhDThesis{ kyas:verifying:2006, + author = {Marcel Kyas}, + title = {Verifying {\acs{ocl}} Specifications of {\acs{uml}} + Models: Tool Support and Compositionality}, + school = {University of Leiden}, + year = 2006, + acknowledgement={brucker, 2007-04-23}, + abstract = {The Unified Modelling Language (\acs{uml}) and the Object + Constraint Language (OCL) serve as specification languages + for embedded and real-time systems used in a + safety-critical environment. In this dissertation class + diagrams, object diagrams, and OCL constraints are + formalised. The formalisation serves as foundation for a + translation of class diagrams, state machines, and + constraints into the theorem prover PVS. This enables the + formal verification of models defined in a subset of + \acs{uml} using the interactive theorem prover. The type + system of OCL makes writing specifications difficult while + the model is still under development. To overcome this + difficulty a new type system is proposed, based on + intersection types, union types, and bounded operator + abstraction. To reduce the complexity of the model and to + increase the structure of the specification, compositional + reasoning is used. The introduction of history variables + allows compositional specifications. Proof rules support + compositional reasoning. The feasibility of the presented + approach is demonstrated by two case-studies. The first one + is the "Sieve of Eratosthenes" and the second one is a part + of the medium altitude reconnaissance system (MARS) + deployed in F-16 fighters of the Royal Dutch Air Force.}, + isbn = {3-86541-142-8}, + publisher = {Lehmanns Media}, + address = {Berlin}, + file = {papers/2006/kyas-verifying-2006.tgz} +} + +@Article{ bertino.ea:object-oriented:1992, + author = {Elisa Bertino and Mauro Negri and Giuseppe Pelagatti and + Licia Sbattella}, + title = {Object-Oriented Query Languages: The Notion and the + Issues}, + journal = j-tkde, + volume = 4, + number = 3, + doi = {10.1109/69.142014}, + library = {DINF-K}, + year = 1992, + pages = {223--237}, + abstract = {The authors describe how the characteristics of an + object-oriented data model, such as object identity, + complex object structure, methods, and class hierarchies, + have an impact on the design of a query language. They also + point out major differences with respect to relational + query languages. The discussion is supported through the + definition of OOPC, a formal object-oriented query language + based on predicate calculus, which incorporates in a + consistent formal notation most features of existing + object-oriented query languages.}, + acknowledgement={brucker, 2007-04-23}, + bibkey = {bertino.ea:object-oriented:1992} +} + +@Article{ beckert.ea:leant-ap:1995, + author = {Bernhard Beckert and Joachim Posegga}, + title = {{\leanTAP}: Lean Tableau-based Deduction}, + journal = j-ar, + volume = 15, + number = 3, + pages = {339--358}, + year = 1995, + publisher = pub-springer, + address = pub-springer:adr, + doi = {10.1007/BF00881804}, + acknowledgement={brucker, 2007-04-23}, + bibkey = {beckert.ea:leant-ap:1995} +} + +@Article{ jackson:alloy:2002, + author = {Daniel Jackson}, + title = {{Alloy}: a lightweight object modelling notation}, + journal = j-tosem, + volume = 11, + number = 2, + year = 2002, + issn = {1049-331X}, + pages = {256--290}, + doi = {10.1145/505145.505149}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={brucker, 2007-04-23}, + abstract = {Alloy is a little language for describing structural + properties. It offers a declaration syntax compatible with + graphical object models, and a set-based formula syntax + powerful enough to express complex constraints and yet + amenable to a fully automatic semantic analysis. Its + meaning is given by translation to an even smaller + (formally defined) kernel. This paper presents the language + in its entirety, and explains its motivation, contributions + and deficiencies. } +} + +@Article{ kobryn:uml:1999, + author = {Cris Kobryn}, + title = {{\acs{uml}} 2001: a standardization odyssey}, + journal = j-cacm, + volume = 42, + number = 10, + year = 1999, + issn = {0001-0782}, + pages = {29--37}, + doi = {10.1145/317665.317673}, + publisher = pub-acm, + address = pub-acm:adr, + language = {USEnglish}, + acknowledgement={brucker, 2007-04-23} +} + +@Article{ basin.ea:natural:1998, + author = {David A. Basin and Se{\'a}n Matthews and Luca Vigan{\`o}}, + title = {Natural Deduction for Non-Classical Logics}, + journal = {Studia Logica}, + doi = {10.1023/A:1005003904639}, + year = 1998, + volume = 60, + number = 1, + publisher = pub-springer, + address = pub-acm:adr, + issn = {0039-3215}, + acknowledgement={brucker, 2007-04-23}, + pages = {119--160}, + language = {USenglish}, + note = {Special issue on \emph{Natural Deduction} edited by Frank + Pfenning and Wilfried Sieg}, + abstract = {We present a framework for machine implementation of + families of non-classical logics with Kripke-style + semantics. We decompose a logic into two interacting parts, + each a natural deduction system: a base logic of labelled + formulae, and a theory of labels characterizing the + properties of the Kripke models. By appropriate + combinations we capture both partial and complete fragments + of large families of non-classical logics such as modal, + relevance, and intuitionistic logics. Our approach is + modular and supports uniform proofs of correctness and + proof normalization. We have implemented our work in the + Isabelle Logical Framework. } +} + +@InProceedings{ nipkow.ea:java-light:1998, + author = {Tobias Nipkow and David von Oheimb}, + title = {{Java$_{{\ell}ight}$} is Type-Safe---Definitely}, + booktitle = {\acs{acm} Symp.\ Principles of Programming Languages + (\acs{popl})}, + publisher = pub-acm, + isbn = {0-89791-979-3}, + address = pub-acm:adr, + year = 1998, + pages = {161--170}, + doi = {10.1145/268946.268960}, + location = {San Diego, California, United States}, + acknowledgement={brucker, 2007-04-23}, + bibkey = {nipkow.ea:java-light:1998} +} + +@Article{ bierman.ea:mj:2003, + author = {Gavin M. Bierman and Matthew J. Parkinson}, + title = {Effects and effect inference for a core {Java} calculus}, + journal = j-entcs, + volume = 82, + number = 7, + year = 2003, + doi = {10.1016/S1571-0661(04)80803-X}, + pages = {1--26}, + booktitle = {WOOD2003, Workshop on Object Oriented Developments + (Satellite Event of ETAPS 2003)}, + acknowledgement={brucker, 2007-04-23}, + bibkey = {bierman.ea:mj:2003}, + publisher = pub-elsevier, + address = pub-elsevier:adr +} + +@Article{ chiorean.ea:ensuring:2004, + author = {Dan Chiorean and Mihai Pasca and Adrian C{\^a}rcu and + Cristian Botiza and Sorin Moldovan}, + title = {Ensuring \acs{uml} Models Consistency Using the \acs{ocl} + Environment.}, + journal = j-entcs, + volume = 102, + booktitle = PROC # { the Workshop, \acs{ocl} 2.0 -- Industry + Standard or Scientific Playground?}, + year = 2004, + acknowledgement={brucker, 2007-04-24}, + pages = {99--110}, + doi = {10.1016/j.entcs.2003.09.005}, + publisher = pub-elsevier, + address = pub-elsevier:adr +} + +@Article{ rauch.ea:formalizing:2003, + abstract = {We present a formal model of the Java two's-complement + integral arithmetics. The model directly formalizes the + arithmetic operations as given in the Java Language + Specification (JLS). The algebraic properties of these + definitions are derived. Underspecifications and + ambiguities in the JLS are pointed out and clarified. The + theory is formally analyzed in Isabelle/HOL, that is, + machine-checked proofs for the ring properties and + divisor/remainder theorems etc. are provided. This work is + suited to build the framework for machine-supported + reasoning over arithmetic formulae in the context of Java + source-code verification.}, + author = {Nicole Rauch and Burkhart Wolff}, + journal = j-entcs, + doi = {10.1016/S1571-0661(04)80808-9}, + acknowledgement={brucker, 2007-04-24}, + publisher = pub-elsevier, + address = pub-elsevier:adr, + title = {Formalizing {Java}'s Two's-Com\-ple\-ment Integral Type in + {Isabelle}/\acs{hol}}, + volume = 80, + year = 2003, + pages = {1--18}, + booktitle = {International Workshop on Formal Methods for Industrial + Critical Systems (\ac{fmics})} +} + +@Article{ kyas.ea:formalizing:2004, + journal = j-entcs, + author = {Kyas, Marcel and Fecher, Harald and de Boer, Frank S. and + van der Zwaag, Mark and Hooman, Jozef and Arons, Tamarah + and Kugler, Hillel}, + title = {Formalizing {\acs{uml}} Models and {\acs{ocl}} Constraints + in {\acs{pvs}}}, + booktitle = {Workshop on Semantic Foundations of Engineering Design + Languages (\acs{sfedl})}, + year = 2004, + doi = {10.1016/j.entcs.2004.09.027 }, + pages = {39--47}, + acknowledgement={brucker, 2007-04-24}, + publisher = pub-elsevier, + address = pub-elsevier:adr +} + +@Proceedings{ grabowski.ea:formal:2005, + editor = {Jens Grabowski and Brian Nielsen}, + title = {Formal Approaches to Software Testing (\textsc{fates})}, + booktitle = {Formal Approaches to Software Testing (\textsc{fates})}, + series = s-lncs, + volume = 3395, + year = 2005, + isbn = {3-540-25109-X}, + doi = {10.1007/b106767}, + acknowledgement={brucker, 2007-04-24}, + publisher = pub-springer, + address = pub-springer:adr +} + +@InCollection{ kerber.ea:tableau:1996, + author = {Manfred Kerber and Michael Kohlhase}, + title = {A Tableau Calculus for Partial Functions}, + pages = {21--49}, + abstract = {Even though it is not very often admitted, partial + functions do play a significant role in many practical + applications of deduction systems. Kleene has already given + a semantic account of partial functions using a + three-valued logic decades ago, but there has not been a + satisfactory mechanization. Recent years have seen a + thorough investigation of the framework of many-valued + truth-functional logics. However, strong Kleene logic, + where quantification is restricted and therefore not + truth-functional, does not fit the framework directly. We + solve this problem by applying recent methods from sorted + logics. This paper presents a tableau calculus that + combines the proper treatment of partial functions with the + efficiency of sorted calculi.}, + acknowledgement={brucker, 2007-04-24}, + bibkey = {kerber.ea:tableau:1996}, + booktitle = {Collegium Logicum---Annals of the Kurt-G{\"o}del-Society}, + volume = 2, + publisher = pub-springer-ny, + address = pub-springer-ny:adr, + isbn = {3-211-82796-X}, + year = 1996 +} + +@InCollection{ paulson:generic:1996, + author = {Lawrence C. Paulson}, + title = {Generic automatic proof tools}, + pages = {23--47}, + abstract = {This paper explores a synthesis between two distinct + traditions in automated reasoning: resolution and + interaction. In particular it discusses Isabelle, an + interactive theorem prover based upon a form of resolution. + It aims to demonstrate the value of proof tools that, + compared with traditional resolution systems, seem absurdly + limited. Isabelle's classical reasoner searches for proofs + using a tableau approach. The reasoner is generic: it + accepts rules proved in applied theories, involving defined + connectives. New constants are not reduced to first-order + logic; the reasoner}, + acknowledgement={brucker, 2007-04-24}, + bibkey = {paulson:generic:1996}, + booktitle = {Automated reasoning and its applications: essays in honor + of {Larry Wos}}, + year = 1997, + editor = {Robert Veroff}, + publisher = pub-mit, + address = pub-mit:adr, + isbn = {978-0-262-22055-2} +} + +@InCollection{ nipkow:order-sorted:1993, + author = {Tobias Nipkow}, + title = {Order-Sorted Polymorphism in {Isabelle}}, + booktitle = {Workshop on Logical Environments}, + editor = {G\'erard Huet and Gordon Plotkin}, + publisher = pub-cup, + address = pub-cup:adr, + year = 1993, + location = {Edinburgh, Scotland}, + pages = {164--188}, + acknowledgement={brucker, 2007-04-24}, + isbn = {0-521-43312-6}, + bibkey = {nipkow:order-sorted:1993} +} + +@InCollection{ hahnle:tableaux:1999, + author = {Reiner H\"ahnle}, + booktitle = {Handbook of Tableau Methods}, + editor = {Marcello D'Agostino and Dov Gabbay and Reiner H\"ahnle and + Joachim Posegga}, + isbn = {978-0-792-35627-1}, + publisher = pub-kluwer, + address = pub-kluwer:adr, + title = {Tableaux for Many-Valued Logics}, + pages = {529--580}, + year = 1999, + public = {no}, + mynote = {Postscript-File nicht weitergeben!}, + abstract = {This article reports on research done in the intersection + between many-valued logics and logical calculi related to + tableaux. A lot of important issues in many-valued logic, + such as algebras arising from many-valued logic, + many-valued function minimization, philosophical topics, or + applications are not discussed here; for these, we refer + the reader to general monographs and overviews such as + [Rosser and Turquette, 1952; Rescher, 1969; Urquhart, 1986; + Bolc and Borowik, 1992; Malinowski, 1993; Hahnle, 1994; + Panti, to appear] . More questionable, perhaps, than the + omissions is the need for a handbook chapter on tableaux + for many-valued logics in the first place.}, + bibkey = {hahnle:tableaux:1999}, + acknowledgement={brucker, 2007-04-24} +} + +@InCollection{ leavens.ea:jml:1999, + author = {Gary T. Leavens and Albert L. Baker and Clyde Ruby}, + title = {{JML}: A Notation for Detailed Design}, + booktitle = {Behavioral Specifications of Businesses and Systems}, + publisher = pub-kluwer, + address = pub-kluwer:adr, + editor = {Haim Kilov and Bernhard Rumpe and Ian Simmonds}, + pages = {175--188}, + year = 1999, + isbn = {978-0-7923-8629-2}, + acknowledgement={brucker, 2007-04-24} +} + +@InCollection{ paulson:fixedpoint:2000, + author = {Lawrence C. Paulson}, + pages = {187--211}, + title = {A fixedpoint approach to (co)inductive and (co)datatype + definitions}, + acknowledgement={brucker, 2007-04-24}, + abstract = {This paper presents a fixedpoint approach to inductive + definitions. In- stead of using a syntactic test such as + ``strictly positive,'' the approach lets definitions + involve any operators that have been proved monotone. It is + conceptually simple, which has allowed the easy + implementation of mutual recursion and iterated + definitions. It also handles coinductive definitions: + simply replace the least fixedpoint by a greatest + fixedpoint. The method has been implemented in two of + Isabelle's logics, zf set theory and higher-order logic. It + should be applicable to any logic in which the + Knaster-Tarski theorem can be proved. Examples include + lists of n elements, the accessible part of a relation and + the set of primitive recursive functions. One example of a + coinductive definition is bisimulations for lazy lists. + Recursive datatypes are examined in detail, as well as one + example of a codatatype: lazy lists. The Isabelle package + has been applied in several large case studies, including + two proofs of the Church-Rosser theorem and a coinductive + proof of semantic consistency. The package can be trusted + because it proves theorems from definitions, instead of + asserting desired properties as axioms. }, + bibkey = {paulson:fixedpoint:2000}, + crossref = {plotkin.ea:proof:2000} +} + +@InCollection{ gordon:from:2000, + author = {Mike Gordon}, + title = {From \acs{lcf} to {\acs{hol}}: a short history}, + pages = {169--185}, + acknowledgement={brucker, 2007-04-24}, + abstract = {The original LCF system was a proof-checking program + developed at Stanford University by Robin Milner in 1972. + Descendents of LCF now form a thriving paradigm in computer + assisted reasoning. Many of the developments of the last 25 + years have been due to Robin Milner, whose influence on the + field of automated reasoning has been diverse and profound. + One of the descendents of LCF is HOL, a proof assistant for + higher order logic originally developed for reasoning about + hardware. The multi-faceted contribution of Robin Milner to + the development of HOL is remarkable. Not only did he + invent the LCF-approach to theorem proving, but he also + designed the ML programming language underlying it and the + innovative polymorphic type system used both by ML and the + LCF and HOL logics. Code Milner wrote is still in use + today, and the design of the hardware verification system + LCF_LSM (a now obsolete stepping stone from LCF to HOL) was + inspired by Milner's Calculus of Communicating Systems + (CCS). }, + crossref = {plotkin.ea:proof:2000} +} + +@Book{ plotkin.ea:proof:2000, + booktitle = {Proof, Language, and Interaction: Essays in Honour of + {Robin Milner}}, + title = {Proof, Language, and Interaction: Essays in Honour of + {Robin Milner}}, + publisher = pub-mit, + address = pub-mit:adr, + year = 2000, + isbn = {978-0-262-16188-6}, + acknowledgement={brucker, 2007-04-24}, + editor = {Gordon Plotkin and Colin Stirling and Mads Tofte} +} + +@InCollection{ pfenning:logical:2001, + author = {Frank Pfenning}, + title = {Logical Frameworks}, + booktitle = {Handbook of Automated Reasoning}, + editor = {Alan Robinson and Andrei Voronkov}, + chapter = 17, + volume = 2, + isbn = {0-444-50812-0}, + pages = {1063--1147}, + publisher = pub-elsevier, + address = pub-elsevier:adr, + year = 2001, + acknowledgement={brucker, 2007-04-24}, + abstract = {The specification of a deductive system usually proceeds + in two stages: first we define the syntax of an object + language and then the axioms and rules of inference. In + order to concentrate on the meanings of expressions we + ignore issues of concrete syntax and parsing and + concentrate on specifying abstract syntax. Di#erent + framework implementations provide di#erent means for + customizing the parser in order to embed the desired + object-language syntax. As an example throughout this + chapter we consider formulations of intuitionistic and + classical first-order logic. In order to keep this chapter + to a manageable length, we restrict ourselves to the + fragment containing implication, negation, and universal + quantification. The reader is invited to test his or her + understanding by extending the development to include a + more complete set of connectives and quantifiers. } +} + +@InProceedings{ pfenning:hoas:1988, + author = {Frank Pfenning and Conal Elliot}, + title = {Higher-Order Abstract Syntax}, + year = 1988, + isbn = {0-89791-269-1}, + location = {Atlanta, Georgia, United States}, + doi = {10.1145/53990.54010}, + pages = {199--208}, + acknowledgement={brucker, 2007-04-24}, + publisher = pub-acm, + address = pub-acm:adr, + booktitle = {Conference on Programming Language Design and + Implementation (\acs{pldi})}, + bibkey = {pfenning:hoas:1988} +} + +@InProceedings{ boulton.ea:experience:1993, + crossref = {stavridou.ea:international:1993}, + author = {Richard Boulton and Andrew Gordon and Michael J. C. Gordon + and John Harrison and John Herbert and John Van Tassel}, + title = {Experience with embedding hardware description languages + in {\acs{hol}}}, + bibkey = {boulton.ea:experience:1993}, + abstract = {The semantics of hardware description languages can be + represented in higher order logic. This provides a formal + definition that is suitable for machine processing. + Experiments are in progress at Cambridge to see whether + this method can be the basis of practical tools based on + the HOL theorem-proving assistant. Three languages are + being investigated: ELLA, Silage and VHDL. The approaches + taken for these languages are compared and current progress + on building semantically-based theorem-proving tools is + discussed.}, + acknowledgement={brucker, 2007-04-24}, + pages = {129--156} +} + +@Proceedings{ stavridou.ea:international:1993, + editor = {Victoria Stavridou and Thomas F. Melham and Raymond T. + Boute}, + booktitle = PROC # { the International Conference on Theorem + Provers in Circuit Design: Theory, Practice and + Experience}, + title = {International Conference on Theorem Provers in Circuit + Design: Theory, Practice and Experience (\acs{tpcd})}, + series = {\acs{ifip} Transactions}, + volume = {A-10}, + isbn = {0-444-89686-4}, + publisher = pub-north, + address = pub-north:adr, + date = {22--24 June 1993}, + acknowledgement={brucker, 2007-04-24}, + year = 1993 +} + +@InProceedings{ khoshafian.ea:object:1986, + author = {Setrag N. Khoshafian and George P. Copeland}, + title = {Object identity}, + booktitle = {Object-oriented programming systems, languages and + applications (\acs{oopsla})}, + year = 1986, + isbn = {0-89791-204-7}, + pages = {406--416}, + location = {Portland, Oregon, United States}, + doi = {10.1145/28697.28739}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={brucker, 2007-04-24} +} + +@InProceedings{ melham:package:1992, + author = {Thomas F. Melham}, + title = {A Package for Inductive Relation Definitions in {\HOL}}, + pages = {350--357}, + isbn = {0-8186-2460-4}, + editor = {Myla Archer and Jennifer J. Joyce and Karl N. Levitt and + Phillip J. Windley}, + booktitle = {International Workshop on the {\HOL} Theorem Proving + System and its Applications (\acs{tphols})}, + month = aug, + publisher = pub-ieee, + address = pub-ieee:adr, + year = 1992, + acknowledgement={brucker, 2007-04-24}, + location = {Davis, California, \acs{usa}}, + bibkey = {melham:package:1992} +} + +@InProceedings{ hahnle:efficient:1994, + doi = {10.1109/ismvl.1994.302195}, + author = {Reiner H\"{a}hnle}, + booktitle = {International Symposium on Multiple-Valued Logics + (\acs{ismvl})}, + location = {Boston/MA, \acs{usa}}, + pages = {240--249}, + isbn = {0-8186-5650-6}, + publisher = pub-ieee, + address = pub-ieee:adr, + title = {Efficient Deduction in Many-Valued Logics}, + year = 1994, + abstract = {This paper tries to identify the basic problems + encountered in automated theorem proving in many-valued + logics and demonstrates to which extent they can be + currently solved. To this end a number of recently + developed techniques are reviewed. We list the avenues of + research in many-valued theorem proving that are in our + eyes the most promising.}, + acknowledgement={brucker, 2007-04-24}, + bibkey = {hahnle:efficient:1994} +} + +@InProceedings{ nipkow.ea:java:2000, + author = {Tobias Nipkow and David von Oheimb and Cornelia Pusch}, + title = {{$\mu$Java}: Embedding a Programming Language in a Theorem + Prover}, + booktitle = {Foundations of Secure Computation}, + series = {\acs{nato} Science Series F: Computer and Systems + Sciences}, + volume = 175, + year = 2000, + publisher = pub-ios, + address = pub-ios:adr, + isbn = {978-1-58603-015-5}, + editor = {Friedrich L. Bauer and Ralf Steinbr{\"u}ggen}, + abstract = {This paper introduces the subset $micro$Java of Java, + essentially by omitting everything but classes. The type + system and semantics of this language (and a corresponding + abstract Machine $micro$JVM) are formalized in the theorem + prover Isabelle/\acs{hol}. Type safety both of $micro$Java + and the $micro$JVM are mechanically verified. + + To make the paper self-contained, it starts with + introductions to Isabelle/\acs{hol} and the art of + embedding languages in theorem provers.}, + pages = {117--144}, + acknowledgement={brucker, 2007-04-24}, + bibkey = {nipkow.ea:java:2000} +} + +@InProceedings{ beckert.ea:translating:2002, + editor = {Serge Autexier and Heiko Mantel}, + pages = {113--123}, + booktitle = {Verification Workshop (\acs{verify})}, + location = {Copenhagen, Denmark}, + author = {Bernhard Beckert and Uwe Keller and Peter H. Schmitt}, + title = {Translating the {O}bject {C}onstraint {L}anguage into + First-order Predicate Logic}, + year = 2002, + abstract = {In this paper, we define a translation of \acs{uml} class + diagrams with OCL constraints into first-order predicate + logic. The goal is logical reasoning about \acs{uml} models. + + We put an emphasis on usability of the formulas resulting + from the translation, and we have developed optimisations + and heuristics to enhance the efficiency of the theorem + proving process. + + The translation has been implemented as part of the KeY + system, but our implementation can also be used + stand-alone. }, + acknowledgement={brucker, 2007-04-24}, + bibkey = {beckert.ea:translating:2002} +} + +@InProceedings{ demuth.ea:generation:2005, + author = {Birgit Demuth and Heinrich Hussmann and Ansgar Konermann}, + title = {Generation of an {\acs{ocl}} 2.0 Parser}, + booktitle = {Workshop on Tool Support for \acs{ocl} and Related + Formalisms---Needs and Trends}, + location = {Montego Bay, Jamaica, October 4, 2005}, + pages = {38--52}, + publisher = {\acs{epfl}}, + year = 2005, + editor = {Thomas Baar}, + series = {Technical Report LGL-REPORT-2005-001}, + acknowledgement={brucker, 2007-04-24} +} + +@InProceedings{ aredo:formalizing:1999, + author = {Demissie B. Aredo}, + booktitle = {\acs{oopsla}'99 Workshop on Rigorous Modeling and Analysis + with the \acs{uml}: Challenges and Limitations, Denver, + Colorado}, + title = {Formalizing {\acs{uml}} Class Diagrams in {\acs{pvs}}}, + year = 1999, + month = nov, + address = {Denver, Colorado, \acs{usa}}, + acknowledgement={brucker, 2007-04-24} +} + +@InProceedings{ jackson.ea:alcoa:2000, + abstract = {Alcoa is a tool for analyzing object models. It has a + range of uses. At one end, it can act as a support tool for + object model diagrams, checking for consistency of + multiplicities and generating sample snapshots. At the + other end, it embodies a lightweight formal method in which + subtle properties of behaviour can be investigated. Alcoa's + input language, Alloy, is a new notation based on Z. Its + development was motivated by the need for a notation that + is more closely tailored to object models (in the style of + \acs{uml}), and more amenable to automatic analysis. Like + Z, Alloy supports the description of systems whose state + involves complex relational structure. State and + behavioural properties are described declaratively, by + conjoining constraints. This makes it possible to develop + and analyze a model incrementally, with Alcoa investigating + the consequences of whatever constraints are given. Alcoa + works by translating constraints to boolean formulas, and + then applying state-of-the-art SAT solvers. It can analyze + billions of states in seconds. }, + author = {Daniel Jackson and Ian Schechter and Ilya Shlyakhter }, + booktitle = {International Conference on Software Engineering + (\acs{icse})}, + language = {USenglish}, + month = jun, + public = {yes}, + doi = {10.1109/ICSE.2000.870482}, + location = {Limerick, Ireland}, + isbn = {1-58113-206-9}, + pages = {730--733}, + topic = {formalism}, + keywords = {Aloca, Alloy}, + title = {{A}lcoa: the {A}lloy Constraint Analyzer }, + year = 2000, + timestamp = 962701274, + acknowledgement={none}, + publisher = pub-acm, + address = pub-acm:adr +} + +@Article{ hahnle:many-valued:2005, + author = {Reiner H\"{a}hnle}, + title = {Many-Valued Logic, Partiality, and Abstraction in Formal + Specification Languages}, + journal = {Logic Journal of the \textsc{igpl}}, + year = 2005, + volume = 13, + pages = {415--433}, + month = jul, + doi = {10.1093/jigpal/jzi032}, + number = 4, + acknowledgement={brucker, 2007-05-04} +} + +@Booklet{ levens.ea:jml:2007, + bibkey = {levens.ea:jml:2007}, + author = {Gary T. Leavens and Erik Poll and Curtis Clifton and + Yoonsik Cheon and Clyde Ruby and David R. Cok and Peter + M\"{u}ller and Joseph Kiniry and Patrice Chalin}, + title = {{\acs{jml}} Reference Manual (Revision 1.2)}, + month = feb, + year = 2007, + organization = {Department of Computer Science, Iowa State University.}, + note = {Available from \url{http://www.jmlspecs.org}}, + acknowledgement={brucker, 2007-04-23} +} + +@InProceedings{ broy.ea:uml2:2006, + bibkey = {broy.ea:uml2:2006}, + author = {Manfred Broy and Michelle L. Crane and J{\"u}rgen Dingel + and Alan Hartman and Bernhard Rumpe and Bran Selic}, + title = {2nd \acs{uml} 2 Semantics Symposium: Formal Semantics for + {\acs{uml}}}, + doi = {10.1007/978-3-540-69489-2_39}, + pages = {318--323}, + abstract = {The purpose of this symposium, held in conjunction with + \acs{models} 2006, was to present the current state of + research of the UML 2 Semantics Project. Equally important + to receiving feedback from an audience of experts was the + opportunity to invite researchers in the field to discuss + their own work related to a formal semantics for the + Unified Modeling Language. This symposium is a follow-on to + our first workshop, held in conjunction with ECMDA 2005.}, + acknowledgement={brucker, 2007-04-23}, + crossref = {kuhne:models:2006} +} + +@InProceedings{ hafner.ea:towards:2006, + author = {Michael Hafner and Muhammad Alam and Ruth Breu}, + title = {Towards a {MOF/QVT}-Based Domain Architecture for Model + Driven Security}, + booktitle = {MoDELS}, + year = 2006, + pages = {275--290}, + ee = {10.1007/11880240_20}, + crossref = {nierstrasz.ea:model:2006} +} + +@Proceedings{ kuhne:models:2006, + doi = {10.1007/978-3-540-69489-2}, + booktitle = {Models in Software Engineering---Workshops and Symposia at + \acs{models} 2006}, + title = {Models in Software Engineering---Workshops and Symposia at + \acs{models} 2006}, + isbn = {978-3-540-69488-5}, + publisher = pub-springer, + paddress = pub-springer:adr, + address = {Genua, Italy}, + volume = 4364, + series = s-lncs, + year = 2006, + acknowledgement={brucker, 2007-04-23}, + editor = {Thomas K{\"u}hne} +} + +@Book{ russell:introduction:1919, + author = {Bertrand Russell}, + title = {Introduction to Mathematical Philosophy}, + publisher = {George Allen \& Unwin}, + year = 1919, + acknowledgement={brucker, 2007-04-23}, + address = {London} +} + +@Article{ bertino.ea:trbac:2001, + author = {Elisa Bertino and Piero Andrea Bonatti and Elena Ferrari}, + title = {TRBAC: A temporal role-based access control model}, + journal = {ACM Trans. Inf. Syst. Secur.}, + volume = 4, + number = 3, + year = 2001, + issn = {1094-9224}, + pages = {191--233}, + doi = {10.1145/501978.501979}, + publisher = pub-acm, + address = pub-acm:adr, + tags = {noTAG}, + clearance = {unclassified}, + timestap = {2008-05-29} +} + +@Article{ moyer.ea:generalized:2001, + title = {Generalized role-based access control}, + author = {Moyer, M.J. and Abamad, M.}, + journal = {Distributed Computing Systems, 2001. 21st International + Conference on.}, + year = 2001, + month = {Apr}, + pages = {391--398}, + keywords = {authorisation, distributed processing, transaction + processingGRBAC, JPEG, RBAC, access control, access control + decisions, access control models, environment roles, + environmental information, expressive power, generalized + role based access control, object roles, object type, rich + access control policies, security policy, security-relevant + characteristics, sensitivity level, subject roles}, + doi = {10.1109/ICDSC.2001.918969}, + abstract = {Generalized Role-Based Access Control (GRBAC) is a new + paradigm for creating and maintaining rich access control + policies. GRBAC leverages and extends the power of + traditional role based access control (RBAC) by + incorporating subject roles, object roles and environment + roles into access control decisions. Subject roles are like + traditional RBAC roles: they abstract the security-relevant + characteristics of subjects into categories that can be + used in defining a security policy. Similarly, object roles + abstract the various properties of objects, such as object + type (e.g., text, JPEG, executable) or sensitivity level + (e.g., classified, top secret) into categories. Environment + roles capture environmental information, such as time of + day or system load so it can be used to mediate access + control. Together, these three types of roles offer + flexibility and expressive power, as well as a degree of + usability not found in current access control models}, + tags = {noTAG}, + clearance = {unclassified}, + timestap = {2008-05-29} +} + +@InProceedings{ zhang.ea:role-based:2002, + author = {Longhua Zhang and Gail-Joon Ahn and Bei-Tseng Chu}, + title = {A role-based delegation framework for healthcare + information systems}, + booktitle = PROC # { the seventh \acs{acm} symposium on Access + control models and technologies (\acs{sacmat})}, + year = 2002, + isbn = {1-58113-496-7}, + pages = {125--134}, + location = {Monterey, California, \acs{usa}}, + doi = {10.1145/507711.507731}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {As organizations implement information strategies that + call for sharing access to resources in the networked + environment, mechanisms must be provided to protect the + resources from adversaries. The proposed delegation + framework addresses the issue of how to advocate selective + information sharing in role-based systems while minimizing + the risks of unauthorized access. We introduce a systematic + approach to specify delegation and revocation policies + using a set of rules. We demonstrate the feasibility of our + framework through policy specification, enforcement, and a + proof-of-concept implementation on specific domains, e.g. + the healthcare environment. We believe that our work can be + applied to organizations that rely heavily on collaborative + tasks.}, + tags = {noTAG}, + clearance = {unclassified}, + timestap = {2008-05-29} +} + +@InProceedings{ wilikens.ea:context-related:2002, + author = {Marc Wilikens and Simone Feriti and Alberto Sanna and + Marcelo Masera}, + title = {A context-related authorization and access control method + based on \acs{rbac}: A case study from the health care + domain}, + booktitle = PROC # { the seventh \acs{acm} symposium on Access + control models and technologies (\acs{sacmat})}, + year = 2002, + isbn = {1-58113-496-7}, + pages = {117--124}, + location = {Monterey, California, USA}, + doi = {10.1145/507711.507730}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {This paper describes an application of authorization and + access control based on the Role Based Access Control + (RBAC) method and integrated in a comprehensive trust + infrastructure of a health care application. The method is + applied to a health care business process that involves + multiple actors accessing data and resources needed for + performing clinical and logistics tasks in the application. + The notion of trust constituency is introduced as a concept + for describing the context of authorisation. In addition, + the applied RBAC covers time constraints, hierarchies and + multi-level authorization rules for coping with the + multi-actor nature and the complexity of the application + domain. The DRIVE RBAC model clearly distinguishes between + static role assignment to users and dynamic allocation of + roles at session time. The paper, while focusing on the + authorization and access control approach, also describes + how the RBAC functions have been integrated in a trust + infrastructure including smart cards.}, + tags = {noTAG}, + clearance = {unclassified}, + timestap = {2008-05-29} +} + +@InProceedings{ ahn.ea:towards:2007, + author = {Gail-Joon Ahn and Hongxin Hu}, + title = {Towards realizing a formal \acs{rbac} model in real + systems}, + booktitle = PROC # { the seventh \acs{acm} symposium on Access + control models and technologies (\acs{sacmat})}, + year = 2007, + isbn = {978-1-59593-745-2}, + pages = {215--224}, + location = {Sophia Antipolis, France}, + doi = {10.1145/1266840.1266875}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {There still exists an open question on how formal models + can be fully realized in the system development phase. The + Model Driven Development (MDD) approach has been recently + introduced to deal with such a critical issue for building + high assurance software systems. + + There still exists an open question on how formal models + can be fully realized in the system development phase. The + Model Driven Development (MDD) approach has been recently + introduced to deal with such a critical issue for building + high assurance software systems. + + The MDD approach focuses on the transformation of + high-level design models to system implementation modules. + However, this emerging development approach lacks an + adequate procedure to address security issues derived from + formal security models. In this paper, we propose an + empirical framework to integrate security model + representation, security policy specification, and + systematic validation of security model and policy, which + would be eventually used for accommodating security + concerns during the system development. We also describe + how our framework can minimize the gap between security + models and the development of secure systems. In addition, + we overview a proof-of-concept prototype of our tool that + facilitates existing software engineering mechanisms to + achieve the above-mentioned features of our framework.}, + tags = {noTAG}, + clearance = {unclassified}, + timestap = {2008-05-29} +} + +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +%%%%%%%% un-checked entries +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +@Book{ woodcock.ea:using:1996, + abstract = {This book contains enough material for three complete + courses of study. It provides an introduction to the world + of logic, sets and relations. It explains the use of the + Znotation in the specification of realistic systems. It + shows how Z specifications may be refined to produce + executable code; this is demonstrated in a selection of + casestudies.This book is both authoritative and + comprehensive. It strikes the right balance between the + formality of mathematics and the practical needs of + industrial softwaredevelopment. It is faithful to the draft + ISO standard for Z. It covers the essentials of + specification, refinement, and proof, revealing techniques + never previouslypublished.}, + author = {Jim Woodcock and Jim Davies}, + price = {\$37.95}, + length = 391, + isbn = {0-13-948472-8}, + language = {USenglish}, + public = {yes}, + publisher = {Prentice Hall}, + title = {Using {Z}: {S}pecification, {R}efinement, and {P}roof}, + series = {Prentice Hall International Series in Computer Science}, + topic = {formalism}, + keywords = {formal methods, Z}, + library = {FAW (25/91): 91: CD/3.2/125}, + url = {http://softeng.comlab.ox.ac.uk/usingz/}, + year = 1996, + timestamp = 962966715, + acknowledgement={none}, + bibkey = {woodcock.ea:using:1996} +} + +@InProceedings{ dick.ea:testing:1993, + bibkey = {dick.ea:testing:1993}, + author = {Jeremy Dick and Alain Faivre}, + title = {Automating the Generation and Sequencing of Test Cases + from Model-Based Specications}, + pages = {268--284}, + booktitle = {Formal Methods Europe 93: Industrial-Strength Formal + Methods}, + editor = {J.C.P. Woodcock and P.G. Larsen}, + month = apr, + year = 1993, + volume = 670, + series = s-lncs, + publisher = pub-springer, + address = pub-springer:adr, + acknowledgement={none} +} + +@Article{ grieskamp.ea:instrumenting:2004, + author = {Wolfgang Grieskamp and Nikolai Tillmann and Margus + Veanes}, + booktitle = {Third International Conference on Quality Software: QSIC + 2003}, + title = {Instrumenting scenarios in a model-driven development + environment}, + journal = {Information and Software Technology}, + year = 2004, + number = 15, + volume = 46, + pages = {1027--1036}, + doi = {10.1016/j.infsof.2004.07.007}, + abstract = {SpecExplorer is an integrated environment for model-driven + development of .NET software. In this paper we discuss how + scenarios can be described in SpecExplorer's modeling + language, Spec#, and how the SpecExplorer tool can be used + to validate those scenarios by various means.}, + acknowledgement={none}, + bibkey = {grieskamp.ea:instrumenting:2004} +} + +@Book{ abrial:event-b:2009, + bibkey = {abrial:event-b:2009}, + title = {Modeling in Event-B: System and Software Design}, + publisher = {Cambridge University Press}, + year = 2009, + author = {Jean-Raymond Abrial}, + acknowledgement={none}, + address = {New York, \acs{ny}, \acs{usa}} +} + +@Book{ abrial:b-book:1996, + bibkey = {abrial:b-book:1996}, + author = {Jean-Raymond Abrial}, + title = {The {B-Book}: assigning programs to meanings}, + year = 1996, + isbn = {0-521-49619-5}, + publisher = {Cambridge University Press}, + acknowledgement={none}, + address = {New York, \acs{ny}, \acs{usa}} +} + +@TechReport{ leino.ea:escjava:2000, + abstract = {The Compaq Extended Static Checker for Java (ESC/Java) is + a programming tool that attempts to find common run-time + errors in Java programs by static analysis of the program + text. Users can control the amount and kinds of checking + that ESC/Java performs by annotating their programs with + specially formatted comments called pragmas. This manual is + designed to serve both as an introduction to ESC/Java and + as a reference manual. It starts by providing an overview + of ESC/Java through an illustrative example of its use and + a summary of its features, and then goes on to document all + the pragmas supported by ESC/Java and all the kinds of + warnings that it generates. Appendices provide a brief + sketch of ESC/Java's implementation, information about + obtaining ESC/Java, and some discussion of its + limitations}, + author = {K. Rustan M. Leino and Greg Nelson and James B. Saxe}, + institution = {Compaq Systems Research Center}, + language = {USenglish}, + month = oct, + public = {yes}, + title = {{\acs{esc}}/{Java} User's Manual}, + url = {http://gatekeeper.dec.com/pub/DEC/SRC/technical-notes/abstracts/src-tn-2000-002.html} + , + number = {SRC-2000-002}, + year = 2000, + keywords = {Java}, + topic = {tools}, + acknowledgement={none}, + bibkey = {leino.ea:escjava:2000} +} + +@Book{ kleppe.ea:mda:2003, + bibkey = {kleppe.ea:mda:2003}, + title = {\acs{mda} Explained. The Model Driven Architecture: + Practice and Promise}, + acknowledgement={none}, + publisher = {Addison-Wesley}, + year = 2003, + author = {Anneke Kleppe and Jos Warmer and Wim Bast} +} + +@Article{ schmidt:mde:2006, + author = {Douglas C. Schmidt}, + title = {Guest Editor's Introduction: Model-Driven Engineering}, + journal = {Computer}, + volume = 39, + number = 2, + acknowledgement={none}, + year = 2006, + issn = {0018-9162}, + pages = {25--31}, + doi = {10.1109/MC.2006.58}, + publisher = {IEEE Computer Society}, + address = {Los Alamitos, \acs{ca}, \acs{usa}} +} + +@InCollection{ gaudel:testing:1995, + author = {Marie Claude Gaudel}, + title = {Testing can be formal, too}, + year = 1995, + booktitle = {\textsc{tapsoft}'95: Theory and Practice of Software + Development}, + isbn = {3-540-59293-8}, + address = pub-springer:adr, + paddress = {Heidelberg, Germany}, + pages = {82--96}, + publisher = pub-springer, + series = s-lncs, + number = 915, + editor = {Peter D. Mosses and Mogens Nielsen and Michael I. + Schwartzbach}, + acknowledgement={none}, + bibkey = {gaudel:testing:1995} +} + +@InProceedings{ jurjens.ea:specification-based:2001, + author = {Jan J{\"u}rjens and Guido Wimmel}, + title = {Specification-Based Testing of Firewalls}, + booktitle = {Ershov Memorial Conference}, + year = 2001, + pages = {308--316}, + crossref = {bjorner.ea:perspectives:2001}, + acknowledgement={none} +} + +@Proceedings{ bjorner.ea:perspectives:2001, + editor = {Dines Bj{\o}rner and Manfred Broy and Alexandre V. + Zamulin}, + title = {Perspectives of System Informatics, 4th International + Andrei Ershov Memorial Conference, PSI 2001, Akademgorodok, + Novosibirsk, Russia, July 2-6, 2001, Revised Papers}, + booktitle = {Ershov Memorial Conference}, + publisher = pub-springer, + adress = pub-springer:adr, + series = s-lncs, + volume = 2244, + year = 2001, + isbn = {3-540-43075-X}, + acknowledgement={none} +} + +@InProceedings{ bishop.ea:rigorous:2005, + author = {Steve Bishop and Matthew Fairbairn and Michael Norrish and + Peter Sewell and Michael Smith and Keith Wansbrough}, + title = {Rigorous specification and conformance testing techniques + for network protocols, as applied to TCP, UDP, and + sockets}, + booktitle = {SIGCOMM}, + year = 2005, + pages = {265--276}, + doi = {10.1145/1080091.1080123}, + crossref = {guerin.ea:proceedings:2005}, + acknowledgement={none} +} + +@Proceedings{ guerin.ea:proceedings:2005, + editor = {Roch Gu{\'e}rin and Ramesh Govindan and Greg Minshall}, + title = PROC # { the ACM SIGCOMM 2005 Conference on + Applications, Technologies, Architectures, and Protocols + for Computer Communications, Philadelphia, Pennsylvania, + \acs{usa}, August 22-26, 2005}, + booktitle = {SIGCOMM}, + publisher = pub-acm, + adress = pub-acm:adr, + year = 2005, + isbn = {1-59593-009-4}, + acknowledgement={none} +} + +@InProceedings{ senn.ea:firewall:2005, + abstract = {Firewalls are widely used to protect networks from + unauthorised access. To ensure that they implement an + organisation's security policy correctly, they need to be + tested. We present an approach that addresses this problem. + Namely, we show how an organisation's network security + policy can be formally specified in a high-level way, and + how this specification can be used to automatically + generate test cases to test a deployed system. In contrast + to other firewall testing methodologies, such as + penetration testing, our approach tests conformance to a + specified policy. Our test cases are organisation-specific + --- i.e.~they depend on the security requirements and on + the network topology ofan organisation --- and can uncover + errors both in the firewall products themselves and in + their configuration.}, + author = {Diana Senn and David A. Basin and Germano Caronni}, + booktitle = {TestCom 2005}, + editor = {Ferhat Khendek and Rachida Dssouli}, + isbn = {3-540-26054-4}, + language = {UKenglish}, + month = {May}, + pages = {226--241}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + title = {Firewall Conformance Testing}, + volume = 3502, + year = 2005, + acknowledgement={none} +} + +@PhDThesis{ bidder:specification:2007, + author = {Diana von Bidder}, + title = {Specification-based Firewall Testing}, + school = {ETH Zurich}, + year = 2007, + public = {yes}, + type = {Ph.D. Thesis}, + acknowledgement={none}, + classification= {thesis}, + note = {\acs{eth} Dissertation No. 17172. Diana von Bidder's + maiden name is Diana Senn.} +} + +@InCollection{ wenzel.ea:building:2007, + abstract = {We present the generic system framework of + Isabelle/Isarunderlying recent versions of Isabelle. Among + other things, Isar provides an infrastructure for Isabelle + plug-ins, comprising extensible state components and + extensible syntax that can be bound to tactical ML + programs. Thus the Isabelle/Isar architecture may be + understood as an extension and refinement of the + traditional LCF approach, with explicit infrastructure for + building derivative systems. To demonstrate the technical + potential of the framework, we apply it to a concrete + formalmethods tool: the HOL-Z 3.0 environment, which is + geared towards the analysis of Z specifications and formal + proof of forward-refinements.}, + author = {Makarius Wenzel and Burkhart Wolff}, + booktitle = {\acs{tphols} 2007}, + editor = {Klaus Schneider and Jens Brandt}, + language = {USenglish}, + acknowledgement={none}, + pages = {352--367}, + publisher = pub-springer, + address = pub-springer:adr, + number = 4732, + series = s-lncs, + title = {Building Formal Method Tools in the {Isabelle}/{Isar} + Framework}, + doi = {10.1007/978-3-540-74591-4_26}, + year = 2007 +} + +@Article{ igarashi.ea:featherweight:2001, + author = {Atsushi Igarashi and Benjamin C. Pierce and Philip + Wadler}, + title = {{Featherweight Java}: a minimal core calculus for {Java} + and {\acs{gj}}}, + journal = j-toplas, + volume = 23, + number = 3, + year = 2001, + issn = {0164-0925}, + pages = {396--450}, + doi = {10.1145/503502.503505}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none}, + abstract = {Several recent studies have introduced lightweight + versions of Java: reduced languages in which complex + features like threads and reflection are dropped to enable + rigorous arguments about key properties such as type + safety. We carry this process a step further, omitting + almost all features of the full language (including + interfaces and even assignment) to obtain a small calculus, + Featherweight Java, for which rigorous proofs are not only + possible but easy. Featherweight Java bears a similar + relation to Java as the lambda-calculus does to languages + such as ML and Haskell. It offers a similar computational + "feel," providing classes, methods, fields, inheritance, + and dynamic typecasts with a semantics closely following + Java's. A proof of type safety for Featherweight Java thus + illustrates many of the interesting features of a safety + proof for the full language, while remaining pleasingly + compact. The minimal syntax, typing rules, and operational + semantics of Featherweight Java make it a handy tool for + studying the consequences of extensions and variations. As + an illustration of its utility in this regard, we extend + Featherweight Java with generic classes in the style of GJ + (Bracha, Odersky, Stoutamire, and Wadler) and give a + detailed proof of type safety. The extended system + formalizes for the first time some of the key features of + GJ. }, + bibkey = {igarashi.ea:featherweight:2001} +} + +@Article{ zhu.ea:software:29-4, + title = {Software Unit Test Coverage and Adequacy}, + author = {Hong Zhu and Patrick A.V. Hall and John H. R. May}, + journal = {ACM Computing Surveys}, + issn = {0360-0300}, + volume = 29, + url = {http://www.cs.bris.ac.uk/Tools/Reports/Abstracts/1997-zhu.html} + , + number = 4, + language = {USenglish}, + pages = {366--427}, + month = dec, + keywords = {Safety_Critical_Systems}, + year = 1997, + abstract = {Objective measurement of test quality is one of the key + issues in software testing. It has been a major research + focus for the last two decades. Many test criteria have + been proposed and studied for this purpose. Various kinds + of rationale have been presented in support of one + criterion or another. This paper surveys the research work + in this area. The notion of adequacy criteria is examined + together with its role in software dynamic testing. A + review of criteria classification is followed by a summary + of the methods for comparison and assessment of criteria.}, + acknowledgement={none}, + bibkey = {zhu.ea:software:29-4} +} + +@PhDThesis{ wenzel:isabelleisar:2002, + author = {Markus M. Wenzel}, + title = {Isabelle/Isar --- a versatile environment for + human-readable formal proof documents}, + school = {TU M{\"u}nchen}, + year = 2002, + url = {http://tumb1.biblio.tu-muenchen.de/publ/diss/in/2002/wenzel.html} + , + abstract = {The basic motivation of this work is to make formal theory + developments with machine-checked proofs accessible to a + broader audience. Our particular approach is centered + around the Isar formal proof language that is intended to + support adequate composition of proof documents that are + suitable for human consumption. Such primary proofs written + in Isar may be both checked by the machine and read by + human-beings; final presentation merely involves trivial + pretty printing of the sources. Sound logical foundations + of Isar are achieved by interpretation within the generic + Natural Deduction framework of Isabelle, reducing all + high-level reasoning steps to primitive inferences. + + The resulting Isabelle/Isar system is generic with respect + to object-logics and proof tools, just as pure Isabelle + itself. The full Isar language emerges from a small core by + means of several derived elements, which may be combined + freely with existing ones. This results in a very rich + space of expressions of formal reasoning, supporting many + viable proof techniques. The general paradigms of Natural + Deduction and Calculational Reasoning are both covered + particularly well. Concrete examples from logic, + mathematics, and computer-science demonstrate that the Isar + concepts are indeed sufficiently versatile to cover a broad + range of applications.}, + address = {M{\"u}nchen}, + month = feb, + acknowledgement={none}, + bibkey = {wenzel:isabelleisar:2002} +} + +@InProceedings{ frantzen.ea:test:2004, + author = {L. Frantzen and J. Tretmans and T.A.C. Willemse}, + title = {Test Generation Based on Symbolic Specifications}, + booktitle = {FATES 2004}, + year = 2004, + month = sep, + abstract = {Classical state-oriented testing approaches are based on + sim- ple machine models such as Labelled Transition Systems + (LTSs), in which data is represented by concrete values. To + implement these theories, data types which have infinite + universes have to be cut down to infinite vari- ants, which + are subsequently enumerated to fit in the model. This leads + to an explosion of the state space. Moreover, exploiting + the syntactical and/or semantical information of the + involved data types is non-trivial after enumeration. To + overcome these problems, we lift the family of test- ing + relations iocoF to the level of Symbolic Transition Systems + (STSs). We present an algorithm based on STSs, which + generates and executes tests on-the-fly on a given system. + It is sound and complete for the ioco F testing + relations.}, + acknowledgement={none}, + bibkey = {frantzen.ea:test:2004} +} + +@Book{ dagostino.ea:handbook:1996, + title = {Handbook of Tableau Methods}, + editor = {Marcello D'Agostino and Dov Gabbay and Reiner H\"ahnle and + Joachim Posegga}, + publisher = {Kluwer, Dordrecht}, + year = 1996, + isbn = {0-7923-5627-6}, + acknowledgement={none}, + bibkey = {dagostino.ea:handbook:1996} +} + +@Article{ visser.ea:test:2004, + author = {Willem Visser and Corina S. P\u{a}s\u{a}reanu and Sarfraz + Khurshid}, + title = {Test input generation with {Java} {PathFinder}}, + journal = {SIGSOFT Softw. Eng. Notes}, + volume = 29, + number = 4, + year = 2004, + issn = {0163-5948}, + pages = {97--107}, + doi = {10.1145/1013886.1007526}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none}, + bibkey = {visser.ea:test:2004} +} + +@InProceedings{ pons.ea:practical:2006, + author = {Claudia Pons and Diego Garcia}, + title = {Practical Verification Strategy for Refinement Conditions + in \acs{uml} Models}, + booktitle = {IFIP Workshop on Advanced Software Engineering}, + year = 2006, + pages = {47--61}, + doi = {10.1007/978-0-387-34831-5_4}, + crossref = {ochoa.ea:ifip:2006} +} + +@Proceedings{ ochoa.ea:ifip:2006, + editor = {Sergio F. Ochoa and Gruia-Catalin Roman}, + title = {IFIP 19th World Computer Congress, First International + Workshop on Advanced Software Engineering, Expanding the + Frontiers of Software Technology, August 25, 2006, + Santiago, Chile}, + booktitle = {IFIP Workshop on Advanced Software Engineering}, + publisher = pub-springer, + address = pub-springer:adr, + series = {IFIP}, + volume = 219, + year = 2006, + isbn = {978-0-387-34828-5} +} + +@Misc{ clearsy:atelier-b:2008, + author = {{Clearsy Inc.}}, + title = {{Atelier B}}, + year = 2008, + note = {\url{http://www.atelierb.eu/}} +} + +@Book{ beckert.ea:key:2007, + editor = {Bernhard Beckert and Reiner H\"ahnle and Peter H. Schmitt}, + title = {Verification of Object-Oriented Software: The {\KeY} + Approach}, + volume = 4334, + series = s-lncs, + doi = {10.1007/978-3-540-69061-0}, + publisher = pub-springer, + address = pub-springer:adr, + year = 2007 +} + +@InProceedings{ rudich.ea:checking:2008, + author = {Arsenii Rudich and {\'A}d{\'a}m Darvas and Peter M{\"u}ller}, + title = {Checking Well-Formedness of Pure-Method Specifications}, + booktitle = {FM}, + year = 2008, + pages = {68--83}, + ee = {http://dx.doi.org/10.1007/978-3-540-68237-0_7}, + crossref = {cuellar.ea:fm:2008} +} + +@Proceedings{ cuellar.ea:fm:2008, + editor = {Jorge Cu{\'e}llar and T. S. E. Maibaum and Kaisa Sere}, + title = {FM 2008: Formal Methods, 15th International Symposium on + Formal Methods, Turku, Finland, May 26-30, 2008, + Proceedings}, + booktitle = {FM}, + series = s-lncs, + publisher = pub-springer, + address = pub-springer:adr, + volume = 5014, + year = 2008, + isbn = {978-3-540-68235-6} +} + +@InProceedings{ ledang:b-based:2004, + author = {Hung Ledang}, + title = {B-based Conistency Checking of UML Diagrams}, + booktitle = PROC # { ICT/RDA2004 : the 2nd National Symposium on + Research, Development and Application of Information and + Communication TechnologyAIR'2005}, + year = 2004, + publisher = {Science and Techniques Publisher} +} + +@InProceedings{ krieger.ea:executing, + author = {Matthias P. Krieger and Alexander Knapp.}, + title = {Executing Underspecified OCL Operation Contracts with a + SAT Solver}, + booktitle = PROC # { the OCL 2008 Workshop}, + editor = {to appear}, + note = {\url{http://www.fots.ua.ac.be/events/ocl2008/}} +} + +@InProceedings{ tretmans.ea:cote:2002, + author = {Jan Tretmans and Edsgar Brinksma}, + title = {C\^ote de Resyste --- Automated Model Based Testing}, + booktitle = {Progress 2002 --- 3rd Workshop on Embedded Systems}, + pages = {246--255}, + year = 2002 +} + +@Article{ tretmans:test:1996, + author = {Jan Tretmans}, + title = {Test Generation with Inputs, Outputs and Repetitive + Quiescence}, + journal = {Software --- Concepts and Tools}, + year = 1996, + volume = 17, + number = 3, + pages = {103--120} +} + +@Article{ jard.ea:tgv:2005, + author = {C. Jard and T. J\'eron}, + title = {TGV: Theory, Principles and Algorithms}, + journal = {Software Tools for Technology Transfer}, + year = 2005, + volume = 7, + number = 4, + pages = {297--315} +} + +@InProceedings{ clarke.ea:stg:2002, + author = {D. Clarke and T. J\'eron and V. Rusu and E. Zinovieva}, + title = {STG: A Symbolic Test Generation Tool}, + pages = {470--475}, + year = 2002, + series = s-lncs, + publisher = pub-springer, + address = pub-springer:adr, + volume = {2280} +} + +@InProceedings{ koch.ea:autolink:1998, + author = {B. Koch and J. Grabowski and D. Hogrefe and M. Schmitt}, + title = {AutoLink --- a Tool for Automatic Test Generation from SDL + Specifications}, + booktitle = {Proc. IEEE Intl. Workshop on Industrial Strength Formal + Specification Techniques (WIFT 1998)}, + pages = {114--127}, + year = 1998 +} + +@InProceedings{ bouquet.ea:mastering:2004, + author = {F. Bouquet and B. Legeard and F. Peureux and E. + Torreborre}, + title = {Mastering Test Generation from Smart Card Software Formal + Models}, + booktitle = {Proc. Intl. Workshop on Construction and Analysis of Safe, + Secure and Interoperable Smart devices}, + pages = {70--85}, + year = 2004, + series = s-lncs, + publisher = pub-springer, + address = pub-springer:adr, + volume = {3362} +} + +@InProceedings{ prowell:tool:2003, + author = {S. Prowell}, + title = {A Tool for Model-based Statistical Testing}, + booktitle = {Proc. HICSS'03, IEEE}, + pages = {337.3}, + year = 2003 +} + +@Article{ walton.ea:generating:2000, + author = {G. Walton and J. Poore}, + title = {Generating transition probabilities to support model-based + software testing}, + journal = {Software: Practice and Experience}, + year = 2000, + volume = 30, + number = 10, + pages = {1095--1106} +} + +@Article{ cohen.ea:aetg:1997, + author = {D. Cohen and S. Dalal and M. Fredman and G. Patton}, + title = {The AETG System: An approach to testing Based on + Combinatorial Design}, + journal = {IEEE Transactions on Software Engineering}, + year = 1997, + volume = 23, + number = 7 +} + +@InProceedings{ bohme.ea:hol-boogie:2008, + author = {Sascha B{\"o}hme and K. Rustan M. Leino and Burkhart + Wolff}, + title = {{\acs{hol}-Boogie}---An Interactive Prover for the + {Boogie} Program-Verifier}, + booktitle = {Theorem Proving in Higher Order Logics}, + year = 2008, + pages = {150--166}, + doi = {10.1007/978-3-540-71067-7_15}, + crossref = {otmane.ea:tphols:2008} +} + +@Proceedings{ otmane.ea:tphols:2008, + editor = {Otmane A\"{\i}t Mohamed and C{\'e}sar Mu{\~n}oz and + Sofi{\`e}ne Tahar}, + title = {Theorem Proving in Higher Order Logics (\acs{tphols})}, + booktitle = {Theorem Proving in Higher Order Logics (\acs{tphols})}, + location = {Montreal, Canada}, + month = aug, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 5170, + year = 2008, + isbn = {978-3-540-71065-3} +} + +@InProceedings{ huisman.ea:inheritance:2000, + author = {Marieke Huisman and Bart Jacobs}, + title = {Inheritance in Higher Order Logic: Modeling and + Reasoning}, + doi = {10.1007/3-540-44659-1_19}, + year = 2000, + pages = {301--319}, + crossref = {aagaard.ea:tphols:2000} +} + +@Proceedings{ aagaard.ea:tphols:2000, + editor = {Mark Aagaard and John Harrison}, + location = {Portland, Oregon, USA}, + month = aug, + title = {Theorem Proving in Higher Order Logics (\acs{tphols})}, + booktitle = {Theorem Proving in Higher Order Logics (\acs{tphols})}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 1869, + year = 2000, + isbn = {3-540-67863-8} +} + +@Book{ roscoe:csp:1998, + author = {A.W. Roscoe}, + title = {Theory and Practice of Concurrency}, + publisher = {Prentice Hall}, + year = 1998, + isbn = {0-13-674409-5} +} + +@Article{ foster:error:1980, + title = {Error Sensitive Test Cases Analysis (ESTCA)}, + author = {Foster, K.A.}, + journal = {Software Engineering, IEEE Transactions on}, + year = 1980, + month = {May}, + volume = {SE-6}, + number = 3, + pages = {258--264}, + abstract = {A hardware failure analysis technique adapted to software + yielded three rules for generating test cases sensitive to + code errors. These rules, and a procedure for generating + these cases, are given with examples. Areas for further + study are recommended.}, + keywords = {null Program correctness, progran testing, software + errors, software reliability, test data generation}, + issn = {0098-5589} +} + +@Book{ myers.ea:art:2004, + author = {Glenford J. Myers and Corey Sandler}, + title = {The Art of Software Testing}, + year = 2004, + isbn = 0471469122, + publisher = {John Wiley \& Sons} +} + +@InProceedings{ tillmann.ea:pex:2008, + author = {Nikolai Tillmann and Jonathan de Halleux}, + title = {{Pex}---White Box Test Generation for {.NET}}, + booktitle = {TAP}, + year = 2008, + pages = {134-153}, + doi = {10.1007/978-3-540-79124-9_10}, + crossref = {beckert.ea:tests:2008}, + abstract = {Pex automatically produces a small test suite with high + code coverage for a .NET program. To this end, Pex performs + a systematic program analysis (using dynamic symbolic + execution, similar to path-bounded model-checking) to + determine test inputs for Parameterized Unit Tests. Pex + learns the program behavior by monitoring execution traces. + Pex uses a constraint solver to produce new test inputs + which exercise different program behavior. The result is an + automatically generated small test suite which often + achieves high code coverage. In one case study, we applied + Pex to a core component of the .NET runtime which had + already been extensively tested over several years. Pex + found errors, including a serious issue.} +} + +@InProceedings{ halleux.ea:parameterized:2008, + author = {Jonathan de Halleux and Nikolai Tillmann}, + title = {Parameterized Unit Testing with {Pex}}, + booktitle = {TAP}, + year = 2008, + pages = {171--181}, + doi = {10.1007/978-3-540-79124-9_12}, + crossref = {beckert.ea:tests:2008}, + abstract = {This hands-on tutorial will teach the principles of + Parameterized Unit Testing [5,4] with Pex [2], an automatic + test input generator for .NET which performs a systematic + program analysis, similar to path bounded model-checking. A + parameterized unit test is simply a method that takes + parameters, calls the code under test, and states + assertions.} +} + +@Proceedings{ beckert.ea:tests:2008, + editor = {Bernhard Beckert and Reiner H{\"a}hnle}, + title = {Tests and Proofs, Second International Conference, TAP + 2008, Prato, Italy, April 9-11, 2008. Proceedings}, + booktitle = {TAP}, + series = s-lncs, + publisher = pub-springer, + address = pub-springer:adr, + volume = 4966, + year = 2008, + isbn = {978-3-540-79123-2} +} + +%%%%%% +@InProceedings{ povey:enforcing:1999, + author = {Dean Povey}, + title = {Enforcing well-formed and partially-formed transactions + for {Unix}}, + booktitle = PROC # { the 8th conference on \acs{usenix} Security + Symposium}, + volume = 8, + year = 1999, + publisher = {\acs{usenix} Association}, + location = {Berkeley, \acs{ca}}, + pages = {5--5} +} + +@InProceedings{ longstaff.ea:model:2000, + author = {J.J. Longstaff and M.A. Lockyer and M.G. Thick}, + title = {A Model of Accountability, Confidentiality and Override + for Healthcare and other Applications}, + booktitle = PROC # { the fifth \acs{acm} workshop on Role-based + access control}, + year = 2000, + isbn = {1-58113-259-X}, + pages = {71--76}, + publisher = pub-acm, + address = pub-acm:adr, + doi = {10.1145/344287.344304} +} + +@InProceedings{ rissanen.ea:towards:2004, + author = {Erik Rissanen}, + title = {Towards a Mechanism for Discretionary Overriding of Access + Control (Transcript of Discussion)}, + booktitle = PROC # { the 12th International Workshop on Security + Protocols}, + year = 2004, + pages = {320--323}, + month = mar, + doi = {10.1007/11861386_39}, + abstract = {Last year, the Swedish Prime Minister was stabbed to death + in a shopping mall in Stockholm, and of course the police + thoroughly investigated it. They had some privacy problems + during the investigation: many policemen just looked at the + case, because there was no access control on the police + system. They didn{\^a}\80\99t have a whole system on-line, + because they cannot really predict the needs of individual + policemen, and they cannot really audit the whole thing + either because there were so many accesses. In the case of + the prime minister we suspect that something was going on + because he was a famous person, and they know from + experience that this tends to happen with famous people, + but in the case of a policemen accessing his + neighbour{\^a}\80\99s data, or something like that, then there + is little reason to notice that something is going on.}, + crossref = {bruce.ea:security:2006} +} + +@Proceedings{ bruce.ea:security:2006, + series = s-lncs, + publisher = pub-springer, + address = pub-springer:adr, + editor = {Bruce Christianson and Bruno Crispo and James A. Malcolm + and Michael Roe}, + title = {Security Protocols, 12th International Workshop, + Cambridge, UK, April 26-28, 2004. Revised Selected Papers}, + booktitle = {Security Protocols Workshop}, + volume = 3957, + year = 2006, + isbn = {3-540-40925-4} +} + +@InProceedings{ rissanen.ea:discretionary:2004, + author = {Erik Rissanen and Babak Sadighi Firozabadi and Marek J. + Sergot}, + title = {Discretionary Overriding of Access Control in the + Privilege Calculus}, + booktitle = PROC # { the Workshop on Formal Aspects Security and + Trust (\acs{fast})}, + year = 2004, + pages = {219--232}, + doi = {10.1007/0-387-24098-5_16}, + crossref = {dimitrakos.ea:formal:2005}, + abstract = {We extend a particular access control framework, the + Privilege Calculus, with a possibility to override denied + access for increased flexibility in hard to define or + unanticipated situations. We require the overrides to be + audited and approved by appropriate managers. In order to + automatically find the authorities who are able to approve + an override, we present an algorithm for authority + resolution. We are able to calculate from the access + control policy who can approve an override without the need + for any additional information.} +} + +@Proceedings{ dimitrakos.ea:formal:2005, + editor = {Theodosis Dimitrakos and Fabio Martinelli}, + title = {Formal Aspects in Security and Trust: Second IFIP TC1 + WG1.7 Workshop on Formal Aspects in Security and Trust + (FAST), an event of the 18th IFIP World Computer Congress, + August 22-27, 2004, Toulouse, France}, + booktitle = {Formal Aspects in Security and Trust}, + publisher = pub-springer, + volume = 173, + address = pub-springer:adr, + year = 2005, + isbn = {0-387-24050-0} +} + +@InProceedings{ alqatawna.ea:overriding:2007, + author = {Ja'far Alqatawna and Erik Rissanen and Babak Sadighi}, + title = {Overriding of Access Control in \textsc{XACML}}, + booktitle = PROC # { the Eighth \acs{ieee} International + Workshop on Policies for Distributed Systems and Networks + (\acs{policy})}, + year = 2007, + isbn = {0-7695-2767-1}, + pages = {87--95}, + doi = {10.1109/POLICY.2007.31}, + address = pub-ieee:adr, + publisher = pub-ieee +} + +@InProceedings{ stevens.ea:new:2002, + author = {Gunnar Stevens and Volker Wulf}, + title = {A new dimension in access control: studying maintenance + engineering across organizational boundaries}, + booktitle = PROC # { the \acs{acm} conference on Computer + supported cooperative work (\acs{cscw})}, + year = 2002, + isbn = {1-58113-560-2}, + pages = {196--205}, + location = {New Orleans, Louisiana, USA}, + doi = {10.1145/587078.587106}, + publisher = pub-acm, + address = pub-acm:adr +} + +@InProceedings{ jaeger.ea:managing:2002, + author = {Trent Jaeger and Antony Edwards and Xiaolan Zhang}, + title = {Managing access control policies using access control + spaces}, + booktitle = PROC # { the seventh \acs{acm} symposium on Access + control models and technologies (\acs{sacmat})}, + year = 2002, + isbn = {1-58113-496-7}, + pages = {3--12}, + location = {Monterey, California, USA}, + doi = {10.1145/507711.507713}, + publisher = pub-acm, + address = pub-acm:adr +} + +@Article{ joshi.ea:generalized:2005, + author = {James B.D. Joshi and Elisa Bertino and Usman Latif and + Arif Ghafoor}, + title = {A Generalized Temporal Role-Based Access Control Model}, + journal = j-tkde, + volume = 17, + number = 1, + issn = {1041-4347}, + year = 2005, + pages = {4--23}, + doi = {10.1109/TKDE.2005.1}, + publisher = pub-ieee, + address = pub-ieee:adr +} + +@InProceedings{ bell.ea:secure:1996, + author = {D. Elliott Bell and Leonard J. LaPadula}, + title = {Secure Computer Systems: A Mathematical Model, Volume + {II}}, + booktitle = {Journal of Computer Security 4}, + year = 1996, + pages = {229--263}, + note = {An electronic reconstruction of \emph{Secure Computer + Systems: Mathematical Foundations}, 1973} +} + +@InProceedings{ bell:looking:2005, + title = {Looking Back at the Bell-La Padula Model}, + author = {D. Elliott Bell}, + journal = PROC # { the 21st Annual Computer Security + Applications Conference}, + year = 2005, + isbn = {1063-9527}, + doi = {10.1109/CSAC.2005.37}, + publisher = {pub-ieee}, + address = pub-ieee:adr, + pages = {337--351} +} + +@Booklet{ oasis:xacml:2005, + title = {{eXtensible Access Control Markup Language (XACML)}, + Version 2.0}, + year = 2005, + url = {http://docs.oasis-open.org/xacml/2.0/XACML-2.0-OS-NORMATIVE.zip} + , + bibkey = {oasis:xacml:2005}, + publisher = {\acs{oases}}, + key = {OASIS}, + language = {USenglish}, + public = {yes} +} + +@InProceedings{ barka.ea:framework:2000, + author = {Ezedin Barka and Ravi Sandhu}, + title = {Framework for Role-based Delegation Models}, + year = 2000, + booktitle = PROC # { the 16th Annual Computer Security + Applications Conference}, + doi = {10.1109/ACSAC.2000.898870}, + isbn = {0-7695-0859-6}, + pages = {168--176}, + publisher = pub-ieee, + address = pub-ieee:adr +} + +@InProceedings{ cheng.ea:fuzzy:2007, + author = {Pau-Chen Cheng and Pankaj Rohatgi and Claudia Keser and + Paul A. Karger and Grant M. Wagner and Angela Schuett + Reninger}, + title = {Fuzzy Multi-Level Security: An Experiment on Quantified + Risk-Adaptive Access Control}, + booktitle = {IEEE Symposium on Security and Privacy}, + year = 2007, + pages = {222--230}, + ee = {http://dx.doi.org/10.1109/SP.2007.21}, + crossref = {ieee:security-privacy:2007} +} + +@Proceedings{ ieee:security-privacy:2007, + title = {2007 IEEE Symposium on Security and Privacy (S{\&}P 2007), + 20-23 May 2007, Oakland, California, USA}, + booktitle = {IEEE Symposium on Security and Privacy}, + publisher = {IEEE Computer Society}, + year = 2007 +} + +@InProceedings{ zhang.ea:toward:2006, + author = {Lei Zhang and Alexander Brodsky and Sushil Jajodia}, + title = {Toward Information Sharing: Benefit And Risk Access + Control (BARAC)}, + booktitle = {POLICY}, + year = 2006, + pages = {45--53}, + doi = {10.1109/POLICY.2006.36}, + crossref = {ieee:policy:2006} +} + +@Proceedings{ ieee:policy:2006, + title = {7th IEEE International Workshop on Policies for + Distributed Systems and Networks (POLICY 2006), 5-7 June + 2006, London, Ontario, Canada}, + booktitle = {POLICY}, + publisher = {IEEE Computer Society}, + year = 2006, + isbn = {0-7695-2598-9} +} + +@InProceedings{ nissanke.ea:risk:2004, + author = {Nimal Nissanke and Etienne J. Khayat}, + title = {Risk Based Security Analysis of Permissions in RBAC}, + booktitle = {WOSIS}, + year = 2004, + pages = {332--341}, + crossref = {fernandez-medina.ea:security:2004} +} + +@Proceedings{ fernandez-medina.ea:security:2004, + editor = {Eduardo Fern{\'a}ndez-Medina and Julio C{\'e}sar + Hern{\'a}ndez Castro and L. Javier Garc\'{\i}a-Villalba}, + title = {Security In Information Systems, Proceedings of the 2nd + International Workshop on Security In Information Systems, + WOSIS 2004, In conjunction with ICEIS 2004, Porto, + Portugal, April 2004}, + booktitle = {WOSIS}, + publisher = {INSTICC Press}, + year = 2004, + isbn = {972-8865-07-4} +} + +@InProceedings{ fisler.ea:verification:2005, + author = {Kathi Fisler and Shriram Krishnamurthi and Leo A. + Meyerovich and Michael Carl Tschantz}, + title = {Verification and change-impact analysis of access-control + policies}, + booktitle = {ICSE}, + year = 2005, + pages = {196--205}, + doi = {10.1145/1062455.1062502}, + crossref = {roman.ea:27th:2005} +} + +@Proceedings{ roman.ea:27th:2005, + editor = {Gruia-Catalin Roman and William G. Griswold and Bashar + Nuseibeh}, + title = {27th International Conference on Software Engineering + (ICSE 2005), 15-21 May 2005, St. Louis, Missouri, USA}, + booktitle = {ICSE}, + address = pub-acm:adr, + publisher = pub-acm, + year = 2005 +} + +@InProceedings{ lin.ea:approach:2007, + author = {Dan Lin and Prathima Rao and Elisa Bertino and Jorge + Lobo}, + title = {An approach to evaluate policy similarity}, + booktitle = {SACMAT}, + year = 2007, + pages = {1--10}, + doi = {10.1145/1266840.1266842}, + crossref = {lotz.ea:sacmat:2007} +} + +@InProceedings{ backes.ea:efficient:2004, + author = {Michael Backes and G{\"u}nter Karjoth and Walid Bagga and + Matthias Schunter}, + title = {Efficient comparison of enterprise privacy policies}, + booktitle = {SAC}, + year = 2004, + pages = {375--382}, + doi = {10.1145/967900.967983}, + crossref = {haddad.ea:proceedings:2004} +} + +@Proceedings{ haddad.ea:proceedings:2004, + editor = {Hisham Haddad and Andrea Omicini and Roger L. Wainwright + and Lorie M. Liebrock}, + title = PROC # { the 2004 ACM Symposium on Applied Computing + (SAC), Nicosia, Cyprus, March 14-17, 2004}, + booktitle = {SAC}, + address = pub-acm:adr, + publisher = pub-acm, + year = 2004, + isbn = {1-58113-812-1} +} + +@InProceedings{ warner.ea:using:2007, + author = {Janice Warner and Vijayalakshmi Atluri and Ravi Mukkamala + and Jaideep Vaidya}, + title = {Using semantics for automatic enforcement of access + control policies among dynamic coalitions}, + booktitle = {SACMAT}, + year = 2007, + pages = {235--244}, + doi = {10.1145/1266840.1266877}, + crossref = {lotz.ea:sacmat:2007} +} + +@Proceedings{ lotz.ea:sacmat:2007, + editor = {Volkmar Lotz and Bhavani M. Thuraisingham}, + title = {SACMAT 2007, 12th ACM Symposium on Access Control Models + and Technologies, Sophia Antipolis, France, June 20-22, + 2007, Proceedings}, + booktitle = {SACMAT}, + address = pub-acm:adr, + publisher = pub-acm, + year = 2007, + isbn = {978-1-59593-745-2} +} + +@InProceedings{ povey:optimistic:1999, + author = {Dean Povey}, + title = {Optimistic Security: A New Access Control Paradigm}, + booktitle = PROC # { the 1999 workshop on New security + paradigms}, + year = 1999, + isbn = {1-58113-149-6}, + pages = {40--45}, + doi = {10.1145/335169.335188}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {Despite the best efforts of security researchers, + sometimes the static nature of authorisation can cause + unexpected risks for users work- ing in a dynamically + changing environment. Disasters, medical emergencies or + time-critical events can all lead to situations where the + ability to relax normal access rules can become critically + impor- tant. + + This paper presents an optimistic access control scheme + where en- forcement of rules is retrospective. The system + administrator is re- lied on to ensure that the system is + not misused, and compensating transactions are used to + ensure that the system integrity can be re- covered in the + case of a breach. It is argued that providing an opti- + mistic scheme alongside a traditional access control + mechanism can provide a useful means for users to exceed + their normal privileges on the rare occasion that the + situation warrants it. + + The idea of a partially-formed transaction is introduced to + show how accesses in an optimistic system might be + constrained. This model is formally described and related + to the Clark-Wilson in- tegrity model.} +} + +@Article{ sandhu.ea:role-based:1996, + author = {Ravi S. Sandhu and Edward J. Coyne and Hal L. Feinstein + and Charles E. Youman}, + title = {Role-Based Access Control Models}, + journal = j-computer, + year = 1996, + volume = 29, + number = 2, + address = pub-ieee:adr, + publisher = pub-ieee, + pages = {38--47}, + url = {http://ite.gmu.edu/list/journals/computer/pdf_ver/i94rbac(org).pdf} + , + abstract = {Abstract This article introduces a family of reference + models for rolebased acce ss control (RBAC) in which + permissions are associated with roles, and users are made + members of appropriate roles. This greatly simplifies + management of permiss ions. Roles are closely related to + the concept of user groups in access control. However, a + role brings together a set of users on one side and a set + of permiss ions on the other, whereas user groups are + typically defined as a set of users o nly. + + The basic concepts of RBAC originated with early multi-user + computer systems. Th e resurgence of interest in RBAC has + been driven by the need for general-purpose customizable + facilities for RBAC and the need to manage the + administration of R BAC itself. As a consequence RBAC + facilities range from simple to complex. This article + describes a novel framework of reference models to + systematically addres s the diverse components of RBAC, and + their interactions.}, + issn = {0018-9162}, + keywords = {Computational linguistics; Computer control systems; + Computer simulation; Computer software; Data abstraction; + Database systems; Discretionary access control; Encoding + (symbols); Integration; Mandator access control; Role based + access control; Semantics; Software encoding; User + interfaces}, + acknowledgement={none}, + bibkey = {sandhu.ea:role-based:1996} +} + +@Booklet{ sarbanes.ea:sox:2002, + title = {{Sarbanes-Oxley} {Act} of 2002}, + author = {P. Sarbanes and G. Oxley and others}, + howpublished = {107th Congress Report, House of Representatives, 2nd + Session, 107--610}, + year = 2002 +} + +@TechReport{ bcbs:baselii:2004, + author = {{Basel Committee on Banking Supervision}}, + title = {{Basel II}: International Convergence of Capital + Measurement and Capital Standards}, + year = 2004, + url = {http://www.bis.org/publ/bcbsca.htm}, + address = {Basel, Switzerland}, + institution = {Bank for International Settlements} +} + +@Book{ dahl.ea:structured:1972, + author = {O.-J. Dahl and E. W. Dijkstra and C. A. R. Hoare}, + title = {Structured Programming}, + publisher = {Academic Press}, + year = 1972, + edition = {3rd}, + volume = 8, + series = {A.P.I.C. Studies in Data Processing}, + address = {London}, + isbn = {0-12-200550-3} +} + +@InProceedings{ bryans:reasoning:2005, + author = {Jery Bryans}, + title = {Reasoning about {XACML} policies using {CSP}}, + booktitle = {SWS '05: Proceedings of the 2005 workshop on Secure web + services}, + year = 2005, + isbn = {1-59593-234-8}, + pages = {28--35}, + location = {Fairfax, VA, USA}, + doi = {10.1145/1103022.1103028}, + address = pub-acm:adr, + publisher = pub-acm +} + +@InProceedings{ chadwick.ea:permis:2002, + author = {David W. Chadwick and Alexander Otenko}, + title = {The \acs{permis} {X.509} role based privilege management + infrastructure}, + booktitle = PROC # { the seventh \acs{acm} symposium on Access + control models and technologies (\acs{sacmat})}, + year = 2002, + isbn = {1-58113-496-7}, + pages = {135--140}, + location = {Monterey, California, USA}, + doi = {10.1145/507711.507732}, + address = pub-acm:adr, + publisher = pub-acm +} + +@InProceedings{ ye.ea:using:2005, + author = {Chunxiao Ye and Zhongfu Wu}, + title = {Using \acs{xml} and \acs{xacml} to Support Attribute Based + Delegation}, + booktitle = {CIT '05: Proceedings of the The Fifth International + Conference on Computer and Information Technology}, + year = 2005, + isbn = {0-7695-2432-X}, + pages = {751--756}, + doi = {10.1109/CIT.2005.196}, + publisher = {IEEE Computer Society}, + address = {Washington, DC, USA} +} + +@Book{ fox.ea:it-sox:2006, + author = {Christopher Fox and Paul Zonneveld}, + abstract = {This publication provides CIOs, IT managers, and control + and assurance professionals with scoping and assessment + ideas, approaches and guidance in support of the IT-related + Committee of Sponsoring Organizations of the Treadway + Commission (COSO) internal control objectives for financial + reporting. + + A streamlined road map is provided to help turn compliance + into a competitive challenge. Lessons learned and + sustaining ideas are also included. + + The second edition is updated for recent US Securities and + Exchange Commission (SEC) and Public Company Accounting and + Oversight Board (PCAOB) guidance related to entity-level + controls, a risk-based/top-down approach, application + controls and the evaluation of deficiencies.}, + title = {\acs{it} Control Objectives for Sarbanes-Oxley: The Role + of \acs{it} in the Design and Implementation of Internal + Control Over Financial Reporting}, + year = 2006, + month = sep, + bibkey = {fox.ea:it-sox:2006}, + num_pages = 128, + edition = {2nd}, + publisher = {IT Governance Institute}, + isbn = {1-933284-76-5}, + address = {Rolling Meadows, IL, USA} +} + +@Article{ basin.ea:automated:2009, + title = {Automated analysis of security-design models}, + journal = j-ist, + volume = 51, + number = 5, + issn = {0950-5849}, + year = 2009, + pages = {815--831}, + doi = {10.1016/j.infsof.2008.05.011}, + author = {David Basin and Manuel Clavel and J{\"u}rgen Doser and + Marina Egea}, + keywords = {Metamodels}, + abstract = {We have previously proposed SecureUML, an expressive + UML-based language for constructing security-design models, + which are models that combine design specifications for + distributed systems with specifications of their security + policies. Here, we show how to automate the analysis of + such models in a semantically precise and meaningful way. + In our approach, models are formalized together with + scenarios that represent possible run-time instances. + Queries about properties of the security policy modeled are + expressed as formulas in UML's Object Constraint Language. + The policy may include both declarative aspects, i.e., + static access-control information such as the assignment of + users and permissions to roles, and programmatic aspects, + which depend on dynamic information, namely the + satisfaction of authorization constraints in a given + scenario. We show how such properties can be evaluated, + completely automatically, in the context of the metamodel + of the security-design language. We demonstrate, through + examples, that this approach can be used to formalize and + check non-trivial security properties. The approach has + been implemented in the SecureMOVA tool and all of the + examples presented have been checked using this tool.}, + note = {Special Issue on Model-Driven Development for Secure + Information Systems}, + publisher = pub-elsevier, + address = pub-elsevier:adr +} + +@InProceedings{ dong.ea:flexible:2008, + author = {Changyu Dong and Giovanni Russello and Naranker Dulay}, + title = {Flexible Resolution of Authorisation Conflicts in + Distributed Systems}, + booktitle = {DSOM}, + year = 2008, + pages = {95--108}, + doi = {10.1007/978-3-540-87353-2_8}, + crossref = {turck.ea:managing:2008}, + abstract = {Managing security in distributed systems requires flexible + and expressive authorisation models with support for + conflict resolution. Models need to be hierarchical but + also non-monotonic supporting both positive and negative + authorisations. In this paper, we present an approach to + resolve the authorisation conflicts that inevitably occur + in such models, with administrator specified conflict + resolution strategies (rules). Strategies can be global or + applied to specific parts of a system and dynamically + loaded for different applications. We use Courteous Logic + Programs (CLP) for the specification and enforcement of + strategies. Authorisation policies are translated into + labelled rules in CLP and prioritised. The prioritisation + is regulated by simple override rules specified or selected + by administrators. We demonstrate the capabilities of the + approach by expressing the conflict resolution strategy for + a moderately complex authorisation model that organises + subjects and objects hierarchically.} +} + +@Proceedings{ turck.ea:managing:2008, + editor = {Filip De Turck and Wolfgang Kellerer and George + Kormentzas}, + title = {Managing Large-Scale Service Deployment, 19th IFIP/IEEE + International Workshop on Distributed Systems: Operations + and Management, DSOM 2008, Samos Island, Greece, September + 22-26, 2008. Proceedings}, + booktitle = {DSOM}, + series = s-lncs, + publisher = pub-springer, + address = pub-springer:adr, + volume = 5273, + year = 2008, + isbn = {978-3-540-85999-4} +} + +@InProceedings{ russello.ea:consent-based:2008, + author = {Giovanni Russello and Changyu Dong and Naranker Dulay}, + title = {Consent-Based Workflows for Healthcare Management}, + booktitle = {9th IEEE International Workshop on Policies for + Distributed Systems and Networks (POLICY 2008), 2-4 June + 2008, Palisades, New York, USA}, + year = 2008, + publisher = {IEEE Computer Society}, + pages = {153--161}, + isbn = {978-0-7695-3133-5}, + doi = {10.1109/POLICY.2008.22}, + abstract = {n this paper, we describe a new framework for healthcare + systems where patients are able to control the disclosure + of their medical data. In our framework, the patient's + consent has a pivotal role in granting or removing access + rights to subjects accessing patient's medical data. + Depending on the context in which the access is being + executed, different consent policies can be applied. + Context is expressed in terms of workflows. The execution + of a task in a given workflow carries the necessary + information to infer whether the consent can be implicitly + retrieved or should be explicitly requested from a patient. + However, patients are always able to enforce their own + decisions and withdraw consent if necessary. Additionally, + the use of workflows enables us to apply the need-to-know + principle. Even when the patient's consent is obtained, a + subject should access medical data only if it is required + by the actual situation. For example, if the subject is + assigned to the execution of a medical diagnosis workflow + requiring access to the patient's medical record. We also + provide a complex medical case study to highlight the + design principles behind our framework. Finally, the + implementation of the framework is outlined.} +} + +@InProceedings{ mitchell-wong.ea:social:2008, + author = {Juliana Mitchell-Wong and Ryszard Kowalczyk and Bao Quoc + Vo}, + title = {Social Network Profile and Policy}, + booktitle = {9th IEEE International Workshop on Policies for + Distributed Systems and Networks (POLICY 2008), 2-4 June + 2008, Palisades, New York, USA}, + year = 2008, + publisher = {IEEE Computer Society}, + isbn = {978-0-7695-3133-5}, + pages = {207--210}, + doi = {10.1109/POLICY.2008.41}, + abstract = {n this paper, we describe a new framework for healthcare + systems where patients are able to control the disclosure + of their medical data. In our framework, the patient's + consent has a pivotal role in granting or removing access + rights to subjects accessing patient's medical data. + Depending on the context in which the access is being + executed, different consent policies can be applied. + Context is expressed in terms of workflows. The execution + of a task in a given workflow carries the necessary + information to infer whether the consent can be implicitly + retrieved or should be explicitly requested from a patient. + However, patients are always able to enforce their own + decisions and withdraw consent if necessary. Additionally, + the use of workflows enables us to apply the need-to-know + principle. Even when the patient's consent is obtained, a + subject should access medical data only if it is required + by the actual situation. For example, if the subject is + assigned to the execution of a medical diagnosis workflow + requiring access to the patient's medical record. We also + provide a complex medical case study to highlight the + design principles behind our framework. Finally, the + implementation of the framework is outlined.} +} + +@Book{ paulson:ml:1996, + author = {Lawrence C. Paulson}, + title = {\acs{ml} for the Working Programmer}, + publisher = {Cambridge Press}, + year = 1996, + acknowledgement={none} +} + +@InProceedings{ kohler.ea:proactive:2008, + title = {Pro Active Access Control for Business Process-driven + Environments}, + author = {Mathias Kohler and Andreas Schaad}, + booktitle = {Annual Computer Security Applications Conference}, + year = 2008 +} + +@InProceedings{ dewin:importance:2002, + author = {Bart De Win and Frank Piessens and Wouter Joosen and Tine + Verhanneman}, + title = {On the importance of the separation-of-concerns principle + in secure software engineering}, + booktitle = {ACSA Workshop on the Application of Engineering Principles + to System Security Design - Final Report (Serban, C., + ed.)}, + year = 2003, + pages = {1--10} +} + +@InProceedings{ priebe:towards:2000, + author = {Torsten Priebe and G\"{u}nther Pernul}, + title = {Towards \textsc{olap} security design --- survey and + research issues}, + booktitle = {DOLAP '00: Proceedings of the 3rd ACM international + workshop on Data warehousing and OLAP}, + year = 2000, + isbn = {1-58113-323-5}, + pages = {33--40}, + location = {McLean, Virginia, United States}, + doi = {10.1145/355068.355313}, + publisher = {ACM}, + address = {New York, NY, USA} +} + +@InProceedings{ atluri:supporting:2005, + author = {Vijayalakshmi Atluri and Janice Warner}, + title = {Supporting conditional delegation in secure workflow + management systems}, + booktitle = PROC # { the tenth \acs{acm} symposium on Access + control models and technologies (\acs{sacmat})}, + year = 2005, + isbn = {1-59593-045-0}, + pages = {49--58}, + location = {Stockholm, Sweden}, + doi = {10.1145/1063979.1063990}, + publisher = pub-acm, + address = pub-acm:adr +} + +@InProceedings{ dimmock:using:2004, + author = {Nathan Dimmock and Andr{\'a} Belokosztolszki and David + Eyers and Jean Bacon and Ken Moody}, + title = {Using trust and risk in role-based access control + policies}, + booktitle = PROC # { the ninth ACM symposium on Access control + models and technologies (\acs{sacmat})}, + year = 2004, + isbn = {1-58113-872-5}, + pages = {156--162}, + location = {Yorktown Heights, New York, USA}, + doi = {10.1145/990036.990062}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {Emerging trust and risk management systems provide a + framework for principals to determine whether they will + exchange resources, without requiring a complete definition + of their credentials and intentions. Most distributed + access control architectures have far more rigid policy + rules, yet in many respects aim to solve a similar problem. + This paper elucidates the similarities between trust + management and distributed access control systems by + demonstrating how the OASIS access control system and its + role-based policy language can be extended to make + decisions on the basis of trust and risk analyses rather + than on the basis of credentials alone. We apply our new + model to the prototypical example of a file storage and + publication service for the Grid, and test it using our + Prolog-based OASIS implementation.} +} + +@Article{ barnett.ea:verification:2004, + author = {Michael Barnett and Robert DeLine and Manuel F{\"a}hndrich + and K. Rustan M. Leino and Wolfram Schulte}, + title = {Verification of Object-Oriented Programs with Invariants}, + journal = {Journal of Object Technology}, + volume = 3, + number = 6, + year = 2004, + pages = {27--56}, + abstract = {An object invariant defines what it means for an object's + data to be in a consistent state. Object invariants are + central to the design and correctness of object-oriented + programs. This paper defines a programming methodology for + using object invariants. The methodology, which enriches a + program's state space to express when each object invariant + holds, deals with owned object components, ownership + transfer, and subclassing, and is expressive enough to + allow many interesting object-oriented programs to be + specified and verified. Lending itself to sound modular + verification, the methodology also provides a solution to + the problem of determining what state a method is allowed + to modify. }, + url = {http://www.jot.fm/issues/issue_2004_06/article2/article2.pdf} + +} + +@Article{ harms.ea:copying:1991, + author = {Douglas E. Harms and Bruce W. Weide}, + title = {Copying and Swapping: Influences on the Design of Reusable + Software Components}, + journal = j-tse, + volume = 17, + number = 5, + year = 1991, + pages = {424--435}, + doi = {10.1109/32.90445 }, + abstract = {The authors argue that a simple alternative to copying as + a data movement primitive-swapping (exchanging) the values + of two variables-has potentially significant advantages in + the context of the design of generic reusable software + components. Specifically, the authors claim that generic + module designs based on a swapping style are superior to + designs based on copying, both in terms of execution-time + efficiency and with respect to the likelihood of + correctness of client programs and module implementations. + Furthermore, designs based on swapping are more reusable + than traditional designs. Specific arguments and examples + to support these positions are presented}, + publisher = pub-ieee, + address = pub-ieee:adr +} + +@InProceedings{ albano.ea:relationship:1991, + author = {Antonio Albano and Giorgio Ghelli and Renzo Orsini}, + title = {A Relationship Mechanism for a Strongly Typed + Object-Oriented Database Programming Language}, + booktitle = {VLDB}, + year = 1991, + pages = {565--575}, + crossref = {lohman.ea:17th:1991} +} + +@Proceedings{ lohman.ea:17th:1991, + editor = {Guy M. Lohman and Am\'{\i}lcar Sernadas and Rafael Camps}, + title = {17th International Conference on Very Large Data Bases, + September 3-6, 1991, Barcelona, Catalonia, Spain, + Proceedings}, + publisher = {Morgan Kaufmann}, + year = 1991, + isbn = {1-55860-150-3} +} + +@InProceedings{ ernst.ea:predicate:1998, + author = {Michael D. Ernst and Craig S. Kaplan and Craig Chambers}, + title = {Predicate Dispatching: A Unified Theory of Dispatch}, + booktitle = {ECOOP}, + year = 1998, + pages = {186--211}, + doi = {10.1007/BFb0054092}, + abstract = {Predicate dispatching generalizes previous method dispatch + mechanisms by permitting arbitrary predicates to control + method applicability and by using logical implication + between predicates as the overriding relationship. The + method selected to handle a message send can depend not + just on the classes of the arguments, as in ordinary + object-oriented dispatch, but also on the classes of + subcomponents, on an argument's state, and on relationships + between objects. This simple mechanism subsumes and extends + object-oriented single and multiple dispatch, ML-style + pattern matching, predicate classes, and classifiers, which + can all be regarded as syntactic sugar for predicate + dispatching. This paper introduces predicate dispatching, + gives motivating examples, and presents its static and + dynamic semantics. An implementation of predicate + dispatching is available.}, + crossref = {jul:eccop98:1998} +} + +@Proceedings{ jul:eccop98:1998, + editor = {Eric Jul}, + title = {ECCOP'98 - Object-Oriented Programming, 12th European + Conference, Brussels, Belgium, July 20-24, 1998, + Proceedings}, + booktitle = {ECOOP}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 1445, + year = 1998, + isbn = {3-540-64737-6} +} + +@Misc{ garbani:future:2009, + author = {Jean-Pierre Garbani}, + title = {Future Trends In The Enterprise Software Market}, + howpublished = {http://www.forrester.com/Research/Document/Excerpt/0,7211,53493,00.html} + , + month = mar # {~9}, + year = 2009, + publisher = {Forrester Research, Inc.}, + address = {Cambridge, USA}, + abstract = {Hardware, software, and people are the three basic + ingredients of enterprise business technology. They provide + the enterprise with an economic advantage through automated + and improved business processes, increased employee + productivity, and more accurate and precise information. + The relationship between these three components has evolved + over time: Initially, hardware reigned supreme; software + was a mere adjunct and free to the mainframe buyer. + Decreasing hardware costs then led to software + "unbundling." In the present era of information technology, + hardware's ever-decreasing costs make it an enabler of + software functions. In the business technology (BT) era, we + predict that managing the third part of the equation -- + people -- will emerge as the dominant focus. As software + applications become business services, the cost of human + resources producing, operating, and managing software will + soon be prohibitive and the new focal point. In this + regard, the current economic downturn, if it persists, may + prove to be a driver that accelerates the shift toward the + BT era.} +} + +@Article{ klein:operating:2009, + author = {Gerwin Klein}, + title = {Operating System Verification --- An Overview}, + journal = {S\={a}dhan\={a}}, + publisher = pub-springer, + year = 2009, + volume = 34, + number = 1, + month = feb, + pages = {27--69}, + abstract = {This paper gives a high-level introduction to the topic of + formal, interactive, machine-checked software verification + in general, and the verification of operating systems code + in particular. We survey the state of the art, the + advantages and limitations of machine-checked code proofs, + and describe two specific ongoing larger-scale verification + projects in more detail.} +} + +@Article{ edwards.ea:resolve:1994, + bibkey = {edwards.ea:resolve:1994}, + author = {Stephen H. Edwards and Wayne D. Heym and Timothy J. Long + and Murali Sitaramanand Bruce W. Weide}, + title = {Part II: specifying components in {RESOLVE}}, + journal = {SIGSOFT Softw. Eng. Notes}, + volume = 19, + number = 4, + year = 1994, + issn = {0163-5948}, + pages = {29--39}, + doi = {10.1145/190679.190682}, + publisher = pub-acm, + address = pub-acm:adr +} + +@PhDThesis{ kassios:theory:2006, + author = {Ioannis T. Kassios}, + title = {A Theory of Object Oriented Refinement}, + school = {University of Toronto}, + abstract = {This thesis introduces a formal theory of object oriented + refinement. The formal design of the theory is based on the + design principles of unification and decoupling, which we + believe have not been adequately used in other object + oriented refinement theories. + + Thanks to the use of these principles, the semantics of the + theory is mathematically simpler than other approaches: the + formalization of most features uses very primitive + mathematical entities. Furthermore, the constructs of the + theory are more general than other approaches. There are + specification constructs more general than classes. + Features that are typically coupled with classes, like + reuse and polymorphism, now apply to these more general + specifications. Finally, our solution to the frame problem + is the only modular approach that we know of that does not + impose any aliasing control policy. + + To demonstrate that the extra generality offers real + advantages to the specifier, we use the theory in some + specification examples that would be impossible with other + approaches. These examples, mainly inspired by Design + Patterns, represent realistic and common software design + situations.}, + year = 2006 +} + +@TechReport{ dewar:setl:1979, + author = {Robert B. K. Dewar}, + title = {The {SETL} Programming Language}, + year = 1979 +} + +@InBook{ chun.ea:risk-based:2008, + author = {Soon Ae Chun and Vijay Atluri}, + editor = {Bhargab B. Bhattacharya and Susmita Sur-Kolay and Subhas + C. Nandy and Aditya Bagchi}, + booktitle = {Algorithms, Architecture and Information Systems + Security}, + title = {Risk-based Access Control for Personal Data Services}, + publisher = {World Scientific Press}, + year = 2008, + volume = 3, + series = {Statistical Science and Interdisciplinary Research}, + isbn = 9789812836236 +} + +@Unpublished{ clark.ea:survey:1997, + author = {John Clark and Jeremy Jacob}, + title = {A Survey of Authentication Protocol: Literature: Version + 1.0}, + year = 1997 +} + +@Unpublished{ dierks.ea:tls:1999, + author = {T. Dierks and C. Allen}, + title = {The TLS Protocol Version 1.0}, + year = 1999, + publisher = {RFC Editor}, + address = {United States}, + note = {RFC 2246} +} + +@InProceedings{ fontaine.ea:expressiveness:2006, + author = {Pascal Fontaine and Jean-Yves Marion and Stephan Merz and + Leonor Prensa Nieto and Alwen Fernanto Tiu}, + title = {Expressiveness + Automation + Soundness: Towards Combining + SMT Solvers and Interactive Proof Assistants}, + booktitle = {TACAS}, + year = 2006, + pages = {167--181}, + doi = {10.1007/11691372_11}, + crossref = {hermanns.ea:tools:2006} +} + +@Proceedings{ hermanns.ea:tools:2006, + editor = {Holger Hermanns and Jens Palsberg}, + title = {Tools and Algorithms for the Construction and Analysis of + Systems, 12th International Conference, TACAS 2006 Held as + Part of the Joint European Conferences on Theory and + Practice of Software, ETAPS 2006, Vienna, Austria, March 25 + - April 2, 2006, Proceedings}, + booktitle = {TACAS}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 3920, + year = 2006, + isbn = {3-540-33056-9} +} + +@InProceedings{ amjad:lcf-style:2008, + author = {Hasan Amjad}, + title = {LCF-Style Propositional Simplification with BDDs and SAT + Solvers}, + booktitle = {TPHOLs}, + year = 2008, + pages = {55--70}, + doi = {10.1007/978-3-540-71067-7_9}, + crossref = {mohamed.ea:theorem:2008} +} + +@Proceedings{ mohamed.ea:theorem:2008, + editor = {Otmane A\"{\i}t Mohamed and C{\'e}sar Mu{\~n}oz and + Sofi{\`e}ne Tahar}, + title = {Theorem Proving in Higher Order Logics, 21st International + Conference, TPHOLs 2008, Montreal, Canada, August 18-21, + 2008. Proceedings}, + booktitle = {TPHOLs}, + publisher = pub-springer, + address = pub-springer:adr, + volume = 5170, + series = s-lncs, + year = 2008, + isbn = {978-3-540-71065-3} +} + +@Article{ weber:integrating:2006, + author = {Tjark Weber}, + title = {Integrating a {SAT} Solver with an {LCF}-style Theorem + Prover}, + editor = {Alessandro Armando and Alessandro Cimatti}, + journal = j-entcs, + month = jan, + year = 2006, + publisher = pub-elsevier, + address = pub-elsevier:adr, + pages = {67--78}, + doi = {10.1016/j.entcs.2005.12.007}, + issn = {1571-0661}, + volume = 144, + number = 2, + note = PROC # { the Third Workshop on Pragmatics of + Decision Procedures in Automated Reasoning (PDPAR 2005)}, + clearance = {unclassified}, + abstract = {This paper describes the integration of a leading SAT + solver with Isabelle/HOL, a popular interactive theorem + prover. The SAT solver generates resolution-style proofs + for (instances of) propositional tautologies. These proofs + are verified by the theorem prover. The presented approach + significantly improves Isabelle's performance on + propositional problems, and furthermore exhibits + counterexamples for unprovable conjectures.} +} + +@Article{ weber.ea:efficiently:2009, + title = {Efficiently checking propositional refutations in HOL + theorem provers}, + journal = {Journal of Applied Logic}, + volume = 7, + number = 1, + pages = {26 -- 40}, + year = 2009, + note = {Special Issue: Empirically Successful Computerized + Reasoning}, + issn = {1570-8683}, + doi = {10.1016/j.jal.2007.07.003}, + author = {Tjark Weber and Hasan Amjad}, + abstract = {This paper describes the integration of zChaff and + MiniSat, currently two leading SAT solvers, with Higher + Order Logic (HOL) theorem provers. Both SAT solvers + generate resolution-style proofs for (instances of) + propositional tautologies. These proofs are verified by the + theorem provers. The presented approach significantly + improves the provers' performance on propositional + problems, and exhibits counterexamples for unprovable + conjectures. It is also shown that LCF-style theorem + provers can serve as viable proof checkers even for large + SAT problems. An efficient representation of the + propositional problem in the theorem prover turns out to be + crucial; several possible solutions are discussed.} +} + +@Article{ wendling:german:2009, + title = {The German {eHealth} programme}, + journal = {Card Technology Today}, + volume = 21, + number = 1, + pages = {10--11}, + year = 2009, + issn = {0965-2590}, + doi = {10.1016/S0965-2590(09)70018-0}, + author = {Dietmar Wendling}, + abstract = {Germany was one of the first countries in the world to use + smart cards for healthcare. Now it is at the starting gate + to roll out a new generation of cards. Dietmar Wendling, + vice president of the eGovernment market sector at SCM + Microsystems reports.} +} + +@Article{ meng.ea:translating:2008, + author = {Jia Meng and Lawrence C. Paulson}, + title = {Translating Higher-Order Clauses to First-Order Clauses}, + journal = j-jar, + volume = 40, + number = 1, + year = 2008, + pages = {35--60}, + doi = {10.1007/s10817-007-9085-y} +} + +@InProceedings{ paulson.ea:source-level:2007, + author = {Lawrence C. Paulson and Kong Woei Susanto}, + title = {Source-Level Proof Reconstruction for Interactive Theorem + Proving}, + booktitle = {TPHOLs}, + year = 2007, + pages = {232--245}, + doi = {10.1007/978-3-540-74591-4_18}, + crossref = {schneider.ea:theorem:2007} +} + +@Proceedings{ schneider.ea:theorem:2007, + editor = {Klaus Schneider and Jens Brandt}, + title = {Theorem Proving in Higher Order Logics, 20th International + Conference, TPHOLs 2007, Kaiserslautern, Germany, September + 10-13, 2007, Proceedings}, + booktitle = {TPHOLs}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 4732, + year = 2007, + isbn = {978-3-540-74590-7} +} + +@Article{ meng.ea:automation:2006, + author = {Jia Meng and Claire Quigley and Lawrence C. Paulson}, + title = {Automation for interactive proof: First prototype}, + journal = {Inf. Comput.}, + volume = 204, + number = 10, + year = 2006, + pages = {1575--1596}, + doi = {10.1016/j.ic.2005.05.010} +} + +@InProceedings{ erkok.ea:using:2008, + location = {Princeton, New Jersey, USA}, + author = {Levent Erk{\"o}k and John Matthews}, + booktitle = {Automated Formal Methods (AFM'08)}, + title = {Using Yices as an Automated Solver in Isabelle/{HOL}}, + year = 2008 +} + +@Article{ jurjens.ea:model-based:2008, + author = {Jan J{\"u}rjens and Rumm, R.}, + title = {Model-based security analysis of the German health card + architecture.}, + journal = {Methods Inf Med}, + year = 2008, + volume = 47, + number = 5, + pages = {409--416}, + keywords = {Patient Identification Systems}, + abstract = {OBJECTIVES: Health-care information systems are + particularly security-critical. In order to make these + applications secure, the security analysis has to be an + integral part of the system design and IT management + process for such systems. METHODS: This work presents the + experiences and results from the security analysis of the + system architecture of the German Health Card, by making + use of an approach to model-based security engineering that + is based on the UML extension UMLsec. The focus lies on the + security mechanisms and security policies of the + smart-card-based architecture which were analyzed using the + UMLsec method and tools. RESULTS: Main results of the paper + include a report on the employment of the UMLsec method in + an industrial health information systems context as well as + indications of its benefits and limitations. In particular, + two potential security weaknesses were detected and + countermeasures discussed. CONCLUSIONS: The results + indicate that it can be feasible to apply a model-based + security analysis using UMLsec to an industrial health + information system like the German Health Card + architecture, and that doing so can have concrete benefits + (such as discovering potential weaknesses, and an increased + confidence that no further vulnerabilities of the kind that + were considered are present).}, + issn = {0026-1270} +} + +@InProceedings{ miseldine:automated:2008, + author = {Philip Miseldine}, + title = {Automated {XACML} policy reconfiguration for evaluation + optimisation}, + booktitle = {SESS}, + year = 2008, + pages = {1--8}, + doi = {10.1145/1370905.1370906}, + crossref = {win.ea:proceedings:2008}, + abstract = {We present a programmatic approach to the optimisation of + XACML policies that specifies how a set of access control + rules should be best represented for optimised evaluation. + The work assumes no changes to the current XACML + specification and methods of interpretation shall be made, + so that those who consume XACML are unaffected + structurally, and those that generate XACML can provide + optimised output. Discussion regarding the flexibility of + the XACML specification to describe the same access rules + with different policy configurations is presented, and is + used to formulate a comprehensive analysis of the + evaluation costs the possible policy configurations will + produce. This leads to the specification of methods that + can be employed to produce optimal forms of policy + description. These are implemented and evaluated to show + the benefits of the approach proposed.} +} + +@Proceedings{ win.ea:proceedings:2008, + editor = {Bart De Win and Seok-Won Lee and Mattia Monga}, + title = PROC # { the Fourth International Workshop on + Software Engineering for Secure Systems, SESS 2008, + Leipzig, Germany, May 17-18, 2008}, + booktitle = {SESS}, + publisher = {ACM}, + year = 2008, + isbn = {978-1-60558-042-5} +} + +@InProceedings{ liu.ea:firewall:2008, + author = {Alex X. Liu and Eric Torng and Chad Meiners}, + title = {Firewall Compressor: An Algorithm for Minimizing Firewall + Policies}, + booktitle = PROC # { the 27th Annual IEEE Conference on Computer + Communications (Infocom)}, + year = 2008, + address = {Phoenix, Arizona}, + month = {April} +} + +@InProceedings{ liu.ea:xengine:2008, + author = {Alex X. Liu and Fei Chen and JeeHyun Hwang and Tao Xie}, + title = {{XEngine}: A Fast and Scalable {XACML} Policy Evaluation + Engine}, + booktitle = PROC # { the International Conference on Measurement + and Modeling of Computer Systems (Sigmetrics)}, + year = 2008, + address = {Annapolis, Maryland}, + month = {June} +} + +@InProceedings{ goubault-larrecq:towards:2008, + address = {Pittsburgh, PA, USA}, + author = {Goubault{-}Larrecq, Jean}, + booktitle = {{P}roceedings of the 21st {IEEE} {C}omputer {S}ecurity + {F}oundations {S}ymposium ({CSF}'08)}, + doi = {10.1109/CSF.2008.21}, + month = jun, + pages = {224--238}, + publisher = {{IEEE} Computer Society Press}, + title = {Towards Producing Formally Checkable Security Proofs, + Automatically}, + year = 2008 +} + +@InProceedings{ weidenbach.ea:spass:2007, + author = {Christoph Weidenbach and Renate A. Schmidt and Thomas + Hillenbrand and Rostislav Rusev and Dalibor Topic}, + title = {System Description: Spass Version 3.0}, + booktitle = {CADE}, + year = 2007, + pages = {514--520}, + doi = {10.1007/978-3-540-73595-3_38}, + crossref = {pfenning:automated:2007} +} + +@Proceedings{ pfenning:automated:2007, + editor = {Frank Pfenning}, + title = {Automated Deduction - CADE-21, 21st International + Conference on Automated Deduction, Bremen, Germany, July + 17-20, 2007, Proceedings}, + booktitle = {CADE}, + series = s-lncs, + publisher = pub-springer, + address = pub-springer:adr, + volume = 4603, + year = 2007, + isbn = {978-3-540-73594-6} +} + +@Article{ paulson:tls:1999, + author = {Lawrence C. Paulson}, + title = {Inductive Analysis of the Internet Protocol {TLS}}, + journal = {ACM Trans. Inf. Syst. Secur.}, + volume = 2, + number = 3, + year = 1999, + pages = {332--351}, + doi = {10.1145/322510.322530} +} + +@Article{ harman.ea:testability:2004, + author = {Mark Harman and Lin Hu and Rob Hierons and Joachim Wegener + and Harmen Sthamer and Andr{\'e} Baresel and Marc Roper}, + title = {Testability Transformation}, + journal = {IEEE Trans. Softw. Eng.}, + volume = 30, + number = 1, + year = 2004, + issn = {0098-5589}, + pages = {3--16}, + doi = {10.1109/TSE.2004.1265732}, + publisher = {IEEE Press}, + address = {Piscataway, NJ, USA} +} + +@Article{ dssouli.ea:communications:1999, + title = {Communications software design for testability: + specification transformations and testability measures}, + journal = {Information and Software Technology}, + volume = 41, + number = {11-12}, + pages = {729--743}, + year = 1999, + issn = {0950-5849}, + doi = {10.1016/S0950-5849(99)00033-6}, + url = {http://www.sciencedirect.com/science/article/B6V0B-3X3TD3J-4/2/0efca94003ffe88571f8aa2d346a1289} + , + author = {R. Dssouli and K. Karoui and K. Saleh and O. Cherkaoui}, + keywords = {Testing} +} + +@Article{ baker:equal:1993, + author = {Henry G. Baker}, + title = {Equal rights for functional objects or, the more things + change, the more they are the same}, + journal = {OOPS Messenger}, + volume = 4, + number = 4, + year = 1993, + pages = {2--27}, + abstract = {We argue that intensional object identity in + object-oriented programming languages and databases is best + defined operationally by side-effect semantics. A corollary + is that "functional" objects have extensional semantics. + This model of object identity, which is analogous to the + normal forms of relational algebra, provides cleaner + semantics for the value-transmission operations and + built-in primitive equality predicate of a programming + language, and eliminates the confusion surrounding + "call-by-value" and "call-by-reference" as well as the + confusion of multiple equality predicates.Implementation + issues are discussed, and this model is shown to have + significant performance advantages in persistent, parallel, + distributed and multilingual processing environments. This + model also provides insight into the "type equivalence" + problem of Algol-68, Pascal and Ada.} +} + +@Article{ hierons.ea:branch-coverage:2005, + author = {Robert M. Hierons and Mark Harman and Chris Fox}, + title = {Branch-Coverage Testability Transformation for + Unstructured Programs}, + journal = {Comput. J.}, + volume = 48, + number = 4, + year = 2005, + pages = {421--436}, + doi = {10.1093/comjnl/bxh093}, + abstract = {Test data generation by hand is a tedious, expensive and + error-prone activity, yet testing is a vital part of the + development process. Several techniques have been proposed + to automate the generation of test data, but all of these + are hindered by the presence of unstructured control flow. + This paper addresses the problem using testability + transformation. Testability transformation does not + preserve the traditional meaning of the program, rather it + deals with preserving test-adequate sets of input data. + This requires new equivalence relations which, in turn, + entail novel proof obligations. The paper illustrates this + using the branch coverage adequacy criterion and develops a + branch adequacy equivalence relation and a testability + transformation for restructuring. It then presents a proof + that the transformation preserves branch adequacy.} +} + +@InProceedings{ harman:open:2008, + title = {Open Problems in Testability Transformation}, + author = {Harman, M.}, + booktitle = {Software Testing Verification and Validation Workshop, + 2008. ICSTW '08. IEEE International Conference on}, + year = 2008, + month = {April}, + pages = {196--209}, + keywords = {data analysis, program testingsearch-based test data + generation, test adequacy criterion, testability + transformation}, + doi = {10.1109/ICSTW.2008.30}, + issn = {978-0-7695-3388-9}, + abstract = {Testability transformation (tetra) seeks to transform a + program in order to make it easier to generate test data. + The test data is generated from the transformed version of + the program, but it is applied to the original version for + testing purposes. A transformation is a testability + transformation with respect to a test adequacy criterion if + all test data that is adequate for the transformed program + is also adequate for the untransformed program. Testability + transformation has been shown to be effective at improving + coverage for search based test data generation. However, + there are many interesting open problems. This paper + presents some of these open problems. The aim is to show + how testability transformation can be applied to a wide + range of testing scenarios.} +} + +@InProceedings{ harman.ea:testability:2008, + author = {Mark Harman and Andr{\'e} Baresel and David Binkley and + Robert M. Hierons and Lin Hu and Bogdan Korel and Phil + McMinn and Marc Roper}, + title = {Testability Transformation - Program Transformation to + Improve Testability}, + booktitle = {Formal Methods and Testing}, + year = 2008, + pages = {320--344}, + doi = {10.1007/978-3-540-78917-8_11}, + crossref = {hierons.ea:formal:2008} +} + +@InProceedings{ veanes.ea:model-based:2008, + author = {Margus Veanes and Colin Campbell and Wolfgang Grieskamp + and Wolfram Schulte and Nikolai Tillmann and Lev + Nachmanson}, + title = {Model-Based Testing of Object-Oriented Reactive Systems + with Spec Explorer}, + booktitle = {Formal Methods and Testing}, + year = 2008, + pages = {39--76}, + doi = {10.1007/978-3-540-78917-8_2}, + abstract = {}, + crossref = {hierons.ea:formal:2008} +} + +@Proceedings{ hierons.ea:formal:2008, + editor = {Robert M. Hierons and Jonathan P. Bowen and Mark Harman}, + title = {Formal Methods and Testing, An Outcome of the FORTEST + Network, Revised Selected Papers}, + booktitle = {Formal Methods and Testing}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 4949, + year = 2008, + isbn = {978-3-540-78916-1} +} + +@Booklet{ omg:uml-infrastructure:2009, + bibkey = {omg:uml-infrastructure:2009}, + key = omg, + publisher = omg, + language = {USenglish}, + note = {Available as \acs{omg} document + \href{http://www.omg.org/cgi-bin/doc?formal/2009-02-04} + {formal/2009-02-04}}, + keywords = {\acs{uml}}, + topic = {formalism}, + public = {yes}, + title = {\acs{uml} 2.2 Infrastructure Specification}, + year = 2009 +} + +@InProceedings{ pirretti.ea:secure:2006, + author = {Matthew Pirretti and Patrick Traynor and Patrick McDaniel + and Brent Waters}, + title = {Secure attribute-based systems}, + booktitle = PROC # {ACM conference on + Computer and communications security (CCS)}, + year = 2006, + isbn = {1-59593-518-5}, + pages = {99--112}, + location = {Alexandria, Virginia, USA}, + doi = {10.1145/1180405.1180419}, + address = pub-acm:adr, + publisher = pub-acm, + abstract = {Attributes define, classify, or annotate the datum to + which they are assigned. However, traditional attribute + architectures and cryptosystems are ill-equipped to provide + security in the face of diverse access requirements and + environments. In this paper, we introduce a novel secure + information management architecture based on emerging + attribute-based encryption (ABE) primitives. A policy + system that meets the needs of complex policies is defined + and illustrated. Based on the needs of those policies, we + propose cryptographic optimizations that vastly improve + enforcement efficiency. We further explore the use of such + policies in two example applications: a HIPAA compliant + distributed file system and a social network. A performance + analysis of our ABE system and example applications + demonstrates the ability to reduce cryptographic costs by + as much as 98\% over previously proposed constructions. + Through this, we demonstrate that our attribute system is + an efficient solution for securely managing information in + large, loosely-coupled, distributed systems.} +} + +@Article{ milicev:semantics:2007, + bibkey = {milicev:semantics:2007}, + title = {On the Semantics of Associations and Association Ends in + UML}, + author = {Dragan Milicev}, + journal = {IEEE Transactions on Software Engineering}, + year = 2007, + month = apr, + volume = 33, + number = 4, + pages = {238--251}, + keywords = {Unified Modeling Language, entity-relationship modelling, + formal specification, object-oriented programming, + programming language semanticsUML, Unified Modeling + Language, association end, conceptual modeling, formal + semantics, formal specification, intentional + interpretation, object-oriented modeling}, + doi = {10.1109/TSE.2007.37}, + issn = {0098-5589} +} + +@InProceedings{ bierman.ea:first-class:2005, + author = {Gavin M. Bierman and Alisdair Wren}, + title = {First-Class Relationships in an Object-Oriented Language}, + booktitle = {ECOOP}, + year = 2005, + pages = {262--286}, + doi = {10.1007/11531142_12}, + crossref = {black:ecoop:2005} +} + +@Proceedings{ black:ecoop:2005, + editor = {Andrew P. Black}, + title = {ECOOP 2005 - Object-Oriented Programming, 19th European + Conference, Glasgow, UK, July 25-29, 2005, Proceedings}, + booktitle = {ECOOP}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 3586, + year = 2005, + isbn = {3-540-27992-X} +} + +@Article{ shafiq.ea:secure:2005, + title = {Secure interoperation in a multidomain environment + employing RBAC policies}, + author = {Basit Shafiq and James B.D. Joshi and Elisa Bertino and + Arif Ghafoor}, + journal = j-tkde, + year = 2005, + month = nov, + volume = 17, + number = 11, + pages = {1557--1577}, + keywords = {Internet, authorisation, integer programming, open systems + Internet-based enterprise, heterogeneous role-based access + control, integer programming, multidomain application + environment, optimality criterion, policy integration + framework, secure interoperation}, + doi = {10.1109/TKDE.2005.185}, + issn = {1041-4347} +} + +@Article{ aedo.ea:rbac:2006, + volume = 11, + number = 4, + month = dec, + year = 2006, + title = {An {RBAC} Model-Based Approach to Specify the Access + Policies of Web-Based Emergency Information Systems}, + author = {Ignacio Aedo and Paloma D{\'i}az and Daniel Sanz}, + journal = {The International Journal of Intelligent Control and + Systems}, + page = {272--283}, + abstract = {One of the main design challenges of any Web-based + Emergency Management Information System (WEMIS) is the + diversity of users and responsibilities to be considered. + Modelling the access capabilities of different communities + of users is a most relevant concern for which the RBAC + (Role-Based Access Control) paradigm provides flexible and + powerful constructs. In this paper we describe how we used + an RBAC model-based approach to specify at different levels + of abstraction the access policy of a specific WEMIS called + ARCE (Aplicaci{\`o}n en Red para Casos de Emergencia). This + approach made possible to face access modelling at earlier + development stages, so that stakeholders got involved in + analytical and empirical evaluations to test the + correctness and completeness of the access policy. + Moreover, since the RBAC meta-model is embedded into a web + engineering method, we put in practice a holistic process + addressing different design perspectives in an integrated + way. } +} + +@InProceedings{ phillips.ea:information:2002, + author = {Charles E. Phillips, Jr. and T.C. Ting and Steven A. + Demurjian}, + title = {Information sharing and security in dynamic coalitions}, + booktitle = {SACMAT '02: Proceedings of the seventh ACM symposium on + Access control models and technologies}, + year = 2002, + isbn = {1-58113-496-7}, + pages = {87--96}, + location = {Monterey, California, USA}, + doi = {10.1145/507711.507726}, + address = pub-acm:adr, + publisher = pub-acm +} + +@InProceedings{ martino.ea:multi-domain:2008, + title = {Multi-domain and privacy-aware role based access control + in eHealth}, + author = {Lorenzo D. Martino and Qun Ni and Dan Lin and Elisa + Bertino}, + booktitle = {Second International Conference on Pervasive Computing + Technologies for Healthcare (PervasiveHealth 2008)}, + year = 2008, + month = {30 2008-Feb. 1}, + pages = {131--134}, + keywords = {authorisation, data privacy, health care, medical + information systemseHealth, electronic medical/health + records, healthcare professionals, multidomain + privacy-aware role based access control, patient safety, + privacy preserving,}, + doi = {10.1109/PCTHEALTH.2008.4571050} +} + +@InProceedings{ kamath.ea:user-credential:2006, + author = {Ajith Kamath and Ramiro Liscano and Abdulmotaleb El + Saddik}, + title = {User-credential based role mapping in multi-domain + environment}, + booktitle = {PST '06: Proceedings of the 2006 International Conference + on Privacy, Security and Trust}, + year = 2006, + isbn = {1-59593-604-1}, + pages = {1--1}, + location = {Markham, Ontario, Canada}, + doi = {10.1145/1501434.1501507}, + address = pub-acm:adr, + publisher = pub-acm +} + +@Article{ geethakumari.ea:cross:2009, + journal = {International Journal of Computer Science \& Applications}, + title = {A Cross -- Domain Role Mapping and Authorization Framework + for {RBAC} in Grid Systems.}, + year = 2009, + volume = {VI}, + issue = {I}, + author = {G. Geethakumari and Atul Negi and V. N. Sastry}, + issn = {0972-9038} +} + +@InProceedings{ gogolla.ea:benchmark:2008, + author = {Martin Gogolla and Mirco Kuhlmann and Fabian B{\"u}ttner}, + title = {A Benchmark for \acs{ocl} Engine Accuracy, + Determinateness, and Efficiency}, + booktitle = {MoDELS}, + year = 2008, + pages = {446--459}, + doi = {10.1007/978-3-540-87875-9_32}, + crossref = {czarnecki.ea:models:2008}, + abstract = {The Object Constraint Language (OCL) is a central element + in modeling and transformation languages like UML, MOF, and + QVT. Consequently approaches for MDE (Model-Driven + Engineering) depend on OCL. However, OCL is present not + only in these areas influenced by the OMG but also in the + Eclipse Modeling Framework (EMF). Thus the quality of OCL + and its realization in tools seems to be crucial for the + success of model-driven development. Surprisingly, up to + now a benchmark for OCL to measure quality properties has + not been proposed. This paper puts forward in the first + part the concepts of a comprehensive OCL benchmark. Our + benchmark covers (A) OCL engine accuracy (e.g., for the + undefined value and the use of variables), (B) OCL engine + determinateness properties (e.g., for the collection + operations any and flatten), and (C) OCL engine efficiency + (for data type and user-defined operations). In the second + part, this paper empirically evaluates the proposed + benchmark concepts by examining a number of OCL tools. The + paper discusses several differences in handling particular + OCL language features and underspecifications in the OCL + standard.} +} + +@InProceedings{ gessenharter:mapping:2008, + author = {Dominik Gessenharter}, + title = {Mapping the {UML2} Semantics of Associations to a {Java} + Code Generation Model}, + booktitle = {MoDELS}, + year = 2008, + pages = {813--827}, + doi = {10.1007/978-3-540-87875-9_56}, + crossref = {czarnecki.ea:models:2008} +} + +@Proceedings{ czarnecki.ea:models:2008, + editor = {Krzysztof Czarnecki and Ileana Ober and Jean-Michel Bruel + and Axel Uhl and Markus V{\"o}lter}, + title = {Model Driven Engineering Languages and Systems, 11th + International Conference, MoDELS 2008, Toulouse, France, + September 28 - October 3, 2008. Proceedings}, + booktitle = {MoDELS}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 5301, + year = 2008, + isbn = {978-3-540-87874-2} +} + +@Article{ aalst.ea:workflow:2003, + author = {Wil M. P. van der Aalst and Arthur H. M. ter Hofstede and + Bartek Kiepuszewski and Alistair P. Barros}, + title = {Workflow Patterns}, + journal = {Distributed and Parallel Databases}, + volume = 14, + number = 1, + year = 2003, + pages = {5--51}, + doi = {10.1023/A:1022883727209}, + abstract = {Differences in features supported by the various + contemporary commercial workflow management systems point + to different insights of suitability and different levels + of expressive power. The challenge, which we undertake in + this paper, is to systematically address workflow + requirements, from basic to complex. Many of the more + complex requirements identified, recur quite frequently in + the analysis phases of workflow projects, however their + implementation is uncertain in current products. + Requirements for workflow languages are indicated through + workflow patterns. In this context, patterns address + business requirements in an imperative workflow style + expression, but are removed from specific workflow + languages. The paper describes a number of workflow + patterns addressing what we believe identify comprehensive + workflow functionality. These patterns provide the basis + for an in-depth comparison of a number of commercially + availablework flow management systems. As such, this paper + can be seen as the academic response to evaluations made by + prestigious consulting companies. Typically, these + evaluations hardly consider the workflow modeling language + and routing capabilities, and focus more on the purely + technical and commercial aspects.} +} + +@InProceedings{ chalin.ea:non-null:2005, + author = {Patrice Chalin and Fr\'{e}d\'{e}ric Rioux}, + title = {Non-null references by default in the {Java} modeling + language}, + booktitle = {SAVCBS '05: Proceedings of the 2005 conference on + Specification and verification of component-based systems}, + year = 2005, + isbn = {1-59593-371-9}, + pages = 9, + location = {Lisbon, Portugal}, + doi = {10.1145/1123058.1123068}, + address = pub-acm:adr, + publisher = pub-acm, + abstract = {Based on our experiences and those of our peers, we + hypothesized that in Java code, the majority of + declarations that are of reference types are meant to be + non-null. Unfortunately, the Java Modeling Language (JML), + like most interface specification and object-oriented + programming languages, assumes that such declarations are + possibly-null by default. As a consequence, developers need + to write specifications that are more verbose than + necessary in order to accurately document their module + interfaces. In practice, this results in module interfaces + being left incompletely and inaccurately specified. In this + paper we present the results of a study that confirms our + hypothesis. Hence, we propose an adaptation to JML that + preserves its language design goals and that allows + developers to specify that declarations of reference types + are to be interpreted as non-null by default. We explain + how this default is safer and results in less writing on + the part of specifiers than null-by-default. The paper also + reports on an implementation of the proposal in some of the + JML tools.} +} + +@Article{ ekman.ea:pluggable:2007, + author = {Torbj{\"o}rn Ekman and G{\"o}rel Hedin}, + title = {Pluggable checking and inferencing of nonnull types for + {Java}}, + journal = {Journal of Object Technology}, + volume = 6, + number = 9, + year = 2007, + pages = {455--475}, + ee = {http://www.jot.fm/issues/issue_2007_10/paper23/index.html} + +} + +@InProceedings{ fahndrich.ea:declaring:2003, + author = {Manuel F{\"a}hndrich and K. Rustan M. Leino}, + title = {Declaring and checking non-null types in an + object-oriented language}, + booktitle = {OOPSLA}, + year = 2003, + pages = {302--312}, + doi = {10.1145/949305.949332}, + crossref = {crocker.ea:proceedings:2003} +} + +@Proceedings{ crocker.ea:proceedings:2003, + editor = {Ron Crocker and Guy L. Steele Jr.}, + title = PROC # { the 2003 ACM SIGPLAN Conference on + Object-Oriented Programming Systems, Languages and + Applications, OOPSLA 2003, October 26-30, 2003, Anaheim, + CA, USA}, + booktitle = {OOPSLA}, + year = 2003, + isbn = {1-58113-712-5}, + address = pub-acm:adr, + publisher = pub-acm +} + +@InProceedings{ lee.ea:lightweight:2007, + author = {Hannah K. Lee and Heiko and Luedemann}, + title = {lightweight decentralized authorization model for + inter-domain collaborations}, + booktitle = {SWS '07: Proceedings of the 2007 ACM workshop on Secure + web services}, + year = 2007, + isbn = {978-1-59593-892-3}, + pages = {83--89}, + location = {Fairfax, Virginia, USA}, + doi = {10.1145/1314418.1314431}, + address = pub-acm:adr, + publisher = pub-acm +} + +@InProceedings{ freudenthal.ea:drbac:2002, + title = {{dRBAC}: distributed role-based access control for dynamic + coalition environments}, + author = {Freudenthal, E. and Pesin, T. and Port, L. and Keenan, E. + and Karamcheti, V.}, + journal = {Distributed Computing Systems, 2002. Proceedings. 22nd + International Conference on}, + year = 2002, + pages = {411--420}, + keywords = {authorisation, distributed processing PKI identities, + continuous monitoring, controlled activities, credential + discovery, credential validation, dRBAC, distributed + role-based access control, dynamic coalition environments, + graph approach, long-lived interactions, multiple + administrative domains, namespaces, policy roots, role + delegation, scalable decentralized access control + mechanism, scalable decentralized trust-management + mechanism, scalar valued attributes, third-party + delegation, transferred permissions, trust domains, trust + relationships}, + doi = {10.1109/ICDCS.2002.1022279}, + issn = {1063-6927 }, + abstract = {distributed role-based access control (dRBAC) is a + scalable, decentralized trust-management and access-control + mechanism for systems that span multiple administrative + domains. dRBAC utilizes PKI identities to define trust + domains, roles to define controlled activities, and role + delegation across domains to represent permissions to these + activities. The mapping of controlled actions to roles + enables their namespaces to serve as policy roots. dRBAC + distinguishes itself from previous approaches by providing + three features: (1) third-party delegation of roles from + outside a domain's namespace, relying upon an explicit + delegation of assignment; (2) modulation of transferred + permissions using scalar valued attributes associated with + roles; and (3) continuous monitoring of trust relationships + over long-lived interactions. The paper describes the dRBAC + model and its scalable implementation using a graph + approach to credential discovery and validation.} +} + +@Article{ liu.ea:role-based:2004, + author = {Duen-Ren Liu and Mei-Yu Wu and Shu-Teng Lee}, + title = {Role-based authorizations for workflow systems in support + of task-based separation of duty}, + journal = {Journal of Systems and Software}, + volume = 73, + number = 3, + year = 2004, + pages = {375--387}, + doi = {10.1016/S0164-1212(03)00175-4}, + abstract = {Role-based authorizations for assigning tasks of workflows + to roles/users are crucial to security management in + workflow management systems. The authorizations must + enforce separation of duty (SoD) constraints to prevent + fraud and errors. This work analyzes and defines several + duty-conflict relationships among tasks, and designs + authorization rules to enforce SoD constraints based on the + analysis. A novel authorization model that incorporates + authorization rules is then proposed to support the + planning of assigning tasks to roles/users, and the + run-time activation of tasks. Different from existing work, + the proposed authorization model considers the AND/XOR + split structures of workflows and execution dependency + among tasks to enforce separation of duties in assigning + tasks to roles/users. A prototype system is developed to + realize the effectiveness of the proposed authorization + model.} +} + +@Booklet{ nipkow.ea:isabelle-hol:2009, + title = {{Isabelle's} Logic: {HOL}}, + author = {Tobias Nipkow and Lawrence C. Paulson and Markus Wenzel}, + year = 2009, + misc = {\url{http://isabelle.in.tum.de/library/HOL/}} +} + +@InProceedings{ garson.ea:security:2008, + author = {Garson, Kathryn and Adams, Carlisle}, + title = {Security and privacy system architecture for an e-hospital + environment}, + booktitle = {IDtrust '08: Proceedings of the 7th symposium on Identity + and trust on the Internet}, + year = 2008, + isbn = {978-1-60558-066-1}, + pages = {122--130}, + location = {Gaithersburg, Maryland}, + doi = {10.1145/1373290.1373306}, + address = pub-acm:adr, + publisher = pub-acm +} + +@Article{ kambourakis.ea:pki-based:2005, + author = {G. Kambourakis and I. Maglogiannis and A. Rouskas}, + title = {PKI-based secure mobile access to electronic health + services and data}, + journal = {Technology and Health Care Journal}, + volume = 13, + number = 6, + year = 2005, + issn = {0928-7329}, + pages = {511--526}, + publisher = pub-ios, + address = pub-ios:adr, + abstract = {Recent research works examine the potential employment of + public-key cryptography schemes in e-health environments. + In such systems, where a Public Key Infrastructure (PKI) is + established beforehand, Attribute Certificates (ACs) and + public key enabled protocols like TLS, can provide the + appropriate mechanisms to effectively support + authentication, authorization and confidentiality services. + In other words, mutual trust and secure communications + between all the stakeholders, namely physicians, patients + and e-health service providers, can be successfully + established and maintained. Furthermore, as the recently + introduced mobile devices with access to computer-based + patient record systems are expanding, the need of + physicians and nurses to interact increasingly with such + systems arises. Considering public key infrastructure + requirements for mobile online health networks, this paper + discusses the potential use of Attribute Certificates (ACs) + in an anticipated trust model. Typical trust interactions + among doctors, patients and e-health providers are + presented, indicating that resourceful security mechanisms + and trust control can be obtained and implemented. The + application of attribute certificates to support medical + mobile service provision along with the utilization of the + de-facto TLS protocol to offer competent confidentiality + and authorization services is also presented and evaluated + through experimentation, using both the 802.11 WLAN and + General Packet Radio Service (GPRS) networks.} +} + +@InProceedings{ sahai.ea:fuzzy:2005, + author = {Amit Sahai and Brent Waters}, + title = {Fuzzy Identity-Based Encryption}, + year = 2005, + pages = {457--473}, + doi = {10.1007/11426639_27}, + crossref = {cramer:advances:2005}, + abstract = {We introduce a new type of Identity-Based Encryption (IBE) + scheme that we call Fuzzy Identity-Based Encryption. In + Fuzzy IBE we view an identity as set of descriptive + attributes. A Fuzzy IBE scheme allows for a private key for + an identity, ohgr, to decrypt a ciphertext encrypted with + an identity, ohgr prime, if and only if the identities ohgr + and ohgr prime are close to each other as measured by the + ldquoset overlaprdquo distance metric. A Fuzzy IBE scheme + can be applied to enable encryption using biometric inputs + as identities; the error-tolerance property of a Fuzzy IBE + scheme is precisely what allows for the use of biometric + identities, which inherently will have some noise each time + they are sampled. Additionally, we show that Fuzzy-IBE can + be used for a type of application that we term + ldquoattribute-based encryptionrdquo. In this paper we + present two constructions of Fuzzy IBE schemes. Our + constructions can be viewed as an Identity-Based Encryption + of a message under several attributes that compose a + (fuzzy) identity. Our IBE schemes are both error-tolerant + and secure against collusion attacks. Additionally, our + basic construction does not use random oracles. We prove + the security of our schemes under the Selective-ID security + model. } +} + +@Proceedings{ cramer:advances:2005, + editor = {Ronald Cramer}, + booktitle = PROC # {International Conference on the Theory and Applications of + Cryptographic Techniques (EUROCRYPT)}, + location = {Advances in Cryptology - EUROCRYPT 2005, 24th Annual + International Conference on the Theory and Applications of + Cryptographic Techniques, Aarhus, Denmark, May 22-26, 2005, + Proceedings}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 3494, + year = 2005, + isbn = {3-540-25910-4} +} + +@InProceedings{ goyal.ea:attribute-based:2006, + author = {Vipul Goyal and Omkant Pandey and Amit Sahai and Brent + Waters}, + title = {Attribute-based encryption for fine-grained access control + of encrypted data}, + booktitle = {CCS '06: Proceedings of the 13th ACM conference on + Computer and communications security}, + year = 2006, + isbn = {1-59593-518-5}, + pages = {89--98}, + location = {Alexandria, Virginia, USA}, + doi = {10.1145/1180405.1180418}, + address = pub-acm:adr, + publisher = pub-acm, + abstract = {As more sensitive data is shared and stored by third-party + sites on the Internet, there will be a need to encrypt data + stored at these sites. One drawback of encrypting data, is + that it can be selectively shared only at a coarse-grained + level (i.e., giving another party your private key). We + develop a new cryptosystem for fine-grained sharing of + encrypted data that we call Key-Policy Attribute-Based + Encryption (KP-ABE). In our cryptosystem, ciphertexts are + labeled with sets of attributes and private keys are + associated with access structures that control which + ciphertexts a user is able to decrypt. We demonstrate the + applicability of our construction to sharing of audit-log + information and broadcast encryption. Our construction + supports delegation of private keys which + subsumesHierarchical Identity-Based Encryption (HIBE).} +} + +@InProceedings{ li.ea:privacy-aware:2009, + author = {Jin Li and Kui Ren and Bo Zhu and Zhiguo Wan}, + title = {Privacy-aware Attribute-based Encryption with User + Accountability}, + booktitle = {The 12th Information Security Conference (ISC'09)}, + location = {September 7-9, 2009, Pisa}, + year = 2009, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + abstract = {As a new public key primitive, attribute-based encryption + (ABE) is envisioned to be a promising tool for implementing + fine-grained access control. To further address the concern + of user access privacy, privacy-aware ABE schemes are being + developed to achieve hidden access policy recently. For the + purpose of secure access control, there is, how- ever, + still one critical functionality missing in the existing + ABE schemes, which is user accountability. Currently, no + ABE scheme can completely prevent the problem of illegal + key sharing among users. In this paper, we tackle this + problem by firstly proposing the notion of accountable, + anonymous, and ciphertext-policy ABE (CP-A3 BE, in short) + and then giving out a concrete construction. We start by + improving the state-of-the-art of anonymous CP-ABE to + obtain shorter public parameters and ciphertext length. In + the proposed CP-A3 BE construction, user accountability can + be achieved in black-box model by embedding additional + user-specific information into the attribute private key + issued to that user, while still maintaining hidden access + policy. The proposed constructions are provably secure.} +} + +@InProceedings{ bobba.ea:pbes:2009, + author = {Rakesh Bobba and and Himanshu Khurana and Musab AlTurki + and Farhana Ashraf}, + title = {PBES: a policy based encryption system with application to + data sharing in the power grid}, + booktitle = {ASIACCS '09: Proceedings of the 4th International + Symposium on Information, Computer, and Communications + Security}, + year = 2009, + isbn = {978-1-60558-394-5}, + pages = {262--275}, + location = {Sydney, Australia}, + doi = {10.1145/1533057.1533093}, + address = pub-acm:adr, + publisher = pub-acm, + abstract = {In distributed systems users need the ability to share + sensitive content with multiple other recipients based on + their ability to satisfy arbitrary policies. One such + system is electricity grids where finegrained sensor data + sharing holds the potential for increased reliability and + efficiency. However, effective data sharing requires + technical solutions that support flexible access policies, + for example, sharing more data when the grid is unstable. + In such systems, both the messages and policies are + sensitive and, therefore, they need to kept be secret. + Furthermore, to allow for such a system to be secure and + usable in the presence of untrusted object stores and + relays it must be resilient in the presence of active + adversaries and provide efficient key management. While + several of these properties have been studied in the past + we address a new problem in the area of policy based + encryption in that we develop a solution with all of these + capabilities. We develop a Policy and Key Encapsulation + Mechanism -- Data Encapsulation Mechanism (PKEM-DEM) + encryption scheme that is a generic construction secure + against adaptive chosen ciphertext attacks and develop a + Policy Based Encryption System (PBES) using this scheme + that provides these capabilities. We provide an + implementation of PBES and measure its performance.} +} + +@InProceedings{ shanqing.ea:attribute-based:2008, + author = {Guo Shanqing and Zeng Yingpei}, + title = {Attribute-based Signature Scheme}, + doi = {10.1109/ISA.2008.111}, + booktitle = {International Conference on Information Security and + Assurance, 2008 (ISA 2008)}, + year = 2008, + pages = {509--511}, + month = apr, + address = pub-ieee:adr, + publisher = pub-ieee, + abstract = {In real life, one requires signatures from people who + satisfy certain criteria like that they should possess some + specific attributes. For example, Alice wants a document to + be signed by some employee in Bob's company. This employee + must have certain attributes such as being part of the IT + staff and at least a junior manager in the cryptography + team or a senior manager in the biometrics team. In order + to satisfy these kinds of needs, we defined a common + Attribute-based signature scheme where the signing member + has to have certain attributes or belong to a certain + group, and we also proved our scheme to be secure.} +} + +@Article{ huang.ea:aspe:2009, + title = {ASPE: attribute-based secure policy enforcement in + vehicular ad hoc networks}, + journal = {Ad Hoc Networks}, + volume = 7, + number = 8, + pages = {1526--1535}, + year = 2009, + mynote = {Privacy and Security in Wireless Sensor and Ad Hoc + Networks}, + issn = {1570-8705}, + doi = {10.1016/j.adhoc.2009.04.011}, + author = {Dijiang Huang and Mayank Verma}, + keywords = {Security, Vehicular networks, Secure group communications, + Key management, Attribute based encryption}, + publisher = pub-elsevier, + address = pub-elsevier:adr, + abstract = {Vehicular ad hoc networks (VANETs) are usually operated + among vehicles moving at high speeds, and thus their + communication relations can be changed frequently. In such + a highly dynamic environment, establishing trust among + vehicles is difficult. To solve this problem, we propose a + flexible, secure and decentralized attribute based secure + key management framework for VANETs. Our solution is based + on attribute based encryption (ABE) to construct an + attribute based security policy enforcement (ASPE) + framework. ASPE considers various road situations as + attributes. These attributes are used as encryption keys to + secure the transmitted data. ASPE is flexible in that it + can dynamically change encryption keys depending on the + VANET situations. At the same time, ASPE naturally + incorporates data access control policies on the + transmitted data. ASPE provides an integrated solution to + involve data access control, key management, security + policy enforcement, and secure group formation in highly + dynamic vehicular communication environments. Our + performance evaluations show that ASPE is efficient and it + can handle large amount of data encryption/decryption flows + in VANETs.} +} + +@InProceedings{ weber:securing:2009, + author = {Stefan G. Weber}, + title = {Securing First Response Coordination With Dynamic + Attribute-Based Encryption}, + booktitle = {World Congress on Privacy, Security, Trust and the + Management of e-Business (CONGRESS)}, + year = 2009, + pages = {58--69}, + isbn = {978-0-7695-3805-1}, + address = pub-ieee:adr, + publisher = pub-ieee, + abstract = {Every minute saved in emergency management processes can + save additional lives of affected victims. Therefore, an + effective coordination of the incident reactions of mobile + first responders is very important, especially in the face + of rapidly changing situations of large scale disasters. + However, tactical communication and messaging between the + headquarter and mobile first responders, initiated for + coordination purposes, has to meet a strong security + requirements: it must preserve confidentiality in order to + prevent malicious third parties from disrupting the + reactions. This paper presents concepts to support the + secure coordination of mobile first responders by providing + means for secure ubiquitous tactical communication. Our + concept harnesses ciphertext-policy attribute-based + encryption (CP-ABE) techniques. We extend current CP-ABE + proposals by additionally taking into account dynamic + factors: our proposed system is able to handle dynamic + attributes, like current status of duty and location of + mobile first responders, in a secure fashion, in order to + support flexible specification of receiver groups of + tactical messages, while end-to-end encryption in the + messaging process is still satisfied.} +} + +@InProceedings{ cardoso:approaches:2006, + author = {Jorge Cardoso}, + title = {Approaches to Compute Workflow Complexity}, + booktitle = {The Role of Business Processes in Service Oriented + Architectures}, + year = 2006, + ee = {http://drops.dagstuhl.de/opus/volltexte/2006/821}, + crossref = {leymann.ea:role:2006}, + abstract = {During the last 20 years, complexity has been an + interesting topic that has been investigated in many fields + of science, such as biology, neurology, software + engineering, chemistry, psychology, and economy. A survey + of the various approaches to understand complexity has lead + sometimes to a measurable quantity with a rigorous but + narrow definition and other times as merely an ad hoc + label. In this paper we investigate the complexity concept + to avoid a vague use of the term `complexity' in workflow + designs. We present several complexity metrics that have + been used for a number of years in adjacent fields of + science and explain how they can be adapted and use to + evaluate the complexity of workflows. } +} + +@Proceedings{ leymann.ea:role:2006, + editor = {Frank Leymann and Wolfgang Reisig and Satish R. Thatte and + Wil M. P. van der Aalst}, + title = {The Role of Business Processes in Service Oriented + Architectures, 16.07. - 21.07.2006}, + booktitle = {The Role of Business Processes in Service Oriented + Architectures}, + publisher = {Internationales Begegnungs- und Forschungszentrum fuer + Informatik (IBFI), S chloss Dagstuhl, Germany}, + series = {Dagstuhl Seminar Proceedings}, + volume = 06291, + year = 2006 +} + +@Article{ parnas:stop:2007, + author = {David Lorge Parnas}, + title = {Stop the numbers game}, + journal = j-cacm, + volume = 50, + number = 11, + year = 2007, + issn = {0001-0782}, + pages = {19--21}, + doi = {10.1145/1297797.1297815}, + address = pub-acm:adr, + publisher = pub-acm +} + + + + +@InBook{ gentry:ibe:2006, + author = {Craig Gentry}, + chapter = {IBE (Identity-Based Encryption)}, + title = {Handbook of Information Security}, + editor = {Hossein Bidgoli}, + volume = 2, + isbn = {0-471-64833-7}, + publisher = {John Wiley \& Sons}, + pages = {575--592}, + month = jan, + year = 2006 +} + +@InProceedings{ bethencourt.ea:ciphertext-policy:2007, + author = {John Bethencourt and Amit Sahai and Brent Waters}, + title = {Ciphertext-Policy Attribute-Based Encryption}, + booktitle = {IEEE Symposium on Security and Privacy}, + pages = {321--334}, + year = 2007, + doi = {10.1109/SP.2007.11}, + abstract = {In several distributed systems a user should only be able + to access data if a user posses a certain set of + credentials or attributes. Currently, the only method for + enforcing such policies is to employ a trusted server to + store the data and mediate access control. However, if any + server storing the data is compromised, then the + confidentiality of the data will be compromised. In this + paper we present a system for realizing complex access + control on encrypted data that we call ciphertext-policy + attribute-based encryption. By using our techniques + encrypted data can be kept confidential even if the storage + server is untrusted; moreover, our methods are secure + against collusion attacks. Previous attribute-based + encryption systems used attributes to describe the + encrypted data and built policies into user's keys; while + in our system attributes are used to describe a user's + credentials, and a party encrypting data determines a + policy for who can decrypt. Thus, our methods are + conceptually closer to traditional access control methods + such as role-based access control (RBAC). In addition, we + provide an implementation of our system and give + performance measurements.}, + publisher = pub-ieee, + address = pub-ieee:adr +} + +@Article{ karedla.ea:caching:1994, + author = {Ramakrishna Karedla and J. Spencer Love and Bradley G. + Wherry}, + title = {Caching strategies to improve disk system performance}, + journal = j-computer, + volume = 27, + number = 3, + issn = {0018-9162}, + year = 1994, + pages = {38--46}, + doi = {10.1109/2.268884}, + publisher = pub-ieee, + address = pub-ieee:adr, + abstract = {I/O subsystem manufacturers attempt to reduce latency by + increasing disk rotation speeds, incorporating more + intelligent disk scheduling algorithms, increasing I/O bus + speed, using solid-state disks, and implementing caches at + various places in the I/O stream. In this article, we + examine the use of caching as a means to increase system + response time and improve the data throughput of the disk + subsystem. Caching can help to alleviate I/O subsystem + bottlenecks caused by mechanical latencies. This article + describes a caching strategy that offers the performance of + caches twice its size. After explaining some basic caching + issues, we examine some popular caching strategies and + cache replacement algorithms, as well as the advantages and + disadvantages of caching at different levels of the + computer system hierarchy. Finally, we investigate the + performance of three cache replacement algorithms: random + replacement (RR), least recently used (LRU), and a + frequency-based variation of LRU known as segmented LRU + (SLRU). } +} + +@InProceedings{ megiddo.ea:arc:2003, + author = {Nimrod Megiddo and Dharmendra S. Modha}, + title = {{ARC}: A Self-Tuning, Low Overhead Replacement Cache}, + booktitle = {FAST '03: Proceedings of the 2nd USENIX Conference on File + and Storage Technologies}, + year = 2003, + pages = {115--130}, + location = {San Francisco, CA}, + publisher = {\acs{usenix} Association}, + address = {Berkeley, CA, USA} +} + +@InProceedings{ chou.ea:evaluation:1985, + author = {Hong-Tai Chou and David J. DeWitt}, + title = {An evaluation of buffer management strategies for + relational database systems}, + booktitle = {VLDB '1985: Proceedings of the 11th international + conference on Very Large Data Bases}, + year = 1985, + pages = {127--141}, + location = {Stockholm, Sweden}, + publisher = {VLDB Endowment} +} + +@InProceedings{ yu.ea:fdac:2009, + author = {Shucheng Yu and Kui Ren and Wenjing Lou}, + title = {{FDAC}: Toward fine-grained distributed data access + control in wireless sensor networks}, + booktitle = PROC # { \acs{ieee} Conference on Computer + Communications (INFOCOM)}, + year = 2009, + address = pub-ieee:adr, + publisher = pub-ieee, + abstract = {Distributed sensor data storage and retrieval has gained + increasing popularity in recent years for supporting + various applications. While distributed architecture enjoys + a more robust and fault-tolerant wireless sensor network + (WSN), such architecture also poses a number of security + challenges especially when applied in mission-critical + applications such as battle field and e-healthcare. First, + as sensor data are stored and maintained by individual + sensors and unattended sensors are easily subject to strong + attacks such as physical compromise, it is significantly + harder to ensure data security. Second, in many + mission-critical applications, fine-grained data access + control is a must as illegal access to the sensitive data + may cause disastrous result and/or prohibited by the law. + Last but not least, sensors usually are resource-scarce, + which limits the direct adoption of expensive cryptographic + primitives. To address the above challenges, we propose in + this paper a distributed data access control scheme that is + able to fulfill fine-grained access control over sensor + data and is resilient against strong attacks such as sensor + compromise and user colluding. The proposed scheme exploits + a novel cryptographic primitive called attribute-based + encryption (ABE), tailors, and adapts it for WSNs with + respect to both performance and security requirements. The + feasibility of the scheme is demonstrated by experiments on + real sensor platforms. To our best knowledge, this paper is + the first to realize distributed fine-grained data access + control for WSNs.} +} + +@InProceedings{ traynor.ea:massive-scale:2008, + author = {Patrick Traynor and Kevin R. B. Butler and William Enck + and Patrick McDaniel}, + title = {Realizing Massive-Scale Conditional Access Systems Through + Attribute-Based Cryptosystems}, + booktitle = PROC # { 15th Annual Network and Distributed System + Security Symposium (NDSS 2008)}, + year = 2008, + abstract = {The enormous growth in the diversity of content services + such as IPtv has highlighted the inadequacy of the + accompanying content security: existing security mechanisms + scale poorly, require complex and often costly dedicated + hardware, or fail to meet basic security requirements. New + security methods are needed. In this paper, we explore the + ability of attribute-based encryption (ABE) to meet the + unique performance and security requirements of conditional + access systems such as subscription radio and payper- view + television. We show through empirical study that costs of + ABE make its direct application inappropriate, but present + constructions that mitigate its incumbent costs. We develop + an extensive simulation that allows us to explore the + performance of a number of virtual hardware configurations + and construction parameters over workloads developed from + real subscription and television audiences. These + simulations show that we can securely deliver high quality + content to viewerships of the highest rated shows being + broadcast today, some in excess of 26,000,000 viewers. It + is through these experiments that we expose the viability + of not only ABE-based content delivery, but applicability + of ABE systems to large-scale distributed systems.}, + url = {http://www.isoc.org/isoc/conferences/ndss/08/papers/06_realizing_massive-scale_conditional.pdf} + , + location = {San Diego, California, USA}, + month = feb, + publisher = {The Internet Society} +} + +@Article{ alpern.ea:defining:1985, + author = {Bowen Alpern and Fred B. Schneider}, + title = {Defining Liveness}, + journal = j-ipl, + volume = 21, + number = 4, + year = 1985, + pages = {181--185} +} + +@Article{ lamport:proving:1977, + title = {Proving the Correctness of Multiprocess Programs}, + author = {Leslie Lamport}, + journal = j-tse, + year = 1977, + month = {March}, + volume = {SE-3}, + number = 2, + pages = {125--143}, + abstract = {The inductive assertion method is generalized to permit + formal, machine-verifiable proofs of correctness for + multiprocess programs. Individual processes are represented + by ordinary flowcharts, and no special synchronization + mechanisms are assumed, so the method can be applied to a + large class of multiprocess programs. A correctness proof + can be designed together with the program by a hierarchical + process of stepwise refinement, making the method practical + for larger programs. The resulting proofs tend to be + natural formalizations of the informal proofs that are now + used.}, + keywords = {null Assertions, concufrent programming, correctness, + multiprocessing, synchronization}, + doi = {10.1109/TSE.1977.229904}, + issn = {0098-5589} +} + +@InProceedings{ goodenough.ea:toward:1975, + author = {John B. Goodenough and Susan L. Gerhart}, + title = {Toward a theory of test data selection}, + booktitle = PROC # { the international conference on Reliable + software}, + year = 1975, + pages = {493--510}, + url = {http://portal.acm.org/citation.cfm?id=808473&dl=ACM&coll=GUIDE#} + , + location = {Los Angeles, California}, + abstract = {This paper examines the theoretical and practical role of + testing in software development. We prove a fundamental + theorem showing that properly structured tests are capable + of demonstrating the absence of errors in a program. The + theorem's proof hinges on our definition of test + reliability and validity, but its practical utility hinges + on being able to show when a test is actually reliable. We + explain what makes tests unreliable (for example, we show + by example why testing all program statements, predicates, + or paths is not usually sufficient to insure test + reliability), and we outline a possible approach to + developing reliable tests. We also show how the analysis + required to define reliable tests can help in checking a + program's design and specifications as well as in + preventing and detecting implementation errors. }, + acknowledgement={none} +} + +@InCollection{ aczel:introduction:1977, + author = {Peter Aczel}, + title = {An Introduction to Inductive Definitions}, + booktitle = {Handbook of Mathematical Logic}, + editor = {Jon Barwise}, + series = {Studies in Logic and the Foundations of Mathematics}, + volume = 90, + chapter = {C.7}, + pages = {739--782}, + publisher = {North-Holland}, + address = {Amsterdam}, + year = 1977 +} + +@InProceedings{ cousot.ea:abstract:1977, + author = {Patrick Cousot and Radhia Cousot}, + title = {Abstract interpretation: a unified lattice model for + static analysis of programs by construction or + approximation of fixpoints}, + booktitle = PROC # { the 4th ACM SIGACT-SIGPLAN symposium on + Principles of programming languages}, + year = 1977, + pages = {238--252}, + location = {Los Angeles, California}, + doi = {10.1145/512950.512973}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none} +} + +@Article{ cardelli.ea:understanding:1985, + author = {Luca Cardelli and Peter Wegner}, + title = {On understanding types, data abstraction, and + polymorphism}, + journal = {ACM Computing Surveys}, + volume = 17, + number = 4, + year = 1985, + issn = {0360-0300}, + acknowledgement={none}, + pages = {471--523}, + doi = {10.1145/6041.6042}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {Our objective is to understand the notion of type in + programming languages, present a model of typed, + polymorphic programming languages that reflects recent + research in type theory, and examine the relevance of + recent research to the design of practical programming + languages. Object-oriented languages provide both a + framework and a motivation for exploring the interaction + among the concepts of type, data abstraction, and + polymorphism, since they extend the notion of type to data + abstraction and since type inheritance is an important form + of polymorphism. We develop a &lgr;-calculus-based model + for type systems that allows us to explore these + interactions in a simple setting, unencumbered by + complexities of production programming languages. The + evolution of languages from untyped universes to + monomorphic and then polymorphic type systems is reviewed. + Mechanisms for polymorphism such as overloading, coercion, + subtyping, and parameterization are examined. A unifying + framework for polymorphic type systems is developed in + terms of the typed &lgr;-calculus augmented to include + binding of types by quantification as well as binding of + values by abstraction. The typed &lgr;-calculus is + augmented by universal quantification to model generic + functions with type parameters, existential quantification + and packaging (information hiding) to model abstract data + types, and bounded quantification to model subtypes and + type inheritance. In this way we obtain a simple and + precise characterization of a powerful type system that + includes abstract data types, parametric polymorphism, and + multiple inheritance in a single consistent framework. The + mechanisms for type checking for the augmented + &lgr;-calculus are discussed. The augmented typed + &lgr;-calculus is used as a programming language for a + variety of illustrative examples. We christen this language + Fun because fun instead of &lgr; is the functional + abstraction keyword and because it is pleasant to deal + with. Fun is mathematically simple and can serve as a basis + for the design and implementation of real programming + languages with type facilities that are more powerful and + expressive than those of existing programming languages. In + afsyped object-oriented languages. } +} + +@InProceedings{ wadler:listlessness:1985, + author = {Philip Wadler}, + title = {Listlessness is better than laziness II: composing + listless functions}, + booktitle = {on Programs as data objects}, + year = 1985, + isbn = {0-387-16446-4}, + pages = {282--305}, + location = {Copenhagen, Denmark}, + publisher = pub-springer, + address = pub-springer:adr, + acknowledgement={none} +} + +@InProceedings{ clement.ea:simple:1986, + author = {Dominique Cl\'ement and Thierry Despeyroux and Gilles Kahn + and Jo\"elle Despeyroux}, + title = {A simple applicative language: {mini-ML}}, + booktitle = {LFP '86: Proceedings of the 1986 ACM conference on LISP + and functional programming}, + year = 1986, + isbn = {0-89791-200-4}, + pages = {13--27}, + location = {Cambridge, Massachusetts, United States}, + doi = {10.1145/319838.319847}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none} +} + +@InProceedings{ hamlet:theoretical:1989, + author = {R. Hamlet}, + title = {Theoretical comparison of testing methods}, + booktitle = PROC # { the ACM SIGSOFT '89 third symposium on + Software testing, analysis, and verification}, + year = 1989, + isbn = {0-89791-342-6}, + pages = {28--37}, + location = {Key West, Florida, United States}, + doi = {10.1145/75308.75313}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none} +} + +@Article{ frankl.ea:applicable:1988, + author = {P. G. Frankl and E. J. Weyuker}, + title = {An Applicable Family of Data Flow Testing Criteria}, + journal = j-ieee-tse, + volume = 14, + number = 10, + year = 1988, + month = oct, + issn = {0098-5589}, + pages = {1483--1498}, + doi = {http://dx.doi.org/10.1109/32.6194}, + publisher = pub-ieee, + abstract = {he authors extend the definitions of the previously + introduced family of data flow testing criteria to apply to + programs written in a large subset of Pascal. They then + define a family of adequacy criteria called feasible data + flow testing criteria, which are derived from the data-flow + testing criteria. The feasible data flow testing criteria + circumvent the problem of nonapplicability of the data flow + testing criteria by requiring the test data to exercise + only those definition-use associations which are + executable. It is shown that there are significant + differences between the relationships among the data flow + testing criteria and the relationships among the feasible + data flow testing criteria. The authors discuss a + generalized notion of the executability of a path through a + program unit. A script of a testing session using their + data flow testing tool, ASSET, is included.}, + acknowledgement={none} +} + +@InProceedings{ teo.ea:use:1988, + author = {Ghee S. Teo and M\'{\i}che{\'a}l Mac an Airchinnigh}, + title = {The Use of VDM in the Specification of Chinese + Characters.}, + booktitle = {VDM Europe}, + year = 1988, + pages = {476--499}, + crossref = {bloomfield.ea:vdm:1988}, + acknowledgement={none} +} + +@Proceedings{ bloomfield.ea:vdm:1988, + editor = {Robin E. Bloomfield and Lynn S. Marshall and Roger B. + Jones}, + title = {VDM '88, VDM - The Way Ahead, 2nd VDM-Europe Symposium, + Dublin, Ireland, September 11-16, 1988, Proceedings}, + booktitle = {VDM Europe}, + publisher = pub-springer, + series = lncs, + volume = 328, + year = 1988, + isbn = {3-540-50214-9}, + acknowledgement={none}, + bibkey = {bloomfield.ea:vdm:1988} +} + +@InProceedings{ cook.ea:inheritance:1990, + author = {William R. Cook and Walter Hill and Peter S. Canning}, + title = {Inheritance is not subtyping}, + booktitle = {POPL '90: Proceedings of the 17th ACM SIGPLAN-SIGACT + symposium on Principles of programming languages}, + year = 1990, + isbn = {0-89791-343-4}, + pages = {125--135}, + location = {San Francisco, California, United States}, + doi = {10.1145/96709.96721}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none}, + abstract = {In typed object-oriented languages the subtype relation is + typically based on the inheritance hierarchy. This + approach, however, leads either to insecure type-systems or + to restrictions on inheritance that make it less flexible + than untyped Smalltalk inheritance. We present a new typed + model of inheritance that allows more of the flexibility of + Smalltalk inheritance within a statically-typed system. + Significant features of our analysis are the introduction + of polymorphism into the typing of inheritance and the + uniform application of inheritance to objects, classes and + types. The resulting notion of type inheritance allows us + to show that the type of an inherited object is an + inherited type but not always a subtype. }, + bibkey = {cook.ea:inheritance:1990} +} + +@Article{ lynch.ea:forward:1996, + author = {Nancy Lynch and Frits Vaandrager}, + title = {Forward and backward simulations II.: timing-based + systems}, + journal = {Inf. Comput.}, + volume = 128, + number = 1, + year = 1996, + issn = {0890-5401}, + pages = {1--25}, + doi = {10.1006/inco.1996.0060}, + publisher = {Academic Press, Inc.}, + acknowledgement={none}, + bibkey = {lynch.ea:forward:1996} +} + +@Article{ lamport.ea:should:1999, + author = {Leslie Lamport and Lawrence C. Paulson}, + title = {Should your specification language be typed.}, + journal = {ACM Trans. Program. Lang. Syst.}, + volume = 21, + number = 3, + year = 1999, + acknowledgement={none}, + pages = {502--526}, + publisher = pub-acm, + address = pub-acm:adr, + issn = {0164-0925}, + doi = {10.1145/319301.319317} +} + +@InProceedings{ muller.ea:formal:1997, + title = {Formal Specification Techniques for Object-Oriented + Programs }, + author = {Peter M{\"u}ller and Arnd Poetzsch-Heffter }, + editor = {Jarke, M. and Pasedach, K. and Pohl, K. }, + booktitle = {Informatik 97: Informatik als Innovationsmotor }, + series = {Informatik Aktuell }, + publisher = pub-springer, + address = pub-springer:adr, + acknowledgement={none}, + year = 1997, + abstract = {Specification techniques for object-oriented programs + relate the operational world of programs to the declarative + world of specifications. We present a formal foundation of + interface specification languages. Based on the formal + foundation, we develop new specification techniques to + describe functional behavior, invariants, and side-effects. + Furthermore, we discuss the influence of program extensions + on program correctness.} +} + +@Article{ wolper:meaning:1997, + author = {Pierre Wolper}, + title = {The Meaning of ``Formal'': From Weak to Strong Formal + Methods.}, + journal = j-sttt, + volume = 1, + publisher = pub-springer, + address = pub-springer:adr, + number = {1-2}, + year = 1997, + pages = {6--8}, + doi = {10.1007/s100090050002}, + acknowledgement={none} +} + +@TechReport{ aredo.ea:towards:1999, + title = {Towards a formalization of {UML} Class Structure in + {PVS}}, + author = {Demissie B. Aredo and I. Traore and K. St{\o}len}, + institution = {Department of Informatics, University of Oslo}, + year = 1999, + month = aug, + number = 272, + acknowledgement={none} +} + +@Book{ derrick.ea:refinement:2001, + author = {John Derrick and Eerke Boiten}, + title = {Refinement in {Z} and {Object-Z}}, + library = {ETH-BIB}, + publisher = pub-springer, + address = pub-springer:adr, + isbn = {1-85233-245-X}, + url = {http://www.cs.kent.ac.uk/people/staff/jd1/books/refine/}, + year = 2001, + acknowledgement={none}, + bibkey = {derrick.ea:refinement:2001} +} + +@TechReport{ boer.ea:towards:2003, + author = {Frank S. de Boer and Cees Pierik}, + year = 2003, + title = {Towards an environment for the verification of annotated + object-oriented programs}, + number = {UU-CS-2003-002}, + institution = {Institute of Information and Computing Sciences, Utrecht + University}, + acknowledgement={none} +} + +@Article{ dybjer.ea:verifying:2004, + author = {Peter Dybjer and Qiao Haiyana and Makoto Takeyama}, + booktitle = {Third International Conference on Quality Software: QSIC + 2003}, + title = {Verifying Haskell programs by combining testing, model + checking and interactive theorem proving}, + journal = {Information and Software Technology}, + year = 2004, + number = 15, + volume = 46, + pages = {1011--1025}, + doi = {10.1016/j.infsof.2004.07.002}, + abstract = {We propose a program verification method that combines + random testing, model checking and interactive theorem + proving. Testing and model checking are used for debugging + programs and specifications before a costly interactive + proof attempt. During proof development, testing and model + checking quickly eliminate false conjectures and generate + counterexamples which help to correct them. With an + interactive theorem prover we also ensure the correctness + of the reduction of a top level problem to subproblems that + can be tested or proved. We demonstrate the method using + our random testing tool and binary decision diagrams-based + (BDDs) tautology checker, which are added to the Agda/Alfa + interactive proof assistant for dependent type theory. In + particular we apply our techniques to the verification of + Haskell programs. The first example verifies the BDD + checker itself by testing its components. The second uses + the tautology checker to verify bitonic sort together with + a proof that the reduction of the problem to the checked + form is correct.}, + acknowledgement={none} +} + +@PhDThesis{ kopylov:type:2004, + author = {Alexei Pavlovich Kopylov}, + title = {Type Theoretical Foundations for Data Structures, Classes, + and Objects}, + school = {Cornell University}, + year = 2004, + abstract = {In this thesis we explore the question of how to represent + programming data structures in a constructive type theory. + The basic data structures in programing languages are + records and objects. Most known papers treat such data + structure as primitive. That is, they add new primitive + type constructors and sup- porting axioms for records and + objects. This approach is not satisfactory. First of all it + complicates a type theory a lot. Second, the validity of + the new axioms is not easily established. As we will see + the naive choice of axioms can lead to contradiction even + in the simplest cases. We will show that records and + objects can be defined in a powerful enough type theory. We + will also show how to use these type constructors to define + abstract data structure. }, + acknowledgement={none}, + month = jan +} + +@InProceedings{ kyas.ea:message:2004, + author = {Marcel Kyas and Frank S. de Boer}, + title = {On Message Specification in {OCL}}, + booktitle = {Compositional Verification in UML}, + year = 2004, + editor = {Frank S. de Boer and Marcello Bonsangue}, + acknowledgement={none}, + series = entcs, + volume = 101, + pages = {73--93}, + publisher = elsevier, + address = elsevier:adr, + abstract = {The object constraint language (OCL) is the established + language for specifying of properties of objects and object + structures. Recently an extension of OCL has been proposed + for the specification of messages sent between objects. In + this paper we present a generalization of this extension + which allows addition- ally to specify causality + constraints. From a pragmatic point of view, such causality + constraints are needed to express, for example, that each + acknowledgment must be preceded by a matching request, + which is frequently required by communication protocols. + Our generalization is based on the introduction of + histories into OCL. Histories describe the external + behavior of objects and groups of objects. Moreover, to + reason compositionally about the behavior of a complex + system we distinguish between local specifications of a + single object and global specifications describing the + interaction between objects. These two types of + specifications are expressed in syntactically difierent + dialects of OCL. Our notion of compositionality, which is + formalized in this paper by a compatibility predicate on + histories, allows the verification of models during the + early stages of a design. } +} + +@InProceedings{ kyas.ea:extended:2004, + author = {Marcel Kyas and Harald Fecher}, + title = {An Extended Type System for {OCL} supporting Templates and + Transformations}, + booktitle = {Formal Methods for Open Object-Based Distributed Systems, + 7th {IFIP} {WG} 6.1 International Conference, {FMOODS} + 2005, Athens, Greece, June 15-17, 2005, Proceedings}, + acknowledgement={none}, + publisher = pub-springer, + address = pub-springer:adr, + year = 2004, + volume = 3535, + editor = {Martin Steffen and Gianluigi Zavattaro}, + isbn = {3-540-26181-8}, + pages = {83--98}, + series = llncs, + doi = {10.1007/11494881_6} +} + +@InProceedings{ giese.ea:simplifying:2005, + author = {Martin Giese and Daniel Larsson}, + title = {Simplifying Transformations of {OCL} Constraints}, + booktitle = {Proceedings, Model Driven Engineering Languages and + Systems (MoDELS) Conference 2005, Montego Bay, Jamaica}, + editor = {Lionel Briand and Clay Williams}, + pages = {309--323}, + volume = 3713, + acknowledgement={none}, + month = oct, + series = lncs, + year = 2005 +} + +@InCollection{ okeefe:improving:2006, + paddress = {Heidelberg}, + address = pub-springer:adr, + author = {Greg O'Keefe}, + booktitle = {{MoDELS} 2006: Model Driven Engineering Languages and + Systems}, + language = {USenglish}, + publisher = pub-springer, + acknowledgement={none}, + series = lncs, + doi = {10.1007/11880240_4}, + number = 4199, + year = 2006, + pages = {42--56}, + editor = {Oscar Nierstrasz and Jon Whittle and David Harel and + Gianna Reggio}, + title = {Improving the Definition of {UML}}, + abstract = {The literature on formal semantics for UML is huge and + growing rapidly. Most contributions open with a brief + remark motivating the work, then quickly move on to the + technical detail. How do we decide whether more rigorous + semantics are needed? Do we currently have an adequate + definition of the syntax? How do we evaluate proposals to + improve the definition? We provide criteria by which these + and other questions can be answered. The growing role of + UML is examined. We compare formal language definition + techniques with those currently used in the definition of + UML. We study this definition for both its content and + form, and conclude that improvements are required. Finally, + we briefly survey the UML formalisation literature, + applying our criteria to determine which of the existing + approaches show the most potential.} +} + +@TechReport{ wasserrab.ea:operational:2005, + author = {Daniel Wasserrab and Tobias Nipkow and Gregor Snelting and + Frank Tip}, + title = {An Operational Semantics and Type Safety Proof for + {\Cpp}-Like Multiple Inheritance}, + institution = {IBM Yorktown Heights}, + number = {RC 23709}, + month = aug, + year = 2005, + abstract = {We present, for the first time, an operational semantics + and a type system for a \Cpp-like object-oriented language + with both shared and repeated multiple inheritance, + together with a machine-checked proof of type safety. The + formalization uncovered several subtle ambiguities in \Cpp, + which \Cpp compilers resolve by ad-hoc means or which even + result in uncontrolled run-time errors. The semantics is + formalized in Isabelle/HOL.}, + acknowledgement={none} +} + +@Unpublished{ crane.ea:class:2006, + author = {Michelle L. Crane and Juergen Dingel and Zinovy Diskin}, + title = {Class Diagrams: Abstract Syntax and Mapping to System + Model}, + note = {Version 1.7.4}, + url = {http://www.cs.queensu.ca/~stl/internal/uml2/documents.htm} + , + pdf = {papers/2006/crane.ea-class-2006.pdf}, + acknowledgement={none}, + year = 2006 +} + +@Article{ henning:rise:2006, + author = {Michi Henning}, + title = {The rise and fall of {CORBA}}, + journal = {Queue}, + volume = 4, + number = 5, + year = 2006, + issn = {1542-7730}, + pages = {28--34}, + doi = {10.1145/1142031.1142044}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none}, + abstract = {Depending on exactly when one starts counting, CORBA is + about 10-15 years old. During its lifetime, CORBA has moved + from being a bleeding-edge technology for early adopters, + to being a popular middleware, to being a niche technology + that exists in relative obscurity. It is instructive to + examine why CORBA---despite once being heralded as the + {\^a}\80\9Cnext-generation technology for e-commerce---suffered + this fate. CORBA{\^a}\80\99s history is one that the computing + industry has seen many times, and it seems likely that + current middleware efforts, specifically Web services, will + reenact a similar history.} +} + +@InProceedings{ briggs.ea:effective:1994, + author = {Preston Briggs and Keith D. Cooper}, + title = {Effective partial redundancy elimination}, + booktitle = PROC # { the ACM SIGPLAN 1994 conference on + Programming language design and implementation}, + year = 1994, + isbn = {0-89791-662-X}, + pages = {159--170}, + location = {Orlando, Florida, United States}, + doi = {10.1145/178243.178257}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none}, + bibkey = {briggs.ea:effective:1994}, + abstract = {Partial redundancy elimination is a code optimization with + a long history of literature and implementation. In + practice, its effectiveness depends on issues of naming and + code shape. This paper shows that a combination of global + reassociation and global value numbering can increase the + effectiveness of partial redundancy elimina- tion. By + imposing a discipline on the choice of names and the shape + of expressions, we are able to expose more redundancies, As + part of the work, we introduce a new algorithm for global + reassociation of expressions. It uses global in- formation + to reorder expressions, creating opportunities for other + optimization. The new algorithm generalizes earlier work + that ordered FORTRAN array address ex- pressions to improve + optimization.} +} + +@Article{ binder:design:1994, + author = {Robert V. Binder}, + title = {Design for testability in object-oriented systems}, + journal = j-cacm, + publisher = pub-acm, + address = pub-acm:adr, + volume = 37, + number = 9, + pages = {87--101}, + month = sep, + year = 1994, + issn = {0001-0782}, + acknowledgement={none}, + keywords = {design; reliability}, + bibkey = {binder:design:1994} +} + +@InProceedings{ bernot.ea:theory:1997, + author = {Gilles Bernot and Laurent Bouaziz and Pascale {Le Gall}}, + title = {A theory of probabilistic functional testing}, + booktitle = PROC # { the 19th international conference on + Software engineering}, + year = 1997, + isbn = {0-89791-914-9}, + pages = {216--226}, + pdf = {papers/1997/p216-bernot.pdf}, + location = {Boston, Massachusetts, United States}, + doi = {10.1145/253228.253273}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none} +} + +@InProceedings{ ntafos:random:1998, + author = {Simeon Ntafos}, + title = {On random and partition testing}, + booktitle = PROC # { ACM SIGSOFT international symposium on + Software testing and analysis}, + year = 1998, + isbn = {0-89791-971-8}, + pages = {42--48}, + location = {Clearwater Beach, Florida, United States}, + doi = {10.1145/271771.271785}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none}, + bibkey = {ntafos:random:1998} +} + +@InProceedings{ hackett.ea:modular:2006, + author = {Brian Hackett and Manuvir Das and Daniel Wang and Zhe + Yang}, + title = {Modular checking for buffer overflows in the large}, + booktitle = {ICSE '06: Proceeding of the 28th international conference + on Software engineering}, + year = 2006, + acknowledgement={none}, + isbn = {1-59593-375-1}, + pages = {232--241}, + location = {Shanghai, China}, + publisher = pub-acm, + address = pub-acm:adr, + doi = {10.1145/1134285.1134319} +} + +@InProceedings{ biere.ea:sat-model-checking:1999, + author = {A. Biere and A. Cimatti and E. M. Clarke and M. Fujita and + Y. Zhu}, + title = {Symbolic model checking using SAT procedures instead of + {BDDs}}, + booktitle = PROC # { the 36th ACM/IEEE conference on Design + automation conference}, + year = 1999, + isbn = {1-58133-109-7}, + pages = {317--320}, + location = {New Orleans, Louisiana, United States}, + doi = {10.1145/309847.309942}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none}, + bibkey = {biere.ea:sat-model-checking:1999} +} + +@Article{ france:problem-oriented:34-10, + author = {Robert France}, + title = {A problem-oriented analysis of basic {UML} static + requirements modeling concepts}, + journal = {ACM SIG-PLAN Notices}, + volume = 34, + number = 10, + pages = {57--69}, + year = 1999, + abstract = {The Unified Modeling Language (UML) is a standard modeling + language in which some of the best object-oriented (OO) + modeling experiences are embedded. In this paper we + illustrate the role formal specification techniques can + play in developing a precise semantics for the UML. We + present a precise characterization of requirements-level + (problem-oriented) Class Diagrams and outline how the + characterization can be used to semantically analyze + requirements Class Diagrams.}, + acknowledgement={none}, + bibkey = {france:problem-oriented:34-10} +} + +@InProceedings{ claessen.ea:quickcheck:2000, + author = {Koen Claessen and John Hughes}, + title = {{QuickCheck}: a lightweight tool for random testing of + {Haskell} programs}, + booktitle = PROC # { the fifth ACM SIGPLAN international + conference on Functional programming}, + year = 2000, + isbn = {1-58113-202-6}, + pages = {268--279}, + doi = {10.1145/351240.351266}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {Quick Check is a tool which aids the Haskell programmer in + formulating and testing properties of programs. Properties + are described as Haskell functions, and can be + automatically tested on random input, but it is also + possible to define custom test data generators. We present + a number of case studies, in which the tool was + successfully used, and also point out some pitfalls to + avoid. Random testing is especially suitable for functional + programs because properties can be stated at a fine grain. + When a function is built from separately tested components, + then random testing suffices to obtain good coverage of the + definition under test. }, + acknowledgement={none}, + bibkey = {claessen.ea:quickcheck:2000} +} + +@Article{ kozen:hoare:2000, + author = {Dexter Kozen}, + title = {On Hoare logic and Kleene algebra with tests}, + journal = {ACM Transactions on Computational Logic}, + volume = 1, + number = 1, + year = 2000, + issn = {1529-3785}, + pages = {60--76}, + doi = {10.1145/343369.343378}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {We show that Kleene algebra with tests (KAT) subsumes + propositional Hoare logic (PHL). Thus the specialized + syntax and deductive apparatus of Hoare logic are + inessential and can be replaced by simple equational + reasoning. In addition, we show that all relationally valid + inference rules are derivable in KAT and that deciding the + relational validity of such rules is PSPACE-complete. }, + acknowledgement={none}, + bibkey = {kozen:hoare:2000} +} + +@InProceedings{ chen.ea:semi-proving:2002, + author = {T. Y. Chen and T. H. Tse and Zhiquan Zhou}, + title = {Semi-proving: an integrated method based on global + symbolic evaluation and metamorphic testing}, + booktitle = PROC # { the international symposium on Software + testing and analysis}, + year = 2002, + isbn = {1-58113-562-9}, + pages = {191--195}, + location = {Roma, Italy}, + doi = {10.1145/566172.566202}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none}, + bibkey = {chen.ea:semi-proving:2002} +} + +@InProceedings{ naumovich.ea:static:2004, + author = {Gleb Naumovich and Paolina Centonze}, + title = {Static Analysis of Role-Based Access Control in {J2EE} + Applications}, + abstract = {This work describes a new technique for analysis of Java + 2, Enterprise Edition (J2EE) applications. In such + applications, Enterprise Java Beans (EJBs) are commonly + used to encapsulate the core computations performed on Web + servers. Access to EJBs is protected by application + servers, according to role-based access control policies + that may be created either at development or deployment + time. These policies may prohibit some types of users from + accessing specific EJB methods. We present a static + technique for analyzing J2EE access control policies with + respect to security-sensitive fields of EJBs and other + server-side objects. Our technique uses points-to analysis + to determine which object fields are accessed by which EJB + methods, directly or indirectly. Based on this information, + J2EE access control policies are analyzed to identify + potential inconsistencies that may lead to security holes. + }, + volume = 29, + number = 5, + month = sep, + year = 2004, + booktitle = {TAV-WEB Proceedings}, + publisher = pub-acm, + address = pub-acm:adr, + acknowledgement={none}, + bibkey = {naumovich.ea:static:2004} +} + +@InProceedings{ altenhofen.ea:high-level:2005, + author = {Michael Altenhofen and Egon B{\"o}rger and Jens Lemcke}, + title = {A High-Level Specification for Mediators(Virtual + Providers)}, + booktitle = {Business Process Management Workshops}, + year = 2005, + pages = {116--129}, + doi = {10.1007/11678564_11}, + crossref = {bussler.ea:business:2006} +} + +@Proceedings{ bussler.ea:business:2006, + editor = {Christoph Bussler and Armin Haller}, + title = {Business Process Management Workshops, BPM 2005 + International Workshops, BPI, BPD, ENEI, BPRM, WSCOBPM, + BPS, Nancy, France, September 5, 2005, Revised Selected + Papers}, + booktitle = {Business Process Management Workshops}, + volume = 3812, + year = 2006, + isbn = {3-540-32595-6} +} + +@Article{ grieskamp.ea:generating:2002, + author = {Wolfgang Grieskamp and Yuri Gurevich and Wolfram Schulte + and Margus Veanes}, + title = {Generating finite state machines from abstract state + machines}, + journal = {SIGSOFT Softw. Eng. Notes}, + volume = 27, + number = 4, + year = 2002, + issn = {0163-5948}, + pages = {112--122}, + doi = {10.1145/566171.566190}, + publisher = pub-acm, + address = pub-acm:adr +} + +@InProceedings{ paulson:isabelle:1988, + author = {Lawrence C. Paulson}, + title = {Isabelle: The Next Seven Hundred Theorem Provers}, + booktitle = {CADE}, + year = 1988, + pages = {772--773}, + crossref = {lusk.ea:cade:1988} +} + +@Proceedings{ lusk.ea:cade:1988, + editor = {Ewing L. Lusk and Ross A. Overbeek}, + title = {9th International Conference on Automated Deduction, + Argonne, Illinois, USA, May 23-26, 1988, Proceedings}, + booktitle = {CADE}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 310, + year = 1988, + isbn = {3-540-19343-X} +} + +@Article{ ntafos:comparison:1988, + author = {S. C. Ntafos}, + title = {A Comparison of Some Structural Testing Strategies}, + journal = j-ieee-tse, + volume = 14, + number = 6, + year = 1988, + pdf = {papers/1988/e0868.pdf}, + issn = {0098-5589}, + pages = {868--874}, + doi = {http://csdl.computer.org/comp/trans/ts/1988/06/e0868abs.htm} + , + publisher = pub-ieee, + address = pub-ieee:adr, + abstract = {Several structural testing strategies are compared in + terms of their relative coverage of the program's structure + and also in terms of the number of test cases needed to + satisfy each strategy. Some of the deficiencies of such + comparisons are discussed}, + acknowledgement={none} +} + +@InProceedings{ lange.ea:flyspeck:2008, + author = {Christoph Lange and Sean McLaughlin and Florian Rabe}, + title = {Flyspeck in a Semantic {Wiki}}, + booktitle = {SemWiki}, + year = 2008, + url = {http://ceur-ws.org/Vol-360/paper-21.pdf}, + crossref = {lange.ea:semwiki:2008} +} + +@Proceedings{ lange.ea:semwiki:2008, + editor = {Christoph Lange and Sebastian Schaffert and Hala + Skaf-Molli and Max V{\"o}lkel}, + title = PROC # {the 3rd Semantic Wiki Workshop (SemWiki + 2008) at the 5th European Semantic Web Conference (ESWC + 2008), Tenerife, Spain, June 2nd, 2008}, + booktitle = {SemWiki}, + publisher = {CEUR-WS.org}, + series = {CEUR Workshop Proceedings}, + volume = 360, + year = 2008 +} + +@Book{ matouvsek.ea:invitation:2008, + author = {Ji{\v}r{\'\i} Matou{\v}sek, and Jaroslav Ne{\v}set{\v}ril}, + title = {Invitation to discrete mathematics.}, + language = {English}, + edition = {2nd}, + publisher = pub-oxford, + address = pub-oxford:adr, + pages = 443, + year = 2008, + isbn = {978-0198570431}, + abstract = {This is the second edition of a delightful textbook, see + [Invitation to discrete mathematics. (Oxford): Clarendon + Press. (1998; Zbl 0901.05001)]. Besides the usual + elimination of a few typos there are some additions, namely + a chapter on partially ordered sets, a section on + Tur\'{a}n's theorem on the number of edges in a + triangle-free graph and a chapter on Ramsey's theorem.\par + New to the second edition are also several proofs of the + Cauchy-Schwarz inequality, a very attractive elegant new + proof of Cayley's theorem on the number of labeled trees on + $n$ vertices via PARTs (Plans of Assembly of a Rooted + Tree), which the authors attribute to Jim Pitman, and + another proof of the determinant formula for counting + spanning trees of a given graph. The newly added geometric + interpretation of the construction of the real projective + plane is aided by the beautiful artistic rendering in the + figure with the caption ``The real projective plane in + moonlight''.}, + keywords = {discrete mathematics; problem solving; counting + techniques; graph theory; trees; algorithms; planar graphs; + Sperner's theorem; finite projective planes; probabilistic + method; generating functions; partially ordered sets; Turan + theorem; Ramsey theorem; Cauchy Schwartz inequality; Cayley + theorem; geometric interpretation; real projective plane} +} + +@Book{ syme.ea:expert-f:2007, + author = {Don Syme and Adam Granicz and Antonio Cisternino}, + title = {Expert F\# (Expert's Voice in {.Net})}, + isbn = 9781590598504, + publisher = {Apress}, + pages = 609, + year = 2007 +} + +@Book{ baier.ea:principles:2008, + abstract = {Our growing dependence on increasingly complex computer + and software systems necessitates the development of + formalisms, techniques, and tools for assessing functional + properties of these systems. One such technique that has + emerged in the last twenty years is model checking, which + systematically (and automatically) checks whether a model + of a given system satisfies a desired property such as + deadlock freedom, invariants, or request-response + properties. This automated technique for verification and + debugging has developed into a mature and widely used + approach with many applications. \_Principles of Model + Checking\_ offers a comprehensive introduction to model + checking that is not only a text suitable for classroom use + but also a valuable reference for researchers and + practitioners in the field. + + The book begins with the basic principles for modeling + concurrent and communicating systems, introduces different + classes of properties (including safety and liveness), + presents the notion of fairness, and provides automata- + based algorithms for these properties. It introduces the + temporal logics LTL and CTL, compares them, and covers + algorithms for verifying these logics, discussing real-time + systems as well as systems subject to random phenomena. + Separate chapters treat such efficiency-improving + techniques as abstraction and symbolic manipulation. The + book includes an extensive set of examples (most of which + run through several chapters) and a complete set of basic + results accompanied by detailed proofs. Each chapter + concludes with a summary, bibliographic notes, and an + extensive list of exercises of both practical and + theoretical nature.}, + author = {Christel Baier and Joost-Pieter Katoen}, + howpublished = {Hardcover}, + isbn = {026202649X}, + month = may, + publisher = pub-mit, + address = pub-mit:adr, + title = {Principles of Model Checking}, + year = 2008 +} + +@Book{ bertot.ea:interactive:2004, + author = {Yves Bertot and Pierre Cast{\'e}ran}, + keywords = {theorem-proving, type-theory, verification}, + title = {Interactive Theorem Proving and Program Development. + Coq'Art: The Calculus of Inductive Constructions}, + pages = 500, + publisher = pub-springer, + address = pub-springer:adr, + isbn = {978-3540208549}, + year = 2004, + abstract = {Coq is an interactive proof assistant for the development + of mathematical theories and formally certified software. + It is based on a theory called the calculus of inductive + constructions, a variant of type theory. This book provides + a pragmatic introduction to the development of proofs and + certified programs using Coq. With its large collection of + examples and exercises it is an invaluable tool for + researchers, students, and engineers interested in formal + methods and the development of zero-fault software.} +} + +@Article{ korel:automated:1990, + author = {Bogdan Korel}, + title = {Automated Software Test Data Generation}, + journal = j-ieee-tse, + volume = 16, + number = 8, + year = 1990, + issn = {0098-5589}, + pages = {870--879}, + doi = {10.1109/32.57624}, + publisher = pub-ieee, + address = pub-ieee:adr, + acknowledgement={none}, + bibkey = {korel:automated:1990} +} + +@Article{ hamlet.ea:partition:1990, + author = {Dick Hamlet and Ross Taylor}, + title = {Partition Testing Does Not Inspire Confidence (Program + Testing)}, + journal = j-ieee-tse, + volume = 16, + number = 12, + year = 1990, + issn = {0098-5589}, + pages = {1402--1411}, + doi = {10.1109/32.62448}, + publisher = pub-ieee, + address = pub-ieee:adr, + acknowledgement={none}, + bibkey = {hamlet.ea:partition:1990} +} + +@TechReport{ sharangpani.ea:statistical:1994, + language = {USenglish}, + author = {H. P. Sharangpani and Ph. D. M. I. Barton }, + title = {Statistical Analysis of Floating Point Flaw in the + Pentium$^{TM}$ Processor}, + institution = {Intel Corporation}, + month = nov, + year = 1994, + keywords = {pentium; flaw; FDIV; bug}, + url = {http://www.intel.com/support/processors/pentium/fdiv/wp/}, + abstract = {A subtle flaw in the hardware divide unit of the Pentium + TM Processor was discovered by Intel. Subsequently, a + characterization of its impact to the end-user application + base was conducted. The flaw is rare and data-dependent, + and causes a reduction in precision of the divide + instruc-tion and certain other operations in certain cases. + The significance of the flaw depends upon (a) the rate of + use of specific FP instructions in the Pentium TM CPU, (b) + the data fed to them, (c) the way in which the results of + these instructions are propagated into further computation + in the application; and (d) the way in which the final + results of the application are interpreted. The thorough + and detailed characterization of the flaw and the + subsequent investigations of its impact on applications + through elaborate surveys, analyses and empirical + observation lead us to the overall conclusion that the flaw + is of no concern to the vast majority of users of Pentium + processor based systems. A few users of applications in the + scientific/engineering and financial engineering fields who + require unusual precision and invoke millions of divides + per day may need t o employ either an updated Pen t ium + processor without the flaw or a software workaround.}, + acknowledgement={none}, + bibkey = {sharangpani.ea:statistical:1994} +} + +@InProceedings{ paulson:formulation:1988, + author = {Lawrence C. Paulson}, + title = {A formulation of the simple theory of types (for + Isabelle).}, + booktitle = {Conference on Computer Logic}, + year = 1988, + acknowledgement={none}, + pages = {246--274}, + doi = {10.1007/3-540-52335-9_58}, + crossref = {martin-lof.ea:international:1990} +} + +@Proceedings{ martin-lof.ea:international:1990, + editor = {Per Martin-L{\"o}f and Grigori Mints}, + title = {International Conference on Computer Logic, Tallinn, USSR, + December 1988, Proceedings}, + booktitle = {Conference on Computer Logic}, + publisher = pub-springer, + series = lncs, + address = pub-springer:adr, + acknowledgement={none}, + volume = 417, + year = 1990, + isbn = {3-540-52335-9} +} + +@Article{ bernot.ea:software:1991, + author = {Gilles Bernot and Marie Claude Gaudel and Bruno Marre}, + title = {Software testing based on formal specifications: a theory + and a tool}, + journal = {Softw. Eng. J.}, + volume = 6, + number = 6, + year = 1991, + issn = {0268-6961}, + pages = {387--405}, + publisher = {Michael Faraday House}, + address = {Herts, UK, UK} +} + +@Article{ chadwick.ea:permis:2008, + author = {David Chadwick and Gansen Zhao and Sassa Otenko and Romain + Laborde and Linying Su and Tuan Anh Nguyen}, + title = {{PERMIS}: a modular authorization infrastructure}, + journal = {Concurrency and Computation: Practice \& Experience}, + volume = 20, + number = 11, + year = 2008, + issn = {1532-0626}, + pages = {1341--1357}, + doi = {10.1002/cpe.v20:11}, + publisher = pub-wiley, + address = pub-wiley:adr, + abstract = {Authorization infrastructures manage privileges and render + access control decisions, allowing applications to adjust + their behavior according to the privileges allocated to + users. This paper describes the PERMIS role-based + authorization infrastructure along with its conceptual + authorization, access control, and trust models. PERMIS has + the novel concept of a credential validation service, which + verifies a user's credentials prior to access control + decision-making and enables the distributed management of + credentials. PERMIS also supports delegation of authority; + thus, credentials can be delegated between users, further + decentralizing credential management. Finally, PERMIS + supports history-based decision-making, which can be used + to enforce such aspects as separation of duties and + cumulative use of resources. Details of the design and the + implementation of PERMIS are presented along with details + of its integration with Globus Toolkit, Shibboleth, and + GridShib. A comparison of PERMIS with other authorization + and access control implementations is given, along with + suggestions where future research and development are still + needed.} +} + +@Article{ altenhofen.ea:asms:2008, + author = {Michael Altenhofen and Andreas Friesen and Jens Lemcke}, + title = {{ASMs} in Service Oriented Architectures}, + journal = j-ucs, + volume = 14, + number = 12, + year = 2008, + pages = {2034--2058}, + abstract = {We give a survey on work we did in the past where we have + successfully applied the ASM methodology to provide + abstract models for a number of problem areas that are + commonly found in Service Oriented Architectures (SOA). In + particular, we summarize our work on (1) service behavior + mediation, (2) service discovery, and (3) service + composition, showing that the corresponding solutions can + be described as variations of a fundamental abstract + processing model{\^a}the Virtual Provider.}, + keywords = {process mediation, service discovery, workflow composition + } +} + +@Book{ borger.ea:abstract:2003, + author = {Egon B{\"o}rger and Robert F. St{\"a}rk}, + title = {Abstract State Machines: A Method for High-Level System + Design and Analysis}, + publisher = pub-springer, + address = pub-springer:adr, + isbn = {3-540-00702-4}, + abstract = {This book introduces into a new software engineering + method which guides the development of systems seamlessly + from requirements capture to coding. The method bridges the + gap between understanding and formulating real-world + problems by humans and the deployment of their solutions by + code-executing machines on changing platforms.It covers + design and analysis for both hardware and software + systems.It has a scientific foundation and improves current + industrial practice by linking the descriptions at the + successive stages of the system development cycle in a + coherent conceptual framework to keep the system models at + related levels synchronized.The method supports the + integration of standard design, analysis and documentation + techniques for model reuse (by abstraction), validation (by + simulation and high-level testing), verification (by + reasoning) and maintenance (by structured documentation). + }, + year = 2003 +} + +@InProceedings{ altenhofen.ea:concurrent:2009, + author = {Michael Altenhofen and Egon B{\"o}rger}, + title = {Concurrent Abstract State Machines and {$^+\mathit{CAL}$} + Programs}, + pages = {1--17}, + doi = {10.1007/978-3-642-03429-9_1}, + abstract = {We apply the ASM semantics framework to define the await + construct in the context of concurrent ASMs. We link + {$^+\mathit{CAL}$} to concurrent control state ASMs with + turbo ASM submachines.}, + booktitle = {Recent Trends in Algebraic Development Techniques}, + series = s-lncs, + publisher = pub-springer, + address = pub-springer:adr, + year = 2009 +} + +@Article{ farahbod.ea:coreasm:2007, + author = {Roozbeh Farahbod and Vincenzo Gervasi and Uwe Gl{\"a}sser}, + title = {{CoreASM}: An Extensible {ASM} Execution Engine}, + journal = {Fundamenta Informaticae}, + publisher = {IOS Press}, + volume = 77, + number = {1-2}, + year = 2007, + issn = {0169-2968}, + pages = {71--103}, + ee = {http://iospress.metapress.com/openurl.asp?genre=article{\&}issn=0169-2968{\&}volume=77{\&}issue=1{\&}spage=71} + , + abstract = {In this paper we introduce a new research effort in making + abstract state machines (ASMs) executable. The aim is to + specify and implement an execution engine for a language + that is as close as possible to the mathematical definition + of pure ASMs. The paper presents the general architecture + of the engine, together with a high-level description of + the extensibility mechanisms that are used by the engine to + accommodate arbitrary backgrounds, scheduling policies, and + new rule forms.}, + keywords = {CoreASM, Abstract state machines, Specification languages, + Executable specification} +} + +@Article{ knuth:literate:1984, + author = {Donald E. Knuth}, + title = {Literate Programming}, + journal = {The Computer Journal}, + volume = 27, + number = 2, + year = 1984, + pages = {97--111}, + doi = {10.1093/comjnl/27.2.97}, + publisher = {Oxford University Press}, + address = {Oxford, UK}, + issn = {0010-4620}, + abstract = {The author and his associates have been experimenting for + the past several years with a programming language and + documentation system called WEB. This paper presents WEB by + example, and discusses why the new system appears to be an + improvement over previous ones.} +} + +@Article{ forgy:rete:1982, + author = {Charles L. Forgy}, + title = {Rete: A Fast Algorithm for the Many Patterns/Many Objects + Match Problem}, + journal = {Artificial Intelligence}, + volume = 19, + number = 1, + year = 1982, + pages = {17--37}, + issn = {0004-3702}, + abstract = {The Rete Match Algorithm is an efficient method for + comparing a large collection of patterns to a large + collection of objects. It finds all the objects that match + each pattern. The algorithm was developed for use in + production system interpreters, and it has been used for + systems containing from a few hundred to more than a + thousand patterns and objects. This article presents the + algorithm in detail. It explains the basic concepts of the + algorithm, it describes pattern and object representations + that are appropriate for the algorithm, and it describes + the operations performed by the pattern matcher. }, + doi = {10.1016/0004-3702(82)90020-0} +} + +@InProceedings{ burrows:chubby:2006, + author = {Mike Burrows}, + title = {The Chubby lock service for loosely-coupled distributed + systems}, + booktitle = {OSDI '06: Proceedings of the 7th symposium on Operating + systems design and implementation}, + year = 2006, + isbn = {1-931971-47-1}, + pages = {335--350}, + publisher = {\acs{usenix} Association}, + location = {Seattle, Washington}, + address = {Berkeley, CA, USA}, + abstract = {We describe our experiences with the Chubby lock service, + which is intended to provide coarse-grained locking as well + as reliable (though low-volume) storage for a + loosely-coupled distributed system. Chubby provides an + interface much like a distributed file system with advisory + locks, but the design emphasis is on availability and + reliability, as opposed to high performance. Many instances + of the service have been used for over a year, with several + of them each handling a few tens of thousands of clients + concurrently. The paper describes the initial design and + expected use, compares it with actual use, and explains how + the design had to be modified to accommodate the + differences.} +} + +@InProceedings{ mccarthy.ea:architecture:1989, + author = {Dennis McCarthy and Umeshwar Dayal}, + title = {The architecture of an active database management system}, + booktitle = {SIGMOD '89: Proceedings of the 1989 ACM SIGMOD + international conference on Management of data}, + year = 1989, + isbn = {0-89791-317-5}, + pages = {215--224}, + location = {Portland, Oregon, United States}, + doi = {10.1145/67544.66946}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {he HiPAC project is investigating active, time-constrained + database management. An active DBMS is one which + automatically executes specified actions when specified + conditions arise. HiPAC has proposed Event-Condition-Action + (ECA) rules as a formalism for active database + capabilities. We have also developed an execution model + that specifies how these rules are processed in the context + of database transactions. The additional functionality + provided by ECA rules makes new demands on the design of an + active DBMS. In this paper we propose an architecture for + an active DBMS that supports ECA rules. This architecture + provides new forms of interaction, in support of ECA rules, + between application programs and the DBMS. This leads to a + new paradigm for constructing database applications.} +} + +@InProceedings{ carioni.ea:scenario-based:2008, + author = {Alessandro Carioni and Angelo Gargantini and Elvinia + Riccobene and Patrizia Scandurra}, + title = {A Scenario-Based Validation Language for {ASMs}}, + booktitle = {ABZ}, + year = 2008, + pages = {71--84}, + doi = {10.1007/978-3-540-87603-8_7}, + abstract = {This paper presents the AValLa language, a domain-specific + modelling language for scenario-based validation of ASM + models, and its supporting tool, the AsmetaV validator. + They have been developed according to the model-driven + development principles as part of the asmeta (ASM + mETAmodelling) toolset, a set of tools around ASMs. As a + proof-of-concepts, the paper reports the results of the + scenario-based validation for the well-known LIFT control + case study.}, + crossref = {borger.ea:abstract:2008} +} + +@Proceedings{ borger.ea:abstract:2008, + editor = {Egon B{\"o}rger and Michael J. Butler and Jonathan P. + Bowen and Paul Boca}, + title = {Abstract State Machines, B and Z, First International + Conference, ABZ 2008, London, UK, September 16-18, 2008. + Proceedings}, + booktitle = {ABZ}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 5238, + year = 2008, + isbn = {978-3-540-87602-1} +} + +@Article{ decandia.ea:dynamo:2007, + author = {Giuseppe DeCandia and Deniz Hastorun and Madan Jampani and + Gunavardhan Kakulapati and Avinash Lakshman and Alex + Pilchin and Swaminathan Sivasubramanian and Peter Vosshall + and Werner Vogels}, + title = {{Dynamo}: {Amazon's} highly available key-value store}, + journal = {ACM SIGOPS Operating Systems Review}, + volume = 41, + number = 6, + year = 2007, + issn = {0163-5980}, + pages = {205--220}, + doi = {10.1145/1323293.1294281}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {Reliability at massive scale is one of the biggest + challenges we face at Amazon.com, one of the largest + e-commerce operations in the world; even the slightest + outage has significant financial consequences and impacts + customer trust. The Amazon.com platform, which provides + services for many web sites worldwide, is implemented on + top of an infrastructure of tens of thousands of servers + and network components located in many datacenters around + the world. At this scale, small and large components fail + continuously and the way persistent state is managed in the + face of these failures drives the reliability and + scalability of the software systems. + + This paper presents the design and implementation of + Dynamo, a highly available key-value storage system that + some of Amazon's core services use to provide an + "always-on" experience. To achieve this level of + availability, Dynamo sacrifices consistency under certain + failure scenarios. It makes extensive use of object + versioning and application-assisted conflict resolution in + a manner that provides a novel interface for developers to + use.} +} + +@InProceedings{ crampton.ea:secondary:2006, + author = {Jason Crampton and Wing Leung and Konstantin Beznosov}, + title = {The secondary and approximate authorization model and its + application to {Bell-LaPadula} policies}, + booktitle = {SACMAT '06: Proceedings of the eleventh ACM symposium on + Access control models and technologies}, + year = 2006, + isbn = {1-59593-353-0}, + pages = {111--120}, + location = {Lake Tahoe, California, USA}, + doi = {10.1145/1133058.1133075}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {We introduce the concept, model, and policy-specific + algorithms for inferring new access control decisions from + previous ones. Our secondary and approximate authorization + model (SAAM) defines the notions of primary vs. secondary + and precise vs. approximate authorizations. Approximate + authorization responses are inferred from cached primary + responses, and therefore provide an alternative source of + access control decisions in the event that the + authorization server is unavailable or slow. The ability to + compute approximate authorizations improves the reliability + and performance of access control sub-systems and + ultimately the application systems themselves.The operation + of a system that employs SAAM depends on the type of access + control policy it implements. We propose and analyze + algorithms for computing secondary authorizations in the + case of policies based on the Bell-LaPadula model. In this + context, we define a dominance graph, and describe its + construction and usage for generating secondary responses + to authorization requests. Preliminary results of + evaluating SAAM BLP algorithms demonstrate a 30\% increase + in the number of authorization requests that can be served + without consulting access control policies.} +} + +@InProceedings{ turkmen.ea:performance:2008, + author = {Fatih Turkmen and Bruno Crispo}, + title = {Performance evaluation of {XACML} {PDP} implementations}, + booktitle = {SWS '08: Proceedings of the 2008 ACM workshop on Secure + web services}, + year = 2008, + isbn = {978-1-60558-292-4}, + pages = {37--44}, + location = {Alexandria, Virginia, USA}, + doi = {10.1145/1456492.1456499}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {eXtensible Access Control Markup Language (XACML), an + OASIS standard, is the most widely used policy specifica- + tion language for access control. Its simplicity in syntax + and strength in coverage makes it suitable for diverse en- + vironments such as Service Oriented Architectures (SOAs) + and P2P systems. There are different implementations of + XACML available. Some of these implementations are open + source and some others are proprietary. In this work we + intended to shed some lights to the performance issues of + XACML engines. We tested 3 open source XACML + implementations with different policy/request settings. Our + experiments revealed some important points to be taken into + consideration when deploying an XACML based access control + system. Besides, our results can be used as hints by policy + writers and system developers for deploying efficient + authorization services.} +} + +@InProceedings{ chen.ea:constraint:2006, + author = {Hong Chen and Ninghui Li}, + title = {Constraint generation for separation of duty}, + booktitle = {SACMAT '06: Proceedings of the eleventh ACM symposium on + Access control models and technologies}, + year = 2006, + isbn = {1-59593-353-0}, + pages = {130--138}, + location = {Lake Tahoe, California, USA}, + doi = {10.1145/1133058.1133077}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {Separation of Duty (SoD) is widely recognized to be a + fundamental principle in computer security. A Static SoD + (SSoD) policy states that in order to have all permissions + necessary to complete a sensitive task, the cooperation of + at least a certain number of users is required. In + Role-Based Access Control (RBAC), Statically Mutually + Exclusive Role (SMER) constraints are used to enforce SSoD + policies. This paper studies the problem of generating sets + of constraints that (a) enforce a set of SSoD policies, (b) + are compatible with the existing role hierarchy, and (c) + are minimal in the sense that there is no other constraint + set that is less restrictive and satisfies (a) and (b).} +} + +@InProceedings{ schaad.ea:case:2005, + author = {Andreas Schaad and Pascal Spadone and Helmut Weichsel}, + title = {A case study of separation of duty properties in the + context of the Austrian {``eLaw''} process.}, + booktitle = {SAC '05: Proceedings of the 2005 ACM symposium on Applied + computing}, + year = 2005, + isbn = {1-58113-964-0}, + pages = {1328--1332}, + location = {Santa Fe, New Mexico}, + doi = {10.1145/1066677.1066976}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {Over the last few years rapid progress has been made in + moving from conceptual studies, "whitepapers" and + initiatives to the actual deployment of e-Government + systems [13]. In this paper we present the case study of an + existing e-Government system (eLaw) which already supports + key legislative processes in the country of Austria1. The + study has been performed in the context of the EU FP6 + project "eJustice".We present a detailed system and + workflow representation referring to the example process of + changing a federal law in Austria. Since such processes and + their results, i.e. the laws of a country, have an enormous + impact on society, they need to be secured against external + and internal alteration, be it inadvertent or malicious. + This is even more important in the electronic world.Instead + of discussing the obvious security requirements like virus + protection or network-level access control, our focus is on + an often neglected form of organisational security and + control properties called separation of duties. We will + analyse and discuss a set of these in terms of the + described eLaw process.} +} + +@InProceedings{ evered.ea:case:2004, + author = {Mark Evered and Serge B{\"o}geholz}, + title = {A case study in access control requirements for a Health + Information System}, + booktitle = {ACSW Frontiers '04: Proceedings of the second workshop on + Australasian information security, Data Mining and Web + Intelligence, and Software Internationalisation}, + year = 2004, + pages = {53--61}, + location = {Dunedin, New Zealand}, + publisher = {Australian Computer Society, Inc.}, + address = {Darlinghurst, Australia, Australia}, + abstract = {We present a detailed examination of the access + constraints for a small real-world Health Information + System with the aim of achieving minimal access rights for + each of the involved principals. We show that, even for + such a relatively simple system, the resulting constraints + are very complex and cannot be expressed easily or clearly + using the static per-method access control lists generally + supported by component-based software. We derive general + requirements for the expressiveness of access constraints + and propose criteria for a more suitable access control + mechanism in the context of component-based systems. We + describe a two-level mechanism which can fulfil these + criteria.} +} + +@Article{ kapsalis.ea:dynamic:2006, + title = {A dynamic context-aware access control architecture for + e-services}, + journal = {Computers \& Security}, + volume = 25, + number = 7, + pages = {507--521}, + year = 2006, + issn = {0167-4048}, + doi = {10.1016/j.cose.2006.05.004}, + author = {Vassilis Kapsalis and Loukas Hadellis and Dimitris Karelis + and Stavros Koubias}, + keywords = {e-Services, Access control, Web services, Context-aware, + Authorization, UML}, + abstract = { The universal adoption of the Internet and the emerging + web services technologies constitutes the infrastructure + that enables the provision of a new generation of + e-services and applications. However, the provision of + e-services through the Internet imposes increased risks, + since it exposes data and sensitive information outside the + client premises. Thus, an advanced security mechanism has + to be incorporated, in order to protect this information + against unauthorized access. In this paper, we present a + context-aware access control architecture, in order to + support fine-grained authorizations for the provision of + e-services, based on an end-to-end web services + infrastructure. Access permissions to distributed web + services are controlled through an intermediary server, in + a completely transparent way to both clients and protected + resources. The access control mechanism is based on a + Role-Based Access Control (RBAC) model, which incorporates + dynamic context information, in the form of context + constraints. Context is dynamically updated and provides a + high level of abstraction of the physical environment by + using the concepts of simple and composite context + conditions. Also, the paper deals with implementation + issues and presents a system that incorporates the proposed + access control mechanism in a web services infrastructure + that conform to the OPC XML-DA specification.} +} + +@InProceedings{ anderson:comparison:2006, + author = {Anne H. Anderson}, + title = {A comparison of two privacy policy languages: {EPAL} and + {XACML}}, + booktitle = {SWS '06: Proceedings of the 3rd ACM workshop on Secure web + services}, + year = 2006, + isbn = {1-59593-546-0}, + pages = {53--60}, + location = {Alexandria, Virginia, USA}, + doi = {10.1145/1180367.1180378}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {Current regulatory requirements in the U.S. and other + countries make it increasingly important for Web Services + to be able to enforce and verify their compliance with + privacy policies. Structured policy languages can play a + major role by supporting automated enforcement of policies + and auditing of access decisions. This paper compares two + policy languages that have been developed for use in + expressing directly enforceable privacy policies -- the + Enterprise Privacy Authorization Language (EPAL) and the + OASIS Standard eXtensible Access Control Markup Language + (XACML), together with its standard privacy profile.} +} + +@Book{ rankl.ea:smart-card:2003, + author = {Wolfgang Rankl and Wolfgang Effing}, + title = {Smart Card Handbook}, + year = 2003, + isbn = 9780470856680, + doi = {10.1002/047085670X}, + publisher = pub-wiley, + address = pub-wiley:adr, + abstract = {The boom in smart card technology reflects the medium's + broad solutions potential. Embedded with a sophisticated + microprocessor, smart cards offer unparalleled memory + capacity and data encryption capability. From providing + secure access to the Internet and mobile radio networks to + performing security-sensitive financial transactions in + credit card schemes, the Electronic Purse and Pay TV + systems, smart card technology is now a multi-billion + dollar industry. The Smart Card Handbook presents a + state-of-the-art overview of the technology from + fundamental information processing through design, + manufacture and operation of smart card schemes. Written in + a highly accessible style the Smart Card Handbook meets the + needs of both novice and expert. This is an essential + reference for computer and electronics engineers and + students in microchip design and security system + development. For professionals developing smart card + products, this unique reference will provide an invaluable + insight to all the facets of this sophisticated + technology.} +} + +@Book{ graham.ea:concrete:1989, + title = {Concrete Mathematics}, + author = {Roland L. Graham and Donald E. Knuth and Oren Patashnik}, + isbn = {0-201-14236-8}, + publisher = pub-aw, + address = pub-aw:adr, + pages = 578, + year = 1989 +} + +@Manual{ iso:ansi-cpp:1998, + bibkey = {iso:ansi-cpp:1998}, + abstract = {Specifies requirements for implementations of the C++ + programming language. This International Standard also + defines C++. Other requirements and relaxations of the + first requirement appear at various places within this + standard.}, + note = {Doc. No. ISO/IEC 14882-1998}, + title = {International Standard: Programming languages -C++ }, + organization = {ANSI/ISO}, + year = 1998, + month = sep, + publisher = {The American National Standards Institute}, + address = {New York} +} + +@Book{ reid:thinking:1990, + abstract = {The book is a result of Glenn Reid's years trying to teach + people to write PostScript programs, during which he + discovered that people tended to try to make PostScript + "look like" other programming languages they already knew. + There is even a chapter in this book entitled "PostScript + Is Not Like C", because it is really a very different + language, and one must learn to "think" in PostScript in + order to be a good programmer. }, + author = {Glenn C. Reid}, + title = {Thinking in Postscript}, + publisher = pub-aw, + address = pub-aw:adr, + month = sep, + year = 1990, + language = {USenglish}, + keywords = {Postscript}, + public = {yes}, + isbn = {0-201-52372-8}, + url = {http://www.rightbrain.com/download/books/ThinkingInPostScript.pdf} + +} + +@Book{ knuth:seminumerical:1981, + author = {Donald E. Knuth}, + series = {The Art of Computer Programming}, + volume = 2, + title = {Seminumerical Algorithms}, + edition = {second}, + isbn = {0-201-03822-6}, + year = 1981, + publisher = pub-aw, + address = pub-aw:adr +} + +@Book{ wegener:complexity:1987, + address = {Stuttgart}, + author = {Ingo Wegener}, + language = {USenglish}, + public = {yes}, + publisher = {John Wiley \& Sons Ltd., and B.G. Teubner}, + title = {The Complexity of Boolean Functions}, + url = {\url{http://ls2-www.informatik.uni-dortmund.de/monographs/bluebook/}} + , + year = 1987, + cover = {1987/wegener:complexity:1987.png}, + timestamp = 948019205 +} + +@Article{ stallman:societal:1997, + author = {Richard Stallman}, + title = {Societal Dimensions: The Right to Read}, + journal = j-cacm, + volume = 40, + number = 2, + pages = {85--87}, + month = feb, + year = 1997, + coden = {CACMA2}, + issn = {0001-0782}, + url = {http://www.acm.org/pubs/citations/journals/cacm/1997-40-2/p85-stallman/} + , + localurl = {papers/1997/p85-stallman.pdf}, + note = {\url{http://www.gnu.org/}}, + acknowledgement=ack-nhfb, + classification= {C0230 (Economic, social and political aspects of + computing); C0310D (Computer installation management)}, + keywords = {Clinton administration; Clipper chip; computer crime; + design; industrial property; key-escrow proposals; legal + aspects; management; pirates; right to read; security; + Software Protection Authority; Software Publisher's + Association}, + subject = {{\bf K.1} Computing Milieux, THE COMPUTER INDUSTRY. {\bf + K.5.0} Computing Milieux, LEGAL ASPECTS OF COMPUTING, + General. {\bf K.4.0} Computing Milieux, COMPUTERS AND + SOCIETY, General.}, + treatment = {G General Review} +} + +@Article{ stallman:societal:1997-b, + author = {Richard Stallman}, + title = {Societal Dimensions: The Right to Read}, + journal = j-cacm, + volume = 40, + number = 2, + pages = {85--87}, + month = feb, + year = 1997, + issn = {0001-0782}, + classification= {C0230 (Economic, social and political aspects of + computing); C0310D (Computer installation management)}, + keywords = {Clinton administration; Clipper chip; computer crime; + design; industrial property; key-escrow proposals; legal + aspects; management; pirates; right to read; security; + Software Protection Authority; Software Publisher's + Association}, + subject = {{\bf K.1} Computing Milieux, THE COMPUTER INDUSTRY. {\bf + K.5.0} Computing Milieux, LEGAL ASPECTS OF COMPUTING, + General. {\bf K.4.0} Computing Milieux, COMPUTERS AND + SOCIETY, General.}, + treatment = {G General Review} +} + +@Article{ dalton.ea:securing:2001, + bibkey = {dalton.ea:securing:2001}, + author = {Chris Dalton and Tse Huong Choo}, + title = {An operating system approach to securing e--services}, + journal = j-cacm, + volume = 44, + number = 2, + pages = {58--64}, + month = feb, + abstract = {This article looks at some of the problems surrounding + application compromise in more detail and puts forward our + approach to solving these problems. We do not attempt to + guarantee that the application services are bug-free (a + difficult problem). Instead, we have found that the effects + of this type of attack, and quite a few others, can be + usefully mitigated by adding specific properties to the OSs + used to host those applications. + + Specifically, we look at Trusted Linux, HP Laboratories' + implementation of a secure version of Linux, which we + believe is an ideal platform for e-service application + hosting.}, + year = 2001, + issn = {0001-0782} +} + +@Article{ cornea-hasegan:proving:1998, + language = {USenglish}, + abstract = {The work presented in this paper was initiated as part of + a study on software alternatives to the hardware + implementations of floating-point operations such as divide + and square root. The results of the study proved the + viability of software implementations, and showed that + certain proposed algorithms are comparable in performance + to current hardware implementations. This paper discusses + two components of that study: + + (1) A methodology for proving the IEEE correctness of the + result of iterative algorithms that implement the + floating-point square root, divide, or remainder operation. + (2) Identification of operands for the floating-point + divide and square root operations that lead to results + representing difficult cases for IEEE rounding. + + Some general properties of floating-point computations are + presented first. The IEEE correctness of the floating-point + square root operation is discussed next. We show how + operands for the floating-point square root that lead to + difficult cases for rounding can be generated, and how to + use this knowledge in proving the IEEE correctness of the + result of iterative algorithms that calculate the square + root of a floating-point number. Similar aspects are + analyzed for the floating-point divide operation, and we + present a method for generating difficult cases for + rounding. In the case of the floating-point divide + operation, however, it is more difficult to use this + information in proving the IEEE correctness of the result + of an iterative algorithm than it is for the floating-point + square root operation. We examine the restrictions on the + method used for square root. Finally, we present possible + limitations due to the finite exponent range.}, + journal = {Intel Technology Journal}, + volume = {Q2}, + year = 1998, + title = {Proving the {IEEE} Correctness of Iterative Floating-Point + Square Root, Divide, and Remainder Algorithms }, + author = {Marius Cornea-Hasegan}, + keywords = {floating-point, IEEE correctness, divide, square root, + remainder}, + url = {\url{http://developer.intel.com/technology/itj/q21998/articles/art_3.htm}} + +} + +@Article{ edelman:mathematics:1997, + author = {Alan Edelman}, + abstract = {Despite all of the publicity surrounding the Pentium bug + of 1994, the mathematical details of the bug are poorly + understood. We discuss these details and supply a new proof + of the Coe--Tang result that the at-risk divisors have six + consecutive ones in positions 5 through 10. Also, we prove + that the worst-case absolute error for arguments in [1,2) + is on the order of 1e--5. }, + journal = {SIAM}, + title = {The Mathematics of the Pentium Division Bug}, + keywords = {Pentium, SRT division, floating point operations }, + year = 1997, + url = {\url{http://epubs.siam.org/sam-bin/dbq/article/29395}}, + pages = {54--67}, + volume = 39, + number = 1, + public = {yes} +} + +@Article{ oleary.ea:formally:1999, + language = {USenglish}, + abstract = {This paper describes the formal specification and + verification of floating-point arithmetic hardware at the + level of IEEE Standard 754. Floating-point correctness is a + crucial problem: the functionality of Intel's + floating-point hardware is architecturally visible and, + once discovered, floating-point bugs are easily reproduced + by the consumer. We have formally specified and verified + IEEE-compliance of the Pentium{\textregistered} Pro + processor's FADD, FSUB, FMUL, FDIV, FSQRT, and FPREM + operations, as well as the correctness of various + miscellaneous operations including conversion to and from + integers. Compliance was verified against the gate-level + descriptions from which the actual silicon is derived and + on which all traditional pre-silicon dynamic validation is + performed. Our results demonstrate that formal functional + verification of gate-level floating-point designs against + IEEE-level specifications is both feasible and practical. + As far as the authors are aware, this is the first such + demonstration. }, + journal = {Intel Technology Journal}, + volume = {Q1}, + year = 1999, + title = {Formally Verifying {IEEE} Compliance of Floating-Point + Hardware}, + author = {John O'Leary and Xudong Zhao and Rob Gerth and Carl-Johan + H. Seger}, + url = {\url{http://developer.intel.com/technology/itj/q11999/articles/art_5.htm}} + , + keywords = {verification; pentium; FDIV bug; flaw; + floating-point-hardware; floating-point; IEEE compliance; + formal verification; model checking; theorem proving } +} + +@Article{ neubauer:feinheiten:1996, + author = {Marion Neubauer}, + title = {Feinheiten bei wissenschaftlichen {P}ublikationen -- + {M}ikrotypographie-{R}egeln, {T}eil {I}}, + journal = dtk, + year = 1996, + volume = {4/96}, + altvolume = 8, + altnumber = 4, + month = feb, + pages = {23--40}, + annote = bretter, + localurl = {papers/1996/dtk96_4_neubauer_feinheiten.pdf}, + url = {\url{http://www.dante.de/dante/DTK/}}, + keywords = {Mikrotypographie, Abk{\"u}rzungen, Akronyme, Einheiten, + Himmelsrichtungen, Anf{\"u}hrungszeichen, Satzzeichen, + Auslassungen, Sonderzeichen, email-Adressen, Ligaturen} +} + +@Article{ neubauer:feinheiten:1997, + author = {Marion Neubauer}, + title = {Feinheiten bei wissenschaftlichen {P}ublikationen -- + {M}ikrotypographie-{R}egeln, {T}eil {II}}, + journal = dtk, + year = 1997, + volume = {1/97}, + altvolume = 9, + altnumber = 1, + month = may, + pages = {25--44}, + url = {\url{http://www.dante.de/dante/DTK/}}, + localurl = {papers/1997/dtk97_1_neubauer_feinheiten.pdf}, + annote = bretter, + keywords = {Mikrotypographie, Zahlen, Ziffern, Nummern, Striche, + Klammern, geschachtelte Klammern, Wortzwischenr{\"a}ume, + Abst{\"a}nde, mathematischer Satz, Worttennungen, + Zeilenumbruch} +} + +@Article{ szpiro:mathematics:2003, + author = {George Szpiro}, + url = {http://www.nature.com/cgi-taf/DynaPage.taf?file=/nature/journal/v424/n6944/full/424012a_fs.html} + , + journal = {Nature}, + pages = {12--13}, + month = jul, + year = 2003, + number = 424, + title = {Mathematics: Does the proof stack up?}, + acknowledgement={none}, + bibkey = {szpiro:mathematics:2003} +} + +@Article{ venet:practical:2008, + author = {Arnaud Venet}, + title = {A practical approach to formal software verification by + static analysis}, + journal = {Ada Lett.}, + volume = {XXVIII}, + number = 1, + year = 2008, + issn = {1094-3641}, + pages = {92--95}, + doi = {10.1145/1387830.1387836}, + address = pub-acm:adr, + publisher = pub-acm, + abstract = {Static analysis by Abstract Interpretation is a promising + way for conducting formal verification of large software + applications. In spite of recent successes in the + verification of aerospace codes, this approach has limited + industrial applicability due to the level of expertise + required to engineer static analyzers. In this paper we + investigate a pragmatic approach that consists of focusing + on the most critical components of the application first. + In this approach the user provides a description of the + usage of functionalities in the critical component via a + simple specification language, which is used to drive a + fully automated static analysis engine. We present + experimental results of the application of this approach to + the verification of absence of buffer overflows in a + critical library of the OpenSSH distribution.} +} + +@InProceedings{ balser.ea:formal:2000, + author = {Michael Balser and Wolfgang Reif and Gerhard Schellhorn + and Kurt Stenzel and Andreas Thums}, + title = {Formal System Development with {KIV}}, + booktitle = {FASE}, + year = 2000, + pages = {363--366}, + doi = {10.1007/3-540-46428-X_25}, + crossref = {maibaum:fundamental:2000}, + abstract = {KIV is a tool for formal systems development. It can be + employed, e.g., 1) for the development of safety critical + systems from formal requirements specifications to + executable code, including the verification of safety + requirements and the correctness of implementations, 2) for + semantical foundations of programming languages from a + specification of the semantics to a verified compiler, 3) + for building security models and architectural models as + they are needed for high level ITSEC [7] or CC [1] + evaluations. } +} + +@Proceedings{ maibaum:fundamental:2000, + editor = {T. S. E. Maibaum}, + title = {Fundamental Approaches to Software Engineering, Third + Internationsl Conference, FASE 2000, Held as Part of the + European Joint Conferences on the Theory and Practice of + Software, ETAPS 2000, Berlin, Germany, March 25 - April 2, + 2000, Proceedings}, + booktitle = {FASE}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 1783, + year = 2000, + isbn = {3-540-67261-3} +} + +@InProceedings{ castillo:asm:2001, + author = {Giuseppe Del Castillo}, + title = {The {ASM} Workbench: A Tool Environment for Computer-Aided + Analysis and Validation of Abstract State Machine Models + Tool Demonstration}, + booktitle = {TACAS}, + year = 2001, + pages = {578--581}, + doi = {10.1007/3-540-45319-9_40}, + abstract = {Gurevich{\^a}s Abstract State Machines (ASMs) constitute a + high-level state-based modelling language, which has been + used in a wide range of applications. The ASM Workbench is + a comprehensive tool environment supporting the development + and computer-aided analysis and validation of ASM models. + It is based on a typed version of the ASM language, called + ASM-SL, and includes features for type-checking, + simulation, debugging, and verification of ASM models.}, + crossref = {margaria.ea:tools:2001} +} + +@Proceedings{ margaria.ea:tools:2001, + editor = {Tiziana Margaria and Wang Yi}, + title = {Tools and Algorithms for the Construction and Analysis of + Systems, 7th International Conference, TACAS 2001 Held as + Part of the Joint European Conferences on Theory and + Practice of Software, ETAPS 2001 Genova, Italy, April 2-6, + 2001, Proceedings}, + booktitle = {TACAS}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 2031, + year = 2001, + isbn = {3-540-41865-2} +} + +@Article{ gurevich.ea:semantic:2005, + title = {Semantic essence of {AsmL}}, + journal = {Theoretical Computer Science}, + volume = 343, + number = 3, + pages = {370--412}, + year = 2005, + note = {Formal Methods for Components and Objects}, + issn = {0304-3975}, + doi = {10.1016/j.tcs.2005.06.017}, + author = {Yuri Gurevich and Benjamin Rossman and Wolfram Schulte}, + keywords = {Abstract state machine,Executable specification language}, + abstract = { The Abstract State Machine Language, AsmL, is a novel + executable specification language based on the theory of + Abstract State Machines. AsmL is object-oriented, provides + high-level mathematical data-structures, and is built + around the notion of synchronous updates and finite choice. + AsmL is fully integrated into the .NET framework and + Microsoft development tools. In this paper, we explain the + design rationale of AsmL and provide static and dynamic + semantics for a kernel of the language.} +} + +@Article{ gargantini.ea:metamodel-based:2008, + author = {Angelo Gargantini and Elvinia Riccobene and Patrizia + Scandurra}, + title = {A Metamodel-based Language and a Simulation Engine for + Abstract State Machines}, + journal = j-ucs, + volume = 14, + number = 12, + year = 2008, + pages = {1949--1983}, + abstract = {In this paper, we present a concrete textual notation, + called AsmetaL, and a general-purpose simulation engine, + called AsmetaS, for Abstract State Machine (ASM) + specifications. They have been developed as part of the + ASMETA (ASMs mETAmodelling) toolset, which is a set of + tools for ASMs based on the metamodelling approach of the + Model-driven Engineering. We briefly present the ASMETA + framework, and we discuss how the language and the + simulator have been developed exploiting the advantages + offered by the metamodelling approach. We introduce the + language AsmetaL used to write ASM specifications, and we + provide the AsmetaL encoding of ASM specifications of + increasing complexity. We explain the AsmetaS architecture, + its kernel engine, and how the simulator works within the + ASMETA tool set. We discuss the features currently + supported by the simulator and how it has been validated.} +} + +@Manual{ schmid:introduction:2001, + author = {Joachim Schmid}, + title = {Introduction to {AsmGofer}}, + year = 2001 +} + +@InProceedings{ miller.ea:czt:2005, + author = {Tim Miller and Leo Freitas and Petra Malik and Mark + Utting}, + title = {{CZT} Support for {Z} Extensions}, + year = 2005, + pages = {227--245}, + doi = {10.1007/11589976_14}, + crossref = {romijn.ea:integrated:2005}, + abstract = {Community Z Tools (CZT) is an integrated framework for the + Z formal specification language. In this paper, we show how + it is also designed to support extensions of Z, in a way + that minimises the work required to build a new Z + extension. The goals of the framework are to maximise + extensibility and reuse, and minimise code duplication and + maintenance effort. To achieve these goals, CZT uses a + variety of different reuse mechanisms, including generation + of Java code from a hierarchy of XML schemas, XML templates + for shared code, and several design patterns for maximising + reuse of Java code. The CZT framework is being used to + implement several integrated formal methods, which add + object-orientation, real-time features and process algebra + extensions to Z. The effort required to implement such + extensions of Z has been dramatically reduced by using the + CZT framework.} +} + +@Proceedings{ romijn.ea:integrated:2005, + editor = {Judi Romijn and Graeme Smith and Jaco van de Pol}, + booktitle = {Integrated Formal Methods (IFM)}, + location = {Eindhoven, The Netherlands}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 3771, + year = 2005, + isbn = {3-540-30492-4} +} + +@TechReport{ ashley.ea:enterprise:2003, + author = {Paul Ashley and Satoshi Hada and G{\"u}nter Karjoth and + Calvin Powers and Matthias Schunter}, + editor = {Calvin Powers and Matthias Schunter}, + title = {Enterprise Privacy Authorization Language ({EPAL} 1.2)}, + institution = {IBM}, + year = 2003, + url = {http://www.zurich.ibm.com/security/enterprise-privacy/epal} + +} + +@TechReport{ cisc:securing:2008, + title = {Securing Cyberspace for the 44th Presidency}, + institution = {Center for Strategic and International Studies (CSIS)}, + month = dec, + year = 2008 +} + +@InProceedings{ wei.ea:authorization:2008, + author = {Qiang Wei and Jason Crampton and Konstantin Beznosov and + Matei Ripeanu}, + title = {Authorization recycling in {RBAC} systems}, + booktitle = {ACM symposium on Access control models and technologies + (SACMAT)}, + year = 2008, + isbn = {978-1-60558-129-3}, + pages = {63--72}, + location = {Estes Park, CO, USA}, + doi = {10.1145/1377836.1377848}, + publisher = pub-acm, + address = pub-acm:adr, + abstract = {As distributed applications increase in size and + complexity, traditional authorization mechanisms based on a + single policy decision point are increasingly fragile + because this decision point represents a single point of + failure and a performance bottleneck. Authorization + recycling is one technique that has been used to address + these challenges. + + This paper introduces and evaluates the mechanisms for + authorization recycling in RBAC enterprise systems. The + algorithms that support these mechanisms allow precise and + approximate authorization decisions to be made, thereby + masking possible failures of the policy decision point and + reducing its load. We evaluate these algorithms + analytically and using a prototype implementation. Our + evaluation results demonstrate that authorization recycling + can improve the performance of distributed access control + mechanisms.} +} + +@Article{ karjoth:access:2003, + author = {G{\"u}nter Karjoth}, + title = {Access control with {IBM} {Tivoli} access manager}, + journal = j-tissec, + publisher = pub-acm, + address = pub-acm:adr, + volume = 6, + number = 2, + year = 2003, + pages = {232--257}, + doi = {10.1145/762476.762479}, + abstract = {Web presence has become a key consideration for the + majority of companies and other organizations. Besides + being an essential information delivery tool, the Web is + increasingly being regarded as an extension of the + organization itself, directly integrated with its operating + processes. As this transformation takes place, security + grows in importance. IBM Tivoli Access Manager offers a + shared infrastructure for authentication and access + management, technologies that have begun to emerge in the + commercial marketplace. This paper describes the + Authorization Service provided by IBM Tivoli Access Manager + for e-business (AM) and its use by AM family members as + well as third-party applications. Policies are defined over + a protected object namespace and stored in a database, + which is managed via a management console and accessed + through an Authorization API. The protected object + namespace abstracts from heterogeneous systems and thus + enables the definition of consistent policies and their + centralized management. ACL inheritance and delegated + management allow these policies to be managed efficiently. + The Authorization API allows applications with their own + access control requirements to decouple authorization logic + from application logic. Policy checking can be externalized + by using either a proxy that sits in front of the Web + servers and application servers or a plug-in that examines + the request. Thus, AM familiy members establish a single + entry point to enforce enterprise policies that regulate + access to corporate data.} +} + +@Book{ heijenoort:from:2002, + abstract = {The fundamental texts of the great classical period in + modern logic, some of them never before available in + English translation, are here gathered together for the + first time. Modern logic, heralded by Leibniz, may be said + to have been initiated by Boole, De Morgan, and Jevons, but + it was the publication in 1879 of Gottlob Frege's + Begriffsschrift that opened a great epoch in the history of + logic by presenting, in full-fledged form, the + propositional calculus and quantification theory. Frege's + book, translated in its entirety, begins the present + volume. The emergence of two new fields, set theory and + foundations of mathematics, on the borders of logic, + mathematics, and philosophy, is depicted by the texts that + follow. Peano and Dedekind illustrate the trend that led to + Principia Mathematica. Burali-Forti, Cantor, Russell, + Richard, and K\"{o}nig mark the appearance of the modern + paradoxes. Hilbert, Russell, and Zermelo show various ways + of overcoming these paradoxes and initiate, respectively, + proof theory, the theory of types, and axiomatic set + theory. Skolem generalizes L\"{o}wenheim's theorem, and he + and Fraenkel amend Zermelo's axiomatization of set theory, + while von Neumann offers a somewhat different system. The + controversy between Hubert and Brouwer during the twenties + is presented in papers of theirs and in others by Weyl, + Bernays, Ackermann, and Kolmogorov. The volume concludes + with papers by Herbrand and by G\"{o}del, including the + latter's famous incompleteness paper. Of the forty-five + contributions here collected all but five are presented in + extenso. Those not originally written in English have been + translated with exemplary care and exactness; the + translators are themselves mathematical logicians as well + as skilled interpreters of sometimes obscure texts. Each + paper is introduced by a note that sets it in perspective, + explains its importance, and points out difficulties in + interpretation. Editorial comments and footnotes are + interpolated where needed, and an extensive bibliography is + included.}, + author = {Jean van Heijenoort}, + howpublished = {Paperback}, + isbn = 0674324498, + keywords = {frege, godel, logic}, + month = {January}, + posted-at = {2006-05-07 22:17:38}, + priority = 2, + publisher = {{Harvard University Press}}, + title = {From Frege to G{\"o}del : A Source Book in Mathematical + Logic, 1879-1931 (Source Books in the History of the + Sciences)}, + year = 2002 +} + +@InProceedings{ kohler.ea:avoiding:2008, + author = {Mathias Kohler and Andreas Schaad}, + title = {Avoiding Policy-based Deadlocks in Business Processes}, + year = 2008, + pages = {709--716}, + doi = {10.1109/ARES.2008.131}, + address = pub-ieee:adr, + publisher = pub-ieee, + abstract = {In the field of business process management, deadlocks + describe a situation where a workflow execution is blocked + and cannot be completed. We speak of policy-based deadlocks + if such a situation is caused by unsatisfiable resource + requirements due to security constraints specified as part + of the business process. In this paper we propose a method + to avoid policy-based deadlocks by analyzing a workflow's + security constraints, determine the minimal required number + of users, and provide an optimal user-activity assignment + for a deadlock-free workflow execution. We will finally + validate our proposed approach by applying it to a + real-world scenario. }, + booktitle = {Third International Conference on Availability, + Reliability and Security (ARES 2008)}, + location = {Technical University of Catalonia, Barcelona , Spain}, + month = mar +} + +@InProceedings{ modersheim.ea:open-source:2009, + author = {Sebastian M{\"o}dersheim and Luca Vigan{\`o}}, + title = {The Open-Source Fixed-Point Model Checker for Symbolic + Analysis of Security Protocols}, + booktitle = {FOSAD}, + year = 2009, + pages = {166--194}, + doi = {10.1007/978-3-642-03829-7_6}, + crossref = {aldini.ea:foundations:2009} +} + +@Proceedings{ aldini.ea:foundations:2009, + editor = {Alessandro Aldini and Gilles Barthe and Roberto Gorrieri}, + title = {Foundations of Security Analysis and Design V, FOSAD + 2007/2008/2009 Tutorial Lectures}, + booktitle = {FOSAD}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 5705, + year = 2009, + isbn = {978-3-642-03828-0}, + doi = {10.1007/978-3-642-03829-7} +} + +@Article{ dolev.ea:security:1981, + author = {D. Dolev and A. C. Yao}, + title = {On the security of public key protocols}, + journal = {Symposium on Foundations of Computer Science}, + volume = 0, + year = 1981, + issn = {0272-5428}, + pages = {350--357}, + doi = {10.1109/SFCS.1981.32}, + publisher = {IEEE Computer Society}, + address = {Los Alamitos, CA, USA} +} + +@InProceedings{ levin.ea:securing:2009, + author = {Timothy E. Levin and Jeffrey S. Dwoskin and Ganesha + Bhaskara and Thuy D. Nguyen and Paul C. Clark and Ruby B. + Lee and Cynthia E. Irvine and Terry Benzel}, + title = {Securing the Dissemination of Emergency Response Data with + an Integrated Hardware-Software Architecture}, + year = 2009, + pages = {133--152}, + doi = {10.1007/978-3-642-00587-9_9}, + crossref = {chen.ea:trusted:2009}, + abstract = {During many crises, access to sensitive emergency-support + information is required to save lives and property. For + example, for effective evacuations first responders need + the names and addresses of non-ambulatory residents. Yet, + currently, access to such information may not be possible + because government policy makers and third-party data + providers lack confidence that today{\^a}s IT systems will + protect their data. Our approach to the management of + emergency information provides first responders with + temporary, transient access to sensitive information, and + ensures that the information is revoked after the + emergency. The following contributions are presented: a + systematic analysis of the basic forms of trusted + communication supported by the architecture; a + comprehensive method for secure, distributed emergency + state management; a method to allow a userspace application + to securely display data; a multifaceted system analysis of + the confinement of emergency information and the secure and + complete revocation of access to that information at the + closure of an emergency.} +} + +@Proceedings{ chen.ea:trusted:2009, + editor = {Liqun Chen and Chris J. Mitchell and Andrew Martin}, + booktitle = PROC # {International Conference on Trusted + Computing (Trust)}, + location = {Trusted Computing, Second International Conference, Trust + 2009, Oxford, UK, April 6-8, 2009, Proceedings}, + series = s-lncs, + publisher = pub-springer, + address = pub-springer:adr, + volume = 5471, + year = 2009, + isbn = {978-3-642-00586-2}, + doi = {10.1007/978-3-642-00587-9} +} + +@Book{ wunder.ea:verteilte:2009, + editor = {Michael Wunder and J{\"u}rgen Grosche}, + title = {Verteilte F{\"u}hrungsinformationssysteme}, + abstract = {R{\"u}ckblick und Sachstand der technologischen Aspekte + bei der Entwicklung verteilter + F{\"u}hrungsinformationssysteme, einer zentralen Aufgabe in + der Bundeswehr sowie bei Beh{\"o}rden und Organisationen + mit Sicherheitsaufgeben (z.B. Polizei, Rettungskr{\"a}fte). + + Vornehmlich Wissenschaftler der Abteilung + Informationstechnik f{\"u}r F{\"u}hrungssysteme des + Forschungsinstituts f{\"u}r Kommunikation, + Informationsverarbeitung und Ergonomie beschreiben + basierend auf einer 40-j{\"a}hrigen Erfahrung in diesem + Anwendungsgebiet Konzepte und Einzelaspekte bei der + Gestaltung von F{\"u}hrungsinformationssystemen. + Reflektiert werden aktuelle Problembereiche bei der + Vernetzung unterschiedlicher Systeme, einer der derzeit + gr{\"o}{\ss}ten Herausforderungen bei der Neugestaltung der + Abl{\"a}ufe und Systeme in der Bundeswehr und in verwandten + Einrichtungen. Dazu werden Informationsstrukturen und + Prozesse untersucht, Systemarchitekturen ausgewertet und + kombiniert sowie Laborstudien und Feldversuche beschrieben. + + In ca. 25 Beitr{\"a}gen wird eine L{\"u}cke in der + verf{\"u}gbaren Literatur geschlossen, die der Vielzahl von + Entwicklern und Anwendern einen Einblick in die aktuelle + Lage und die zuk{\"u}nftigen Gestaltungsm{\"o}glichkeiten bietet.}, + isbn = {978-3-642-00508-4}, + language = {German}, + doi = {10.1007/978-3-642-00509-1}, + year = 2009, + publisher = pub-springer, + address = pub-springer:adr +} + +@InProceedings{ levin.ea:idea:2009, + author = {Timothy E. Levin and Cynthia E. Irvine and Terry Benzel + and Thuy D. Nguyen and Paul C. Clark and Ganesha Bhaskara}, + title = {Idea: Trusted Emergency Management}, + booktitle = {ESSoS}, + year = 2009, + pages = {32--36}, + doi = {10.1007/978-3-642-00199-4_3}, + abstract = {Through first-responder access to sensitive information + for which they have not been pre-vetted, lives and property + can be saved. We describe enhancements to a trusted + emergency information management (EIM) system that securely + allows for extraordinary access to sensitive information + during a crisis. A major component of the architecture is + the end-user device, the security of which is enhanced with + processor-level encryption of memory. This paper introduces + an approach to more efficiently use the + processor-encryption feature for secure data storage, as + well as ISA instructions for the management of emergency + state.}, + crossref = {massacci.ea:engineering:2009} +} + +@Proceedings{ massacci.ea:engineering:2009, + editor = {Fabio Massacci and Samuel T. Redwine Jr. and Nicola + Zannone}, + title = {Engineering Secure Software and Systems, First + International Symposium ESSoS 2009, Leuven, Belgium, + February 4-6, 2009. Proceedings}, + booktitle = {ESSoS}, + volume = 5429, + year = 2009, + isbn = {978-3-642-00198-7}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + doi = {10.1007/978-3-642-00199-4} +} + +@Article{ phan.ea:survey:2008, + author = {Tan Phan and Jun Han and Jean-Guy Schneider and Tim + Ebringer and Tony Rogers}, + title = {A Survey of Policy-Based Management Approaches for Service + Oriented Systems}, + journal = {Australian Software Engineering Conference}, + volume = 0, + year = 2008, + issn = {1530-0803}, + pages = {392--401}, + doi = {10.1109/ASWEC.2008.56}, + address = pub-ieee:adr, + publisher = pub-ieee, + abstract = {Policy based management in Service Oriented Architecture + (SOA) allows organizations to apply rules and regulations + on their business processes. Policy has long been employed + in the management of traditional distributed systems and + many policy frameworks have been proposed. However,SOA + differs in several aspects to traditional systems; thus, + there is a unique set of requirements for an effective SOA + policy system. In this paper, we evaluate five popular + policy frameworks which are IETF, Ponder, KAoS, Rei and + WS-Policy against a number of general and SOA-specific + criteria to identify what features of these existing + systems can be adopted for SOA and what are not. We then, + based on their feature sets, discuss the applicability of + the frameworks for SOA management.} +} + +@InProceedings{ sevinc.ea:securing:2007, + author = {Paul E. Sevin\c{c} and Mario Strasser and David A. Basin}, + title = {Securing the Distribution and Storage of Secrets with + Trusted Platform Modules}, + booktitle = {WISTP}, + year = 2007, + pages = {53--66}, + doi = {10.1007/978-3-540-72354-7_5}, + crossref = {sauveron.ea:information:2007}, + abstract = {We present a protocol that allows servers to securely + distribute secrets to trusted platforms. The protocol + maintains the confidentiality of secrets in the face of + eavesdroppers and careless users. Given an ideal + (tamper-proof) trusted platform, the protocol can even + withstand attacks by dishonest users. As an example of its + use, we present an application to secure document + processing.} +} + +@Proceedings{ sauveron.ea:information:2007, + editor = {Damien Sauveron and Constantinos Markantonakis and Angelos + Bilas and Jean-Jacques Quisquater}, + title = {Information Security Theory and Practices. Smart Cards, + Mobile and Ubiquitous Computing Systems, First IFIP TC6 / + WG 8.8 / WG 11.2 International Workshop, WISTP 2007, + Heraklion, Crete, Greece, May 9-11, 2007, Proceedings}, + booktitle = {WISTP}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 4462, + year = 2007, + isbn = {978-3-540-72353-0} +} + +@InProceedings{ liu.ea:fabric:2009, + author = {Jed Liu and Michael D. George and K. Vikram and Xin Qi and + Lucas Waye and Andrew C. Myers}, + title = {Fabric: a platform for secure distributed computation and + storage}, + booktitle = {SOSP '09: Proceedings of the ACM SIGOPS 22nd symposium on + Operating systems principles}, + year = 2009, + isbn = {978-1-60558-752-3}, + pages = {321--334}, + location = {Big Sky, Montana, USA}, + doi = {10.1145/1629575.1629606}, + address = pub-acm:adr, + publisher = pub-acm, + abstract = {Fabric is a new system and language for building secure + distributed information systems. It is a decentralized + system that allows heterogeneous network nodes to securely + share both information and computation resources despite + mutual distrust. Its high-level programming language makes + distribution and persistence largely transparent to + programmers. Fabric supports data-shipping and + function-shipping styles of computation: both computation + and information can move between nodes to meet security + requirements or to improve performance. Fabric provides a + rich, Java-like object model, but data resources are + labeled with confidentiality and integrity policies that + are enforced through a combination of compile-time and + run-time mechanisms. Optimistic, nested transactions ensure + consistency across all objects and nodes. A peer-to-peer + dissemination layer helps to increase availability and to + balance load. Results from applications built using Fabric + suggest that Fabric has a clean, concise programming model, + offers good performance, and enforces security.} +} + +@InProceedings{ ferreira.ea:how:2009, + author = {Ana Ferreira and David Chadwick and Pedro Farinha and + Gansen Zhao and Rui Chilro and Ricardo Cruz-Correia and + Luis Antunes}, + title = {How to securely break into RBAC: the BTG-RBAC model}, + booktitle = {Annual Computer Security Applications Conference (ACSAC)}, + year = 2009, + abstract = {Access control models describe frameworks that dictate how + subjects (e.g. users) access resources. In the Role-Based + Access Control (RBAC) model access to resources is based on + the role the user holds within the organization. Although + flexible and easier to manage within large-scale + authorization frameworks, RBAC is usually a static model + where access control decisions have only two output + options: Grant or Deny. Break The Glass (BTG) policies can + be provided in order to break or override the access + controls within an access control policy but in a + controlled and justifiable manner. The main objective of + this paper is to integrate BTG within the NIST/ANSI RBAC + model in a transparent and secure way so that it can be + adopted generically in any domain where unanticipated or + emergency situations may occur. The new proposed model, + called BTG-RBAC, provides a third decision option BTG. This + allows break the glass policies to be implemented in any + application without any major changes to either the + application or the RBAC authorization infrastructure, apart + from the decision engine. Finally, in order to validate the + model, we discuss how the BTG-RBAC model is being + introduced within a Portuguese healthcare institution where + the legislation requires that genetic information must be + accessed by a restricted group of healthcare professionals. + These professionals, advised by the ethical committee, have + required and asked for the implementation of the BTG + concept in order to comply with the said legislation.} +} + +@Article{ moggi:notions:1991, + author = {Eugenio Moggi}, + title = {Notions of Computation and Monads}, + journal = {Information and Computation}, + volume = 93, + number = 1, + year = 1991, + pages = {55--92} +} + +@InProceedings{ wadler:monads:1995, + author = {Philip Wadler}, + title = {Monads for Functional Programming}, + booktitle = {Advanced Functional Programming}, + year = 1995, + pages = {24--52}, + crossref = {jeuring.ea:advanced:1995} +} + +@Proceedings{ jeuring.ea:advanced:1995, + editor = {Johan Jeuring and Erik Meijer}, + title = {Advanced Functional Programming, First International + Spring School on Advanced Functional Programming + Techniques, B{\aa}stad, Sweden, May 24-30, 1995, Tutorial + Text}, + booktitle = {Advanced Functional Programming}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 925, + year = 1995, + isbn = {3-540-59451-5} +} + +@InProceedings{ grieskamp.ea:model-based:2008, + author = {Wolfgang Grieskamp and Nicolas Kicillof and Dave MacDonald + and Alok Nandan and Keith Stobie and Fred L. Wurden}, + title = {Model-Based Quality Assurance of Windows Protocol + Documentation}, + booktitle = {Software Testing, Verification, and Validation (ICST)}, + year = 2008, + pages = {502--506}, + doi = {10.1109/ICST.2008.50}, + abstract = {Microsoft is producing high-quality documentation for + Windows client-server and server-server protocols. Our + group in the Windows organization is responsible for + verifying the documentation to ensure it is of the highest + quality. We are applying various test-driven methods + including, when appropriate, a model-based approach. This + paper describes certain aspects of the quality assurance + process we put in place, and specifically focuses on + model-based testing (MBT). Our experiences so far confirm + that MBT works and that it scales, provided it is + accompanied by sound tool support and clear methodological + guidance.}, + location = {Lillehammer, Norway, April 9-11, 2008}, + volume = 0, + isbn = {978-0-7695-3127-4}, + publisher = pub-ieee, + address = pub-ieee:adr +} + +@InProceedings{ berghofer.ea:random:2004, + author = {Stefan Berghofer and Tobias Nipkow}, + title = {Random Testing in Isabelle/HOL}, + booktitle = {Software Engineering and Formal Methods (SEFM)}, + year = 2004, + pages = {230--239}, + doi = {10.1109/SEFM.2004.36}, + abstract = {When developing non-trivial formalizations in a theorem + prover, a considerable amount of time is devoted to + "debugging" specifications and conjectures by failed proof + attempts. To detect such problems early in the proof and + save development time, we have extended the Isabelle + theorem prover with a tool for testing specifications by + evaluating propositions under an assignment of random + values to free variables. Distribution of the test data is + optimized via mutation testing. The technical contributions + are an extension of earlier work with inductive definitions + and a generic method for randomly generating elements of + recursive datatypes.}, + location = {28-30 September 2004, Beijing, China}, + publisher = pub-ieee, + address = pub-ieee:adr, + isbn = {0-7695-2222-X} +} + +@TechReport{ gallaher.ea:economic:2002, + institution = {National Institute of Standards \& Technology}, + number = {Planning Report 02-03}, + year = 2002, + month = may, + author = {M.P. Gallaher and B.M. Kropp}, + title = {The Economic Impacts of Inadequate Infrastructure for + Software Testing}, + abstract = {Software has become an intrinsic part of business over the + last decade. Virtually every business in the U.S. in every + sector depends on it to aid in the development, production, + marketing, and support of its products and services. + Advances in computers and related technology have provided + the building blocks on which new industries have evolved. + Innovations in the fields of robotic manufacturing, + nanotechnologies, and human genetics research all have been + enabled by low cost computational and control capabilities + supplied by computers and software. In 2000, total sales of + software reached approximately \$180 billion. Rapid growth + has created a significant and high-paid workforce, with + 697,000 employed as software engineers and an additional + 585,000 as computer programmers. Reducing the cost of + software development and improving software quality are + important objectives of the U.S. software industry. + However, the complexity of the underlying software needed + to support the U.S.'s computerized economy is increasing at + an alarming rate. The size of software products is no + longer measured in terms of thousands of lines of code, but + millions of lines of code. This increasing complexity along + with a decreasing average market life expectancy for many + software products has heightened concerns over software + quality.} +} + +@InProceedings{ tej.ea:corrected:1997, + author = {Haykal Tej and Burkhart Wolff}, + title = {A Corrected Failure Divergence Model for {CSP} in + {Isabelle/HOL}}, + year = 1997, + pages = {318--337}, + doi = {10.1007/3-540-63533-5_17}, + abstract = {We present a failure-divergence model for CSP following + the concepts of [BR 85]. Its formal representation within + higher order logic in the theorem prover Isabelle/HOL [Pau + 94] revealed an error in the basic definition of CSP + concerning the treatment of the termination symbol tick. A + corrected model has been formally proven consistent with + Isabelle/HOL. Moreover, the changed version maintains the + essential algebraic properties of CSP. As a result, there + is a proven correct implementation of a ldquoCSP + workbenchrdquo within Isabelle.}, + crossref = {fitzgerald.ea:formal:1997} +} + +@Proceedings{ fitzgerald.ea:formal:1997, + editor = {John S. Fitzgerald and Cliff B. Jones and Peter Lucas}, + booktitle = {Formal Methods Europe (FME)}, + location = {Graz, Austria, September 15-19, 1997, Proceedings}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 1313, + year = 1997, + isbn = {3-540-63533-5} +} + +@InProceedings{ bentakouk.ea:formal:2009, + author = {Lina Bentakouk and Pascal Poizat and Fatiha Za\"{\i}di}, + title = {A Formal Framework for Service Orchestration Testing Based + on Symbolic Transition Systems}, + year = 2009, + pages = {16--32}, + doi = {10.1007/978-3-642-05031-2_2}, + crossref = {nunez.ea:testing:2009}, + abstract = {The pre-eminent role played by software composition, and + more particularly service composition, in modern software + development, together with the complexity of workflow + languages such as WS-BPEL have made composite service + testing a topical issue. In this article we contribute to + this issue with an automatic testing approach for WS-BPEL + orchestrations. Compared to related work, we support + WS-BPEL data computations and exchanges, while overcoming + the consequential state explosion problem. This is achieved + through the use of symbolic transition system models and + their symbolic execution. Throughout the article, we + illustrate our approach on a realistic medium-size example. + + } +} + +@Proceedings{ nunez.ea:testing:2009, + editor = {Manuel N{\'u}{\~n}ez and Paul Baker and Mercedes G. Merayo}, + title = {Testing of Software and Communication Systems, 21st IFIP + WG 6.1 International Conference, TESTCOM 2009 and 9th + International Workshop, FATES 2009, Eindhoven, The + Netherlands, November 2-4, 2009. Proceedings}, + booktitle = {TestCom/FATES}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 5826, + year = 2009, + isbn = {978-3-642-05030-5}, + doi = {10.1007/978-3-642-05031-2} +} + +@InProceedings{ anand.ea:demand-driven:2008, + author = {Saswat Anand and Patrice Godefroid and Nikolai Tillmann}, + title = {Demand-Driven Compositional Symbolic Execution}, + booktitle = {TACAS}, + year = 2008, + pages = {367--381}, + doi = {10.1007/978-3-540-78800-3_28}, + crossref = {ramakrishnan.ea:tools:2008}, + abstract = {We discuss how to perform symbolic execution of large + programs in a manner that is both compositional (hence more + scalable) and demand-driven. Compositional symbolic + execution means finding feasible interprocedural program + paths by composing symbolic executions of feasible + intraprocedural paths. By demand-driven, we mean that as + few intraprocedural paths as possible are symbolically + executed in order to form an interprocedural path leading + to a specific target branch or statement of interest (like + an assertion). A key originality of this work is that our + demand-driven compositional interprocedural symbolic + execution is performed entirely using first-order logic + formulas solved with an off-the-shelf SMT + (Satisfiability-Modulo-Theories) solver {\^a} no procedure + in-lining or custom algorithm is required for the + interprocedural part. This allows a uniform and elegant way + of summarizing procedures at various levels of detail and + of composing those using logic formulas. We have + implemented a prototype of this novel symbolic execution + technique as an extension of Pex, a general automatic + testing framework for .NET applications. Preliminary + experimental results are encouraging. For instance, our + prototype was able to generate tests triggering assertion + violations in programs with large numbers of program paths + that were beyond the scope of non-compositional test + generation. } +} + +@Proceedings{ ramakrishnan.ea:tools:2008, + editor = {C. R. Ramakrishnan and Jakob Rehof}, + title = {Tools and Algorithms for the Construction and Analysis of + Systems, 14th International Conference, TACAS 2008, Held as + Part of the Joint European Conferences on Theory and + Practice of Software, ETAPS 2008, Budapest, Hungary, March + 29-April 6, 2008. Proceedings}, + booktitle = {TACAS}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 4963, + year = 2008, + isbn = {978-3-540-78799-0} +} + +@InProceedings{ boyapati.ea:korat:2002, + author = {Chandrasekhar Boyapati and Sarfraz Khurshid and Darko + Marinov}, + title = {{Korat}: automated testing based on {Java} predicates}, + booktitle = {ISSTA}, + year = 2002, + pages = {123--133}, + doi = {10.1145/566172.566191}, + abstract = {This paper presents Korat, a novel framework for automated + testing of Java programs. Given a formal specification for + a method, Korat uses the method precondition to + automatically generate all (nonisomorphic) test cases up to + a given small size. Korat then executes the method on each + test case, and uses the method postcondition as a test + oracle to check the correctness of each output.To generate + test cases for a method, Korat constructs a Java predicate + (i.e., a method that returns a boolean) from the method's + pre-condition. The heart of Korat is a technique for + automatic test case generation: given a predicate and a + bound on the size of its inputs, Korat generates all + (nonisomorphic) inputs for which the predicate returns + true. Korat exhaustively explores the bounded input space + of the predicate but does so efficiently by monitoring the + predicate's executions and pruning large portions of the + search space.This paper illustrates the use of Korat for + testing several data structures, including some from the + Java Collections Framework. The experimental results show + that it is feasible to generate test cases from Java + predicates, even when the search space for inputs is very + large. This paper also compares Korat with a testing + framework based on declarative specifications. Contrary to + our initial expectation, the experiments show that Korat + generates test cases much faster than the declarative + framework.} +} + +@Article{ visser.ea:model:2003, + author = {Willem Visser and Klaus Havelund and Guillaume P. Brat and + Seungjoon Park and Flavio Lerda}, + title = {Model Checking Programs}, + journal = {Autom. Softw. Eng.}, + volume = 10, + number = 2, + year = 2003, + pages = {203--232}, + doi = {10.1023/A:1022920129859}, + abstract = {The majority of work carried out in the formal methods + community throughout the last three decades has (for good + reasons) been devoted to special languages designed to make + it easier to experiment with mechanized formal methods such + as theorem provers, proof checkers and model checkers. In + this paper we will attempt to give convincing arguments for + why we believe it is time for the formal methods community + to shift some of its attention towards the analysis of + programs written in modern programming languages. In + keeping with this philosophy we have developed a + verification and testing environment for Java, called Java + PathFinder (JPF), which integrates model checking, program + analysis and testing. Part of this work has consisted of + building a new Java Virtual Machine that interprets Java + bytecode. JPF uses state compression to handle big states, + and partial order and symmetry reduction, slicing, + abstraction, and runtime analysis techniques to reduce the + state space. JPF has been applied to a real-time avionics + operating system developed at Honeywell, illustrating an + intricate error, and to a model of a spacecraft controller, + illustrating the combination of abstraction, runtime + analysis, and slicing with model checking.} +} + +@InProceedings{ bjorner.ea:path:2009, + author = {Nikolaj Bj{\o}rner and Nikolai Tillmann and Andrei + Voronkov}, + title = {Path Feasibility Analysis for String-Manipulating + Programs}, + booktitle = {TACAS}, + year = 2009, + pages = {307--321}, + doi = {10.1007/978-3-642-00768-2_27}, + crossref = {kowalewski.ea:tools:2009}, + abstract = {We discuss the problem of path feasibility for programs + manipulating strings using a collection of standard string + library functions. We prove results on the complexity of + this problem, including its undecidability in the general + case and decidability of some special cases. In the context + of test-case generation, we are interested in an efficient + finite model finding method for string constraints. To this + end we develop a two-tier finite model finding procedure. + First, an integer abstraction of string constraints are + passed to an SMT (Satisfiability Modulo Theories) solver. + The abstraction is either unsatisfiable, or the solver + produces a model that fixes lengths of enough strings to + reduce the entire problem to be finite domain. The + resulting fixed-length string constraints are then solved + in a second phase. We implemented the procedure in a + symbolic execution framework, report on the encouraging + results and discuss directions for improving the method + further.} +} + +@Proceedings{ kowalewski.ea:tools:2009, + editor = {Stefan Kowalewski and Anna Philippou}, + title = {Tools and Algorithms for the Construction and Analysis of + Systems, 15th International Conference, TACAS 2009, Held as + Part of the Joint European Conferences on Theory and + Practice of Software, ETAPS 2009, York, UK, March 22-29, + 2009. Proceedings}, + booktitle = {TACAS}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 5505, + year = 2009, + isbn = {978-3-642-00767-5}, + doi = {10.1007/978-3-642-00768-2} +} + +@InProceedings{ huima:implementing:2007, + author = {Antti Huima}, + title = {Implementing Conformiq Qtronic}, + booktitle = {TestCom/FATES}, + year = 2007, + pages = {1--12}, + doi = {10.1007/978-3-540-73066-8_1}, + crossref = {petrenko.ea:testing:2007}, + abstract = {Conformiq Qtronic is a commercial tool for model driven + testing. It derives tests automatically from behavioral + system models. These are black-box tests [1] by nature, + which means that they depend on the model and the + interfaces of the system under test, but not on the + internal structure (e.g. source code) of the + implementation. In this essay, which accompanies my invited + talk, I survey the nature of Conformiq Qtronic, the main + implementation challenges that we have encountered and how + we have approached them.} +} + +@Proceedings{ petrenko.ea:testing:2007, + editor = {Alexandre Petrenko and Margus Veanes and Jan Tretmans and + Wolfgang Grieskamp}, + title = {Testing of Software and Communicating Systems, 19th IFIP + TC6/WG6.1 International Conference, TestCom 2007, 7th + International Workshop, FATES 2007, Tallinn, Estonia, June + 26-29, 2007, Proceedings}, + booktitle = {TestCom/FATES}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 4581, + year = 2007, + isbn = {978-3-540-73065-1} +} + +@InProceedings{ tretmans.ea:torx:2003, + howpublished = {http://eprints.eemcs.utwente.nl/9475/}, + month = {December}, + author = {G. J. Tretmans and H. Brinksma}, + booktitle = {First European Conference on Model-Driven Software + Engineering, Nuremberg, Germany}, + editor = {A. Hartman and K. Dussa-Ziegler}, + abstract = {Systematic testing is very important for assessing and + improving the quality of software systems. Yet, testing + turns out to be expensive, laborious, time-consuming and + error-prone. The Dutch research and development project + C\^ote de Resyste worked on methods, techniques and tools + for automating specification based testing using formal + methods. The main achievement of the project is a test + tool, baptized TorX, which integrates automatic test + generation, test execution, and test analysis in an + on-the-fly manner. On the one hand, TorX is based on + well-defined theory, viz. the ioco-test theory, which has + its roots in the theory of testing- and + refusal-equivalences for transition systems. On the other + hand, the applicability of TorX has been demonstrated by + testing several academic and industrial case studies. This + paper summarizes the main results of the project C\^ote de + Resyste.}, + title = {TorX: Automated Model-Based Testing}, + year = 2003, + pages = {31--43}, + location = {Nuremberg, Germany}, + trnumber = 9475, + event_dates = {December 11-12, 2003}, + num_pages = 13 +} + +@InProceedings{ jaffuel.ea:leirios:2007, + author = {Eddie Jaffuel and Bruno Legeard}, + title = {LEIRIOS Test Generator: Automated Test Generation from B + Models}, + booktitle = {B}, + year = 2007, + pages = {277--280}, + doi = {10.1007/11955757_29}, + crossref = {julliand.ea:b:2006}, + abstract = {Since 2003, automated test generation from B abstract + machines has been trying out in the smart card industry, + using LEIRIOS Test Generator (LTG) for SmartCard tool. Now + the major card manufacturers, such as Gemalto and Giesecke + & Devrient, are regularly deploying model-based testing in + their validation processes. The purpose is black-box + functional testing: from the specifications (a standard or + specific requirements), a B formal model is developed which + is the basis for test generation. Generated test cases are + then translated into executable test scripts and then run + on the application.} +} + +@Proceedings{ julliand.ea:b:2006, + editor = {Jacques Julliand and Olga Kouchnarenko}, + title = {B 2007: Formal Specification and Development in B, 7th + International Conference of B Users, Besan\c{c}on, France, + January 17-19, 2007, Proceedings}, + booktitle = {B}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 4355, + year = 2006, + isbn = {3-540-68760-2} +} + +@InProceedings{ hu.ea:enabling:2008, + author = {Hongxin Hu and Gail-Joon Ahn}, + title = {Enabling verification and conformance testing for access + control model}, + booktitle = {ACM symposium on Access control models and technologies + (SACMAT)}, + year = 2008, + isbn = {978-1-60558-129-3}, + pages = {195--204}, + location = {Estes Park, CO, USA}, + doi = {10.1145/1377836.1377867}, + address = pub-acm:adr, + publisher = pub-acm, + abstract = {Verification and testing are the important step for + software assurance. However, such crucial and yet + challenging tasks have not been widely adopted in building + access control systems. In this paper we propose a + methodology to support automatic analysis and conformance + testing for access control systems, integrating those + features to Assurance Management Framework (AMF). Our + methodology attempts to verify formal specifications of a + role-based access control model and corresponding policies + with selected security properties. Also, we systematically + articulate testing cases from formal specifications and + validate conformance to the system design and + implementation using those cases. In addition, we + demonstrate feasibility and effectiveness of our + methodology using SAT and Alloy toolset.} +} + +@Article{ grindal.ea:combination:2005, + author = {Mats Grindal and Jeff Offutt and Sten F. Andler}, + title = {Combination testing strategies: a survey}, + journal = {Softw. Test., Verif. Reliab.}, + volume = 15, + number = 3, + year = 2005, + pages = {167--199}, + doi = {10.1002/stvr.319}, + abstract = {Combination strategies are test case selection methods + that identify test cases by combining values of the + different test object input parameters based on some + combinatorial strategy. This survey presents 16 different + combination strategies, covering more than 40 papers that + focus on one or several combination strategies. This + collection represents most of the existing work performed + on combination strategies. This survey describes the basic + algorithms used by the combination strategies. Some + properties of combination strategies, including coverage + criteria and theoretical bounds on the size of test suites, + are also included in this description. This survey paper + also includes a subsumption hierarchy that attempts to + relate the various coverage criteria associated with the + identified combination strategies.} +} + +@InProceedings{ goga:comparing:2001, + author = {Nicolae Goga}, + title = {Comparing TorX, Autolink, TGV and UIO Test Algorithms}, + booktitle = {SDL Forum}, + year = 2001, + pages = {379--402}, + doi = {10.1007/3-540-48213-X}, + crossref = {reed.ea:sdl:2001}, + abstract = {This paper presents a comparison of four algorithms for + test derivation: TorX, TGV, Autolink and UIO algorithms. + The algorithms are classified according to the detection + power of their conformance rela- tions. Because Autolink + does not have an explicit conformance relation, a + conformance relation is reconstructed for it. The + experimental results obtained by applying TorX, Autolink, + UIO and TGV to the Conference Protocol case study are + consistent with the theoretical results of this paper.} +} + +@Proceedings{ reed.ea:sdl:2001, + editor = {Rick Reed and Jeanne Reed}, + title = {SDL 2001: Meeting UML, 10th International SDL Forum + Copenhagen, Denmark, June 27-29, 2001, Proceedings}, + booktitle = {SDL Forum}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 2078, + year = 2001, + isbn = {3-540-42281-1} +} + +@InProceedings{ belinfante.ea:tools:2004, + author = {Axel Belinfante and Lars Frantzen and Christian + Schallhart}, + title = {Tools for Test Case Generation}, + booktitle = {Model-Based Testing of Reactive Systems}, + year = 2004, + pages = {391--438}, + doi = {10.1007/11498490_18}, + crossref = {broy.ea:model-based:2005} +} + +@Proceedings{ broy.ea:model-based:2005, + editor = {Manfred Broy and Bengt Jonsson and Joost-Pieter Katoen and + Martin Leucker and Alexander Pretschner}, + title = {Model-Based Testing of Reactive Systems, Advanced Lectures + [The volume is the outcome of a research seminar that was + held in Schloss Dagstuhl in January 2004]}, + booktitle = {Model-Based Testing of Reactive Systems}, + volume = 3472, + year = 2005, + isbn = {3-540-26278-4}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs +} + +@InProceedings{ gardner.ea:securing:2009, + author = {Ryan W. Gardner and Sujata Garera and Matthew W. Pagano + and Matthew Green and Aviel D. Rubin}, + title = {Securing medical records on smart phones}, + booktitle = {ACM workshop on Security and privacy in medical and + home-care systems (SPIMACS)}, + year = 2009, + isbn = {978-1-60558-790-5}, + pages = {31--40}, + location = {Chicago, Illinois, USA}, + doi = {10.1145/1655084.1655090}, + address = pub-acm:adr, + publisher = pub-acm, + abstract = {There is an inherent conflict between the desire to + maintain privacy of one's medical records and the need to + make those records available during an emergency. To + satisfy both objectives, we introduce a flexible + architecture for the secure storage of medical records on + smart phones. In our system, a person can view her records + at any time, and emergency medical personnel can view the + records as long as the person is present (even if she is + unconscious). Our solution allows for efficient revocation + of access rights and is robust against adversaries who can + access the phone's storage offline.} +} + + +@Article{ govaerts.ea:formal:2008, + author = {John Govaerts and Arosha K. Bandara and Kevin Curran}, + title = {A formal logic approach to firewall packet filtering + analysis and generation}, + journal = {Artif. Intell. Rev.}, + volume = 29, + number = {3-4}, + year = 2008, + pages = {223--248}, + doi = {10.1007/s10462-009-9147-0}, + publisher = pub-springer, + address = pub-springer:adr, + abstract = {Recent years have seen a significant increase in the usage + of computers and their capabilities to communicate with + each other. With this has come the need for more security + and firewalls have proved themselves an important piece of + the overall architecture, as the body of rules they + implement actually realises the security policy of their + owners. Unfor- tunately, there is little help for their + administrators to understand the actual meaning of the + firewall rules. This work shows that formal logic is an + important tool in this respect, because it is particularly + apt at modelling real-world situations and its formalism is + conductive to reason about such a model. As a consequence, + logic may be used to prove the properties of the models it + represents and is a sensible way to go in order to create + those models on com- puters to automate such activities. We + describe here a prototype which includes a description of a + network and the body of firewall rules applied to its + components. We were able to detect a number of anomalies + within the rule-set: inexistent elements (e.g. hosts or + services on destination components), redundancies in rules + defining the same action for a network and hosts belonging + to it, irrelevance as rules would involve traffic that + would not pass through a filtering device, and + contradiction in actions applied to elements or to a + network and its hosts. The prototype produces actual + firewall rules as well, generated from the model and + expressed in the syntax of IPChains and Cisco's PIX. } +} + +@InProceedings{ jalili.ea:specification:2002, + author = {Rasool Jalili and Mohsen Rezvani}, + title = {Specification and Verification of Security Policies in + Firewalls}, + booktitle = {EurAsia-ICT}, + year = 2002, + pages = {154--163}, + doi = {10.1007/3-540-36087-5_18}, + crossref = {shafazand.ea:eurasia-ict:2002}, + abstract = {Rules are used as a way of managing and configuring + firewalls to fulfill security requirements in most cases. + Managers have to specify their organizational security + policies using low level and order-dependent rules. + Furthermore, dependency of firewalls to the network + topology, frequent changes in network topology (specially + in dynamic networks), and lack of a method for analysis and + verification of specified security policy may reduce to + inconsistencies and security holes. Existence of a higher + level environment for security policy specification can + rectify part of the problems. In this paper we present a + language for high level and formal specification of + security policy in firewalls. Using the language, a + security manager can configure its firewall based on his + required security policy independent of the network + topology. The language is used as a framework for analysis + and verification of security policies. We designed and + implemented a tool based on theorem proving for detecting + inconsistencies, coverage, as well as applying a query on + the specified policy. Results of analysis can be used to + detect security vulnerabilities.} +} + +@Proceedings{ shafazand.ea:eurasia-ict:2002, + editor = {Hassan Shafazand and A. Min Tjoa}, + title = {EurAsia-ICT 2002: Information and Communication + Technology, First EurAsian Conference, Shiraz, Iran, + October 29-31, 2002, Proceedings}, + booktitle = {EurAsia-ICT}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 2510, + year = 2002, + isbn = {3-540-00028-3} +} + +@Article{ sutter:free:2005, + author = {Herb Sutter}, + journal = {Dr. Dobb{\^a}s Journal}, + number = 3, + pages = {202--210}, + title = {The Free Lunch Is Over: A Fundamental Turn Toward + Concurrency in Software}, + url = {http://www.gotw.ca/publications/concurrency-ddj.htm}, + volume = 30, + year = 2005 +} + + +@InProceedings{ lachner.ea:information:2008, + author = {Janine Lachner and Hermann Hellwagner}, + title = {Information and Communication Systems for Mobile Emergency + Response}, + year = 2008, + pages = {213--224}, + doi = {10.1007/978-3-540-78942-0_22}, + crossref = {kaschek.ea:information:2008}, + abstract = {This discussion paper attempts to propose emergency + response and disaster management as worthwhile areas of + applied research for the information system community. The + typical requirements, entities and activities involved in + specifically mobile emergency response operations are + summarized. Recent research contributions in this area are + exemplarily reviewed in order to give a deeper insight into + the role and use of mobile information and communication + systems. Finally, the major challenges and research needs + regarding information systems are summarized, with a view + to draw the attention of information systems researchers to + this interesting and important field.} +} + +@Proceedings{ kaschek.ea:information:2008, + editor = {Roland Kaschek and Christian Kop and Claudia Steinberger + and G{\"u}nther Fliedl}, + title = {Information Systems and e-Business Technologies, 2nd + International United Information Systems Conference, + UNISCON 2008, Klagenfurt, Austria, April 22-25, 2008, + Proceedings}, + booktitle = {Information Systems and e-Business Technologies (UNISCON)}, + series = s-lnbip, + volume = 5, + year = 2008, + publisher = pub-springer, + address = pub-springer:adr, + isbn = {978-3-540-78941-3}, +} +@InProceedings{ johnson:complexity:2008, + author = {Chris W. Johnson}, + title = {Complexity, Structured Chaos and the Importance of + Information Management for Mobile Computing in the {UK} + Floods of 2007}, + booktitle = {Mobile Response}, + year = 2008, + pages = {1--11}, + doi = {10.1007/978-3-642-00440-7_1}, + crossref = {loffler.ea:mobile:2009}, + abstract = {Many research teams have developed mobile computing + architectures to support the emergency and rescue services + in a range of civil contingencies. These proposals are + based on innovative technologies and show considerable + creativity in the design of their user interfaces. In + contrast, this paper presents lessons learned from the 2007 + UK floods. Mobile telecommunications failed in many + different ways and from many different causes, including + physical damage to handsets, as well as the loss of base + stations and UPSs. The insights gained from the floods are + being used to inform the design of next generation mobile + digital communications systems for UK responders. However, + the technical problems are arguably less important than the + insights that were obtained about {\^a}systemic{\^a} + failures in the interfaces between local government, + emergency services and the variety of agencies that must + cooperate in major civil contingencies. Problems in + information management led to inconsistencies and + incompatibilities. In consequence, the output from one + application could not easily be used as input to systems + operated by other agencies. These issues must be addressed + before we are overwhelmed by the increased bandwidth + afforded by new mobile devices and novel sensing + technologies. It is concluded that unless we understand the + chaos, complexity and the contextual issues that + characterise previous emergency situations then there is + little prospect that we will be able to design effective + mobile technologies for future incidents.} +} + +@Proceedings{ loffler.ea:mobile:2009, + editor = {Jobst L{\"o}ffler and Markus Klann}, + location = {Mobile Response, Second International Workshop on Mobile + Information Technology for Emergency Response, + MobileResponse 2008. Bonn, Germany, May 29-30, 2008, + Revised Selected Papers}, + title = {Mobile Information Technology for Emergency Response, + (MobileResponse)}, + booktitle = {Mobile Information Technology for Emergency Response, + (MobileResponse)}, + publisher = pub-springer, + address = pub-springer:adr, + series = s-lncs, + volume = 5424, + year = 2009, + isbn = {978-3-642-00439-1}, + doi = {10.1007/978-3-642-00440-7} +} + + +@book{Hoare:1985:CSP:3921, + author = {Hoare, C. A. R.}, + title = {Communicating Sequential Processes}, + year = {1985}, + isbn = {0-13-153271-5}, + publisher = {Prentice-Hall, Inc.}, + address = {Upper Saddle River, NJ, USA}, +} + +@InProceedings{brookes-roscoe85, +author="Brookes, S. D. and Roscoe, A. W.", +editor="Brookes, Stephen D. and Roscoe, Andrew William and Winskel, Glynn", +title="An improved failures model for communicating processes", +booktitle="Seminar on Concurrency", +year="1985", +publisher="Springer Berlin Heidelberg", +address="Berlin, Heidelberg", +pages="281--305", +abstract="We extend the failures model of communicating processes to allow a more satisfactory treatment of divergence in addition to deadlock. The relationship between the revised model and the old model is discussed, and we make some connections with various models proposed by other authors.", +isbn="978-3-540-39593-5" +} + +@techreport{KriegBrueckner95, + author = {Krieg-Br\"uckner, B. and Peleska, J. and Olderog, E.-R. and Balzer, D. and Baer, A.}, + title = "Uniform Workbench --- Universelle Entwicklungsumgebung f{\"u}r formale Methoden", + institution = "Technischer Bericht 8/95", + year=1995, + address="Univ. Bremen", + note="\url{http://www.informatik.uni-bremen.de/~uniform}" +} + +@InProceedings{Camilleri91, +author="Camilleri, Albert J.", +editor="Birtwistle, Graham", +title="A Higher Order Logic Mechanization of the CSP Failure-Divergence Semantics", +booktitle="IV Higher Order Workshop, Banff 1990", +year="1991", +publisher="Springer London", +address="London", +pages="123--150", +abstract="Reasoning using process algebras often involves doing complex proofs, and computer-based support to facilitate the task is therefore desirable. In this paper we show how a general-purpose theorem prover based on higher order logic provides a natural framework for mechanizing the process algebra CSP. This is done by defining the semantics of the CSP operators in the logic and proving the high-level algebraic laws from the definitions as theorems. We mechanize a variation on the failure-divergence semantics that does not use alphabets at the syntactic level, but embeds them in the semantics. Our approach abstracts further from the explicit use of alphabets by modelling them as type variables. The result is a mechanized theory for a polymorphic formalization of CSP.", +isbn="978-1-4471-3182-3" +} + + +@InProceedings{FDRTutorial2000, +author="Broadfoot, Philippa +and Roscoe, Bill", +editor="Havelund, Klaus +and Penix, John +and Visser, Willem", +title="Tutorial on FDR and Its Applications", +booktitle="SPIN Model Checking and Software Verification", +year="2000", +publisher="Springer Berlin Heidelberg", +address="Berlin, Heidelberg", +pages="322--322", +abstract="FDR [1] is a refinement checker for the process algebra CSP [2,4], based on that language's well-established semantic models. FDR stands for Failures-Divergences Refinement, after the premier model. In common with many other model checkers, it works by ``determinising'' (or normalising) a specification and enumerating states in the cartesian product of this and the implementation. Unlike most, the specification and implementation are written in the same language. Under development by its creators, Formal Systems (a spin-off of the Computing Laboratory) since 1991, it now offers a range of state compression methods. On current workstations it can work at up to 20M states/hour with only a small degradation on moving to disc-based storage.", +isbn="978-3-540-45297-3" +} + +@article{IsobeRoggenbach2010, + title={CSP-Prover: a Proof Tool for the Verification of Scalable Concurrent Systems}, + author={Yoshinao Isobe and Markus Roggenbach}, + journal={Information and Media Technologies}, + volume={5}, + number={1}, + pages={32-39}, + year={2010}, + doi={10.11185/imt.5.32} +} + +@article{Noninterference_CSP-AFP, + author = {Pasquale Noce}, + title = {Noninterference Security in Communicating Sequential Processes}, + journal = {Archive of Formal Proofs}, + month = may, + year = 2014, + note = {\url{http://isa-afp.org/entries/Noninterference_CSP.html}, + Formal proof development}, + ISSN = {2150-914x}, +} + +@article{Noninterference_Sequential_Composition-AFP, + author = {Pasquale Noce}, + title = {Conservation of CSP Noninterference Security under Sequential Composition}, + journal = {Archive of Formal Proofs}, + month = apr, + year = 2016, + note = {\url{http://isa-afp.org/entries/Noninterference_Sequential_Composition.html}, + Formal proof development}, + ISSN = {2150-914x}, +} + +@article{Noninterference_Concurrent_Composition-AFP, + author = {Pasquale Noce}, + title = {Conservation of CSP Noninterference Security under Concurrent Composition}, + journal = {Archive of Formal Proofs}, + month = jun, + year = 2016, + note = {\url{http://isa-afp.org/entries/Noninterference_Concurrent_Composition.html}, + Formal proof development}, + ISSN = {2150-914x}, +} + + +@article{Feliachi-Wolff-Gaudel-AFP12, + author = {Abderrahmane Feliachi and Burkhart Wolff and Marie-Claude Gaudel}, + title = {Isabelle/Circus}, + journal = {Archive of Formal Proofs}, + month = jun, + year = 2012, + note = {\url{http://afp.sourceforge.net/entries/Circus.shtml}, + Formal proof development}, + area = {logical_representations}, + abstract = {The Circus specification language combines elements for + complex data and behavior specifications, using an integration of Z and CSP + with a refinement calculus. Its semantics is based on Hoare and He's + Unifying Theories of Programming (UTP). Isabelle/Circus is a formalization of + the UTP and the Circus language in Isabelle/HOL. It contains proof rules and + tactic support that allows for proofs of refinement for Circus processes + (involving both data and behavioral aspects). + The Isabelle/Circus environment supports a syntax for the semantic + definitions which is close to textbook presentations of Circus. This + article contains an extended version of corresponding VSTTE Paper together + with the complete formal development of its underlying commented theories.}, + issn = {2150-914x} +} + + + +@inproceedings{feliachigw12, + author = {Abderrahmane Feliachi and + Marie-Claude Gaudel and + Burkhart Wolff}, + title = {Isabelle/{C}ircus: A Process Specification and Verification Environment}, + booktitle = {VSTTE}, + year = {2012}, + pages = {243-260}, + series = {Lecture Notes in Computer Science}, + volume = {LNCS 7152}, + isbn = {978-3-642-27704-7}, + ee = {http://dx.doi.org/10.1007/978-3-642-27705-4_20}, + doi = {10.1007/978-3-642-27705-4_20}, + abstract = {The Circus specification language combines elements for complex data and behavior specifications, + using an integration of Z and CSP with a refinement calculus. Its semantics is based on Hoare and + He's unifying theories of programming (UTP). We develop a machine-checked, formal semantics based + on a "shallow embedding" of Circus in Isabelle/UTP (our semantic theory of UTP based on Isabelle/HOL). + We derive proof rules from this semantics and imple- ment tactic support that finally allows for + proofs of refinement for Circus processes (involving both data and behavioral aspects). + This proof environment supports a syntax for the semantic definitions which is close to textbook + presentations of Circus. }, + pdf = {http://www.lri.fr/~wolff/papers/conf/VSTTE-IsabelleCircus11.pdf}, + classification = {conference}, + area = {logical_representations}, + bibsource = {DBLP, http://dblp.uni-trier.de} +} + + +@article{feliachi-wolff:SymbTestgenCirta:2013, + author = {Abderrahmane Feliachi and Marie-Claude Gaudel and Burkhart Wolff}, + title = {Symbolic Test-generation in {HOL}-{T}est{G}en/{C}irTA: A Case Study}, + journal = {Int. J. Software Informatics}, + volume = {9}, + number = {2}, + year = {2015}, + pages = {177--203}, + abstract = {HOL-TestGen/CirTA is a theorem-prover based test generation + environment for specifications written in Circus, a + process-algebraic specification language in the tradition + of CSP. HOL-TestGen/CirTA is based on a formal embedding of + its semantics in Isabelle/HOL, allowing to derive rules over + specification constructs in a logically safe way. + Beyond the derivation of algebraic laws and calculi for + process refinement, the originality of HOL-TestGen/ CirTA + consists in an entire derived theory for the generation of + symbolic test-traces, including optimized rules for test-generation + as well as rules for symbolic execution. The deduction process + is automated by Isabelle tactics, allowing to protract the + state-space explosion resulting from blind enumeration of data. + + The implementation of test-generation procedures in CirTA is + completed by an integrated tool chain that transforms the + initial Circus specification of a system into a set of + equivalence classes (or ⤽symbolic tests⤝), which were compiled + to conventional JUnit test-drivers. + + This paper describes the novel tool-chain based on prior + theoretical work on semantics and test-theory and attempts + an evaluation via a medium-sized case study performed on a + component of a real-world safety-critical medical monitoring + system written in Java. We provide experimental measurements + of the kill-capacity of implementation mutants.}, + public = yes, + url = {https://www.lri.fr/~wolff/papers/journals/2015-cirta-case-study.pdf} +} + +@book{donovan2015go, + title={The Go Programming Language}, + author={Donovan, A.A.A. and Kernighan, B.W.}, + isbn={9780134190563}, + series={Addison-Wesley Professional Computing Series}, + year={2015}, + publisher={Pearson Education} +} + + +@ARTICLE{Barret95, +author={G. {Barrett}}, +journal={IEEE Transactions on Software Engineering}, +title={Model checking in practice: the T9000 virtual channel processor}, +year={1995}, +volume={21}, +number={2}, +pages={69-78}, +keywords={transputers;circuit CAD;formal specification;circuit analysis computing;model checking;T9000 virtual channel processor;formal method;industrial products;visual specification styles;case study;nonstandard CSP operators;visual specification style;INMOS CAD system;Design engineering;Design methodology;Hardware;Testing;Product design;Refining;Design automation;Manufacturing processes;Geometry;Process design}, +doi={10.1109/32.345823}, +ISSN={2326-3881}, +month={Feb}} + +@article{BrookesHR84, + author = {S. D. Brookes and + C. A. R. Hoare and + A. W. Roscoe}, + title = {A Theory of Communicating Sequential Processes}, + journal = {J. {ACM}}, + volume = {31}, + number = {3}, + pages = {560--599}, + year = {1984} +} + +@InProceedings{scott:cpo:1972, +author="Scott, Dana", +editor="Lawvere, F. W.", +title="Continuous lattices", +booktitle="Toposes, Algebraic Geometry and Logic", +year="1972", +publisher="Springer", +address="Berlin, Heidelberg", +pages="97--136", +abstract="Starting from the topological point of view a certain wide class of To-spaces is +introduced having a very strong extension property for continuous functions with values in these +spaces. It is then shown that all such spaces are complete lattices whose lattice structure +determines the topology --- these are the continuous lattices --- and every such lattice has +the extension property. With this foundation the lattices are studied in detail with respect +to projections, subspaces, embeddings, and constructions such as products, sums, function spaces, +and inverse limits. The main result of the paper is a proof that every topological space can be +embedded in a continuous lattice which is homeomorphic (and isomorphic) to its own function space. +The function algebra of such spaces provides mathematical models for the +Church-Curry $\lambda$-calculus.", +isbn="978-3-540-37609-5" +} + + +@article{HOL-CSP-AFP, + author = {Safouan Taha and Lina Ye and Burkhart Wolff}, + title = {{HOL-CSP Version 2.0}}, + journal = {Archive of Formal Proofs}, + month = apr, + year = 2019, + note = {\url{http://isa-afp.org/entries/HOL-CSP.html}}, + ISSN = {2150-914x}, +} + + +@incollection{HOL-CSP-iFM2020, + keywords = {Process-Algebra, Concurrency, Computational Models, Parametric System Verification}, + author = {Safouan Taha and Lina Ye and Burkhart Wolff}, + booktitle = {Integrated Formal Methods (iFM)}, + language = {USenglish}, + publisher = {Springer-Verlag}, + address = {Heidelberg}, + series = {Lecture Notes in Computer Science}, + number = {12546}, + doi = {10.1007/978-3-030-63461-2_23}, + editor = {Carlo A. Furia}, + title = {{P}hilosophers may {D}ine - {D}efinitively!}, + classification = {conference}, + areas = {formal methods, software}, + year = {2020}, + public = {yes} +} diff --git a/thys/CSP_RefTK/document/root.tex b/thys/CSP_RefTK/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/CSP_RefTK/document/root.tex @@ -0,0 +1,85 @@ +\documentclass[11pt,a4paper]{book} +\usepackage{isabelle,isabellesym} +\usepackage{graphicx} +\graphicspath {{figures/}} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + + +\usepackage{latexsym} +\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage[greek,english]{babel} + %option greek for \ + %option english (default language) for \, \ + +%\usepackage[latin1]{inputenc} + %for \, \, \, \, + %\, \, \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{The HOL-CSP Refinement Toolkit} +\author{ Safouan Taha \and Burkhart Wolff \and Lina Ye } +\maketitle +\chapter*{Abstract} + + Recently, a modern version of Roscoes and Brookes \cite{brookes-roscoe85} + Failure-Divergence Semantics for CSP has been formalized in Isabelle \cite{HOL-CSP-AFP}. + + We use this formal development called HOL-CSP2.0 to analyse a family of refinement + notions, comprising classic and new ones. + This analysis enables to derive a number of properties that allow to deepen the + understanding of these notions, in particular with respect to specification + decomposition principles for the case of infinite sets of events. The established + relations between the refinement relations help to clarify some obscure points in the + CSP literature, but also provide a weapon for shorter refinement proofs. Furthermore, + we provide a framework for state-normalisation allowing to formally reason on + parameterised process architectures. + + As a result, we have a modern environment for formal proofs of concurrent systems that + allow for the combination of general infinite processes with locally finite ones in a + logically safe way. We demonstrate these verification-techniques for classical, + generalised examples: The CopyBuffer for arbitrary data and the Dijkstra's Dining + Philosopher Problem of arbitrary size. + + If you consider to cite this work, please refer to \cite{HOL-CSP-iFM2020}. + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +% optional bibliography +\bibliographystyle{abbrv} +\bibliography{adb-long,root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,573 +1,574 @@ ADS_Functor AI_Planning_Languages_Semantics AODV AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BNF_CC BNF_Operations Banach_Steinhaus Bell_Numbers_Spivey Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Bondy Boolean_Expression_Checkers Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRDT CYK CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean ClockSynchInst Closest_Pair_Points CofGroups Coinductive Coinductive_Languages Collections Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Consensus_Refined Constructive_Cryptography Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties +CSP_RefTK DFS_Framework DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Example-Submission Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Harrison FOL_Seq_Calc1 Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_SSA Formula_Derivatives Fourier Free-Boolean-Algebra Free-Groups FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Heard_Of Hello_World HereditarilyFinite Hermite Hidden_Markov_Models Higher_Order_Terms Hoare_Time HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL IEEE_Floating_Point IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interval_Arithmetic_Word32 Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Of_Medians_Selection Menger Mersenne_Primes MiniML Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_Assembly_Kit_Security Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Myhill-Nerode Name_Carrying_Type_Inference Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PCF PLM POPLmark-deBruijn PSemigroupsConvolution Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Program-Conflict-Analysis Projective_Geometry Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem QHLProver QR_Decomposition Quantales Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Regular-Sets Regular_Algebras Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rep_Fin_Groups Residuated_Lattices Resolution_FOL Rewriting_Z Ribbon_Proofs Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Routing Roy_Floyd_Warshall SATSolverVerification SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sort_Encodings Source_Coding_Theorem Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares SuperCalc Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God UPF UPF_Firewall UTP Universal_Turing_Machine UpDown_Scheme Valuation VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Well_Quasi_Orders Winding_Number_Eval Word_Lib WorkerWrapper XML ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL