Import of UPF release afp-UPF-2015-05-27 (Isabelle 2015).

This commit is contained in:
Achim D. Brucker 2016-08-10 10:27:32 +01:00
parent 38468e46bb
commit 9686a2dee3
12 changed files with 484 additions and 498 deletions

View File

@ -42,7 +42,7 @@
(* $Id: Analysis.thy 10953 2014-11-24 11:23:40Z wolff $ *)
header{* Properties on Policies *}
section{* Properties on Policies *}
theory
Analysis
imports
@ -54,14 +54,14 @@ text {*
In this theory, several standard policy properties are paraphrased in UPF terms.
*}
section{* Basic Properties *}
subsection{* Basic Properties *}
subsection{* A Policy Has no Gaps *}
subsubsection{* A Policy Has no Gaps *}
definition gap_free :: "('a \<mapsto> 'b) \<Rightarrow> bool"
where "gap_free p = (dom p = UNIV)"
subsection{* Comparing Policies *}
subsubsection{* Comparing Policies *}
text {* Policy p is more defined than q: *}
definition more_defined :: "('a \<mapsto> 'b) \<Rightarrow>('a \<mapsto> 'b) \<Rightarrow>bool"
where "more_defined p q = (dom q \<subseteq> dom p)"
@ -128,7 +128,7 @@ lemma "A\<^sub>I \<sqsubseteq>\<^sub>A p"
unfolding more_permissive_def allow_all_fun_def allow_pfun_def allow_all_id_def
by(auto split: option.split decision.split)
section{* Combined Data-Policy Refinement *}
subsection{* Combined Data-Policy Refinement *}
definition policy_refinement ::
"('a \<mapsto> 'b) \<Rightarrow> ('a' \<Rightarrow> 'a) \<Rightarrow>('b' \<Rightarrow> 'b) \<Rightarrow> ('a' \<mapsto> 'b') \<Rightarrow> bool"
@ -163,7 +163,7 @@ apply(erule_tac x=" (f' a')" in allE, auto)
done
section {* Equivalence of Policies *}
subsection {* Equivalence of Policies *}
subsubsection{* Equivalence over domain D *}
definition p_eq_dom :: "('a \<mapsto> 'b) \<Rightarrow> 'a set \<Rightarrow> ('a \<mapsto> 'b) \<Rightarrow>bool" ("_ \<approx>\<^bsub>_\<^esub> _" [60,60,60]60)
@ -190,7 +190,7 @@ shows "no_conflicts p q"
apply (metis)+
done
subsection{* Miscellaneous *}
subsubsection{* Miscellaneous *}
lemma dom_inter: "\<lbrakk>dom p \<inter> dom q = {}; p x = \<lfloor>y\<rfloor>\<rbrakk> \<Longrightarrow> q x = \<bottom>"
by (auto)

View File

@ -41,7 +41,7 @@
******************************************************************************)
(* $Id: ElementaryPolicies.thy 10945 2014-11-21 12:50:43Z wolff $ *)
header{* Elementary Policies *}
section{* Elementary Policies *}
theory
ElementaryPolicies
imports
@ -54,7 +54,7 @@ text{*
policies defined in this theory.
*}
section{* The Core Policy Combinators: Allow and Deny Everything *}
subsection{* The Core Policy Combinators: Allow and Deny Everything *}
definition
deny_pfun :: "('\<alpha> \<rightharpoonup>'\<beta>) \<Rightarrow> ('\<alpha> \<mapsto> '\<beta>)" ("AllD")
@ -111,7 +111,7 @@ lemma neq_Allow_Deny: "pf \<noteq> \<emptyset> \<Longrightarrow> (deny_pfun pf)
apply (auto)
done
section{* Common Instances *}
subsection{* Common Instances *}
definition allow_all_fun :: "('\<alpha> \<Rightarrow> '\<beta>) \<Rightarrow> ('\<alpha> \<mapsto> '\<beta>)" ("A\<^sub>f")
where "allow_all_fun f = allow_pfun (Some o f)"
@ -161,7 +161,7 @@ lemma deny_left_cancel :"dom pf = UNIV \<Longrightarrow> (deny_pfun pf) \<Oplus>
apply (auto simp: deny_pfun_def option.splits)
done
section{* Domain, Range, and Restrictions *}
subsection{* Domain, Range, and Restrictions *}
text{*
Since policies are essentially maps, we inherit the basic definitions for
@ -196,6 +196,7 @@ lemma sub_ran : "ran p \<subseteq> Allow \<union> Deny"
apply (auto simp: Allow_def Deny_def ran_def full_SetCompr_eq[symmetric])
apply (case_tac "x")
apply (simp_all)
apply (rename_tac \<alpha>)
apply (erule_tac x="\<alpha>" in allE)
apply (simp)
done

View File

@ -41,14 +41,14 @@
******************************************************************************)
(* $Id: Monads.thy 10922 2014-11-10 15:41:49Z wolff $ *)
header {* Basic Monad Theory for Sequential Computations *}
section {* Basic Monad Theory for Sequential Computations *}
theory
Monads
imports
Main
begin
section{* General Framework for Monad-based Sequence-Test *}
subsection{* General Framework for Monad-based Sequence-Test *}
text{*
As such, Higher-order Logic as a purely functional specification formalism has no built-in
mechanism for state and state-transitions. Forms of testing involving state require therefore
@ -69,7 +69,7 @@ text{*
\end{enumerate}
*}
subsection{* State Exception Monads *}
subsubsection{* State Exception Monads *}
type_synonym ('o, '\<sigma>) MON\<^sub>S\<^sub>E = "'\<sigma> \<rightharpoonup> ('o \<times> '\<sigma>)"
definition bind_SE :: "('o,'\<sigma>)MON\<^sub>S\<^sub>E \<Rightarrow> ('o \<Rightarrow> ('o','\<sigma>)MON\<^sub>S\<^sub>E) \<Rightarrow> ('o','\<sigma>)MON\<^sub>S\<^sub>E"
@ -262,7 +262,7 @@ by(simp add: malt_SE_def)
lemma malt_SE_cons [simp]: "\<Sqinter>\<^sub>S\<^sub>E (a # S) = (a \<sqinter>\<^sub>S\<^sub>E (\<Sqinter>\<^sub>S\<^sub>E S))"
by(simp add: malt_SE_def)
subsection{* State-Backtrack Monads *}
subsubsection{* State-Backtrack Monads *}
text{*This subsection is still rudimentary and as such an interesting
formal analogue to the previous monad definitions. It is doubtful that it is
interesting for testing and as a computational structure at all.
@ -301,7 +301,7 @@ lemma bind_assoc_SB: "(y := (x := m; k); h) = (x := m; (y := k; h))"
apply (simp add: unit_SB_def bind_SB_def split_def)
done
subsection{* State Backtrack Exception Monad *}
subsubsection{* State Backtrack Exception Monad *}
text{*
The following combination of the previous two Monad-Constructions allows for the semantic
foundation of a simple generic assertion language in the style of Schirmer's Simpl-Language or
@ -370,7 +370,7 @@ next
by(rule_tac x="(a, b)" in bexI, simp_all)
next
case goal4 then show ?case
apply (erule_tac Q="None = ?X" in contrapos_pp)
apply (erule_tac Q="None = X" for X in contrapos_pp)
apply (erule_tac x="(aa,b)" and P="\<lambda> x. None \<noteq> split (\<lambda>out. k) x" in ballE)
apply (auto simp: aux (*Option.not_None_eq*) image_def split_def intro!: rev_bexI)
done
@ -390,7 +390,7 @@ next
qed
section{* Valid Test Sequences in the State Exception Monad *}
subsection{* Valid Test Sequences in the State Exception Monad *}
text{*
This is still an unstructured merge of executable monad concepts and specification oriented
high-level properties initiating test procedures.
@ -553,7 +553,7 @@ lemma [code]: "(\<sigma> \<Turnstile> m) = (case (m \<sigma>) of None \<Rightar
apply (auto)
done
section{* Valid Test Sequences in the State Exception Backtrack Monad *}
subsection{* Valid Test Sequences in the State Exception Backtrack Monad *}
text{*
This is still an unstructured merge of executable monad concepts and specification oriented
high-level properties initiating test procedures.

View File

@ -40,7 +40,7 @@
******************************************************************************)
(* $Id: Normalisation.thy 10879 2014-10-26 11:35:31Z brucker $ *)
header{* Policy Transformations *}
section{* Policy Transformations *}
theory
Normalisation
imports
@ -54,7 +54,7 @@ text{*
study~\cite{brucker.ea:formal-fw-testing:2014}.
*}
section{* Elementary Operators *}
subsection{* Elementary Operators *}
text{*
We start by providing several operators and theorems useful when reasoning about a list of
rules which should eventually be interpreted as combined using the standard override operator.
@ -236,7 +236,7 @@ proof (induct p rule: rev_induct)
qed
section{* Distributivity of the Transformation. *}
subsection{* Distributivity of the Transformation. *}
text{*
The scenario is the following (can be applied iteratively):
\begin{itemize}

View File

@ -40,7 +40,7 @@
******************************************************************************)
(* $Id: NormalisationTestSpecification.thy 10879 2014-10-26 11:35:31Z brucker $ *)
header {* Policy Transformation for Testing *}
section {* Policy Transformation for Testing *}
theory
NormalisationTestSpecification
imports

View File

@ -40,7 +40,7 @@
******************************************************************************)
(* $Id: ParallelComposition.thy 10879 2014-10-26 11:35:31Z brucker $ *)
header{* Parallel Composition*}
section{* Parallel Composition*}
theory
ParallelComposition
imports
@ -62,7 +62,7 @@ text{*
flattening.
*}
section{* Parallel Combinators: Foundations *}
subsection{* Parallel Combinators: Foundations *}
text {*
There are four possible semantics how the decision can be combined, thus there are four
parallel composition operators. For each of them, we prove several properties.
@ -212,7 +212,7 @@ lemma mt_prod_2_id[simp]:"\<emptyset> \<Otimes>\<^sub>2\<^sub>I p = \<emptyset>"
apply (simp add: prod_2_id_def prod_2_def)
done
section{* Combinators for Transition Policies *}
subsection{* Combinators for Transition Policies *}
text {*
For constructing transition policies, two additional combinators are required: one combines
state transitions by pairing the states, the other works equivalently on general maps.
@ -231,7 +231,7 @@ where
"p1 \<Otimes>\<^sub>S p2 = (p1 \<Otimes>\<^sub>M p2) o (\<lambda> (a,b,c). ((a,b),a,c))"
section{* Range Splitting *}
subsection{* Range Splitting *}
text{*
The following combinator is a special case of both a parallel composition operator and a
range splitting operator. Its primary use case is when combining a policy with state transitions.
@ -257,7 +257,7 @@ lemma comp_ran_split_charn:
apply (auto)
done
section {* Distributivity of the parallel combinators *}
subsection {* Distributivity of the parallel combinators *}
lemma distr_or1_a: "(F = F1 \<Oplus> F2) \<Longrightarrow> (((N \<Otimes>\<^sub>1 F) o f) =
(((N \<Otimes>\<^sub>1 F1) o f) \<Oplus> ((N \<Otimes>\<^sub>1 F2) o f))) "

View File

@ -40,7 +40,7 @@
******************************************************************************)
(* $Id: SeqComposition.thy 10879 2014-10-26 11:35:31Z brucker $ *)
header{* Sequential Composition *}
section{* Sequential Composition *}
theory
SeqComposition
imports
@ -52,7 +52,7 @@ text{*
the second policy to the output of the first one. Again, there are four possibilities how the
decisions can be combined. *}
section {* Flattening *}
subsection {* Flattening *}
text{*
A key concept of sequential policy composition is the flattening of nested decisions. There are
four possibilities, and these possibilities will give the various flavours of policy composition.
@ -65,22 +65,20 @@ where "flat_orA(allow(allow y)) = allow y"
lemma flat_orA_deny[dest]:"flat_orA x = deny y \<Longrightarrow> x = deny(deny y)"
apply (case_tac "x")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
done
lemma flat_orA_allow[dest]: "flat_orA x = allow y \<Longrightarrow> x = allow(allow y)
\<or> x = allow(deny y)
\<or> x = deny(allow y)"
apply (case_tac "x")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
done
fun flat_orD :: "('\<alpha> decision) decision \<Rightarrow> ('\<alpha> decision)"
@ -91,22 +89,20 @@ where "flat_orD(allow(allow y)) = allow y"
lemma flat_orD_allow[dest]: "flat_orD x = allow y \<Longrightarrow> x = allow(allow y)"
apply (case_tac "x")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
done
lemma flat_orD_deny[dest]: "flat_orD x = deny y \<Longrightarrow> x = deny(deny y)
\<or> x = allow(deny y)
\<or> x = deny(allow y)"
apply (case_tac "x")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
done
fun flat_1 :: "('\<alpha> decision) decision \<Rightarrow> ('\<alpha> decision)"
@ -117,20 +113,18 @@ where "flat_1(allow(allow y)) = allow y"
lemma flat_1_allow[dest]: "flat_1 x = allow y \<Longrightarrow> x = allow(allow y) \<or> x = allow(deny y)"
apply (case_tac "x")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
done
lemma flat_1_deny[dest]: "flat_1 x = deny y \<Longrightarrow> x = deny(deny y) \<or> x = deny(allow y)"
apply (case_tac "x")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
done
fun flat_2 :: "('\<alpha> decision) decision \<Rightarrow> ('\<alpha> decision)"
@ -141,23 +135,21 @@ where "flat_2(allow(allow y)) = allow y"
lemma flat_2_allow[dest]: "flat_2 x = allow y \<Longrightarrow> x = allow(allow y) \<or> x = deny(allow y)"
apply (case_tac "x")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
done
lemma flat_2_deny[dest]: "flat_2 x = deny y \<Longrightarrow> x = deny(deny y) \<or> x = allow(deny y)"
apply (case_tac "x")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (case_tac "\<alpha>")
apply (simp_all)
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
apply (rename_tac \<alpha>)
apply (case_tac "\<alpha>", simp_all)[1]
done
section{* Policy Composition *}
subsection{* Policy Composition *}
text{*
The following definition allows to compose two policies. Denies and allows are transferred.
*}

View File

@ -40,7 +40,7 @@
******************************************************************************)
(* $Id: Service.thy 10945 2014-11-21 12:50:43Z wolff $ *)
header {* Secure Service Specification *}
section {* Secure Service Specification *}
theory
Service
imports
@ -51,8 +51,8 @@ text {*
that allows the staff in a hospital to access health care records of patients.
*}
section{* Datatypes for Modelling Users and Roles*}
subsection {* Users *}
subsection{* Datatypes for Modelling Users and Roles*}
subsubsection {* Users *}
text{*
First, we introduce a type for users that we use to model that each
staff member has a unique id:
@ -64,7 +64,7 @@ text {*
*}
type_synonym patient = int (* Each patient gets a unique id *)
subsection {* Roles and Relationships*}
subsubsection {* Roles and Relationships*}
text{* In our example, we assume three different roles for members of the clinical staff: *}
datatype role = ClinicalPractitioner | Nurse | Clerical
@ -83,8 +83,8 @@ type_synonym \<Sigma> = "patient \<rightharpoonup> LR"
text{* The user context stores the roles the users are in. *}
type_synonym \<upsilon> = "user \<rightharpoonup> role"
section {* Modelling Health Records and the Web Service API*}
subsection {* Health Records *}
subsection {* Modelling Health Records and the Web Service API*}
subsubsection {* Health Records *}
text {* The content and the status of the entries of a health record *}
datatype data = dummyContent
datatype status = Open | Closed
@ -93,7 +93,7 @@ type_synonym entry = "status \<times> user \<times> data"
type_synonym SCR = "(entry_id \<rightharpoonup> entry)"
type_synonym DB = "patient \<rightharpoonup> SCR"
subsection {* The Web Service API *}
subsubsection {* The Web Service API *}
text{* The operations provided by the service: *}
datatype Operation = createSCR user role patient
| appendEntry user role patient entry_id entry
@ -207,7 +207,7 @@ fun allContentStatic where
|"allContentStatic [] = True"
section{* Modelling Access Control*}
subsection{* Modelling Access Control*}
text {*
In the following, we define a rather complex access control model for our
scenario that extends traditional role-based access control
@ -217,7 +217,7 @@ text {*
for a general motivation and explanation of break-the-glass access control).
*}
subsection {* Sealed Envelopes *}
subsubsection {* Sealed Envelopes *}
type_synonym SEPolicy = "(Operation \<times> DB \<mapsto> unit) "
@ -259,7 +259,7 @@ definition SEPolicy :: SEPolicy where
lemmas SEsimps = SEPolicy_def get_entry_def userHasAccess_def
editEntrySE_def deleteEntrySE_def readEntrySE_def
subsection {* Legitimate Relationships *}
subsubsection {* Legitimate Relationships *}
type_synonym LRPolicy = "(Operation \<times> \<Sigma>, unit) policy"
@ -365,7 +365,7 @@ definition FunPolicy where
removeLRFunPolicy \<Oplus> readSCRFunPolicy \<Oplus>
addLRFunPolicy \<Oplus> createFunPolicy \<Oplus> A\<^sub>U"
subsection{* Modelling Core RBAC *}
subsubsection{* Modelling Core RBAC *}
type_synonym RBACPolicy = "Operation \<times> \<upsilon> \<mapsto> unit"
@ -389,9 +389,9 @@ definition RBACPolicy :: RBACPolicy where
then \<lfloor>allow ()\<rfloor>
else \<lfloor>deny ()\<rfloor>)"
section {* The State Transitions and Output Function*}
subsection {* The State Transitions and Output Function*}
subsection{* State Transition *}
subsubsection{* State Transition *}
fun OpSuccessDB :: "(Operation \<times> DB) \<rightharpoonup> DB" where
"OpSuccessDB (createSCR u r p,S) = (case S p of \<bottom> \<Rightarrow> \<lfloor>S(p\<mapsto>\<emptyset>)\<rfloor>
@ -434,7 +434,7 @@ fun OpSuccessSigma :: "(Operation \<times> \<Sigma>) \<rightharpoonup> \<Sigma>"
fun OpSuccessUC :: "(Operation \<times> \<upsilon>) \<rightharpoonup> \<upsilon>" where
"OpSuccessUC (f,u) = \<lfloor>u\<rfloor>"
subsection {* Output *}
subsubsection {* Output *}
type_synonym Output = unit
@ -445,7 +445,7 @@ fun OpSuccessOutput :: "(Operation) \<rightharpoonup> Output" where
fun OpFailOutput :: "Operation \<rightharpoonup> Output" where
"OpFailOutput x = \<lfloor>()\<rfloor>"
section {* Combine All Parts *}
subsection {* Combine All Parts *}
definition SE_LR_Policy :: "(Operation \<times> DB \<times> \<Sigma>, unit) policy" where
"SE_LR_Policy = (\<lambda>(x,x). x) o\<^sub>f (SEPolicy \<Otimes>\<^sub>\<or>\<^sub>D LR_Policy) o (\<lambda>(a,b,c). ((a,b),a,c))"

View File

@ -40,7 +40,7 @@
******************************************************************************)
(* $Id: ServiceExample.thy 10954 2014-11-24 12:43:29Z wolff $ *)
header {* Instantiating Our Secure Service Example *}
section {* Instantiating Our Secure Service Example *}
theory
ServiceExample
imports
@ -52,7 +52,7 @@ text {*
two patients:
*}
section {* Access Control Configuration *}
subsection {* Access Control Configuration *}
definition alice :: user where "alice = 1"
definition bob :: user where "bob = 2"
definition charlie :: user where "charlie = 3"
@ -86,11 +86,11 @@ definition LR1 :: LR where
definition \<Sigma>0 :: \<Sigma> where
"\<Sigma>0 = (empty(patient1\<mapsto>LR1))"
section {* The Initial System State *}
subsection {* The Initial System State *}
definition \<sigma>0 :: "DB \<times> \<Sigma>\<times>\<upsilon>" where
"\<sigma>0 = (Spine0,\<Sigma>0,UC0)"
section{* Basic Properties *}
subsection{* Basic Properties *}
lemma [simp]: "(case a of allow d \<Rightarrow> \<lfloor>X\<rfloor> | deny d2 \<Rightarrow> \<lfloor>Y\<rfloor>) = \<bottom> \<Longrightarrow> False"
by (case_tac a,simp_all)

View File

@ -41,7 +41,7 @@
******************************************************************************)
(* $Id: UPF.thy 10879 2014-10-26 11:35:31Z brucker $ *)
header {* Putting Everything Together: UPF *}
section {* Putting Everything Together: UPF *}
theory
UPF
imports

View File

@ -41,7 +41,7 @@
******************************************************************************)
(* $Id: UPFCore.thy 10951 2014-11-21 21:54:46Z wolff $ *)
header{* The Core of the Unified Policy Framework (UPF) *}
section{* The Core of the Unified Policy Framework (UPF) *}
theory
UPFCore
imports
@ -49,7 +49,7 @@ imports
begin
section{* Foundation *}
subsection{* Foundation *}
text{*
The purpose of this theory is to formalize a somewhat non-standard view
on the fundamental concept of a security policy which is worth outlining.
@ -112,7 +112,7 @@ definition Deny :: "('\<alpha> decision) set"
where "Deny = range deny"
section{* Policy Constructors *}
subsection{* Policy Constructors *}
text{*
Most elementary policy constructors are based on the
update operation @{thm [source] "Fun.fun_upd_def"} @{thm Fun.fun_upd_def}
@ -172,7 +172,7 @@ lemma pol_upd_neq1 [simp]: "m(a\<mapsto>\<^sub>+x) \<noteq> n(a\<mapsto>\<^sub>-
by(auto dest: map_upd_eqD1)
section{* Override Operators *}
subsection{* Override Operators *}
text{*
Key operators for constructing policies are the override operators. There are four different
versions of them, with one of them being the override operator from the Map theory. As it is
@ -269,7 +269,7 @@ lemma override_D_assoc: "p1 \<Oplus>\<^sub>D (p2 \<Oplus>\<^sub>D p3) = (p1 \<Op
apply (simp add: override_D_def split: decision.splits option.splits)
done
section{* Coercion Operators *}
subsection{* Coercion Operators *}
text{*
Often, especially when combining policies of different type, it is necessary to
adapt the input or output domain of a policy to a more refined context.

View File

@ -36,13 +36,6 @@
\urlstyle{rm}
\isabellestyle{it}
\renewcommand{\isamarkupheader}[1]{\section{#1}}
\renewcommand{\isamarkupsection}[1]{\subsection{#1}}
\renewcommand{\isamarkupsubsection}[1]{\subsubsection{#1}}
\renewcommand{\isamarkupsubsubsection}[1]{\paragraph{#1}}
\renewcommand{\isamarkupsect}[1]{\subsection{#1}}
\renewcommand{\isamarkupsubsect}[1]{\susubsection{#1}}
\renewcommand{\isamarkupsubsubsect}[1]{\paragraph{#1}}
\renewcommand{\isastyle}{\isastyleminor}
\pagestyle{empty}