Import of official AFP entry for Isabelle 2019.
afp-mirror/UPF_Firewall/master This commit looks good Details

This commit is contained in:
Achim D. Brucker 2019-06-22 23:31:30 +01:00
parent fe59480469
commit ea91410aec
47 changed files with 315 additions and 315 deletions

2
.ci/Jenkinsfile vendored
View File

@ -3,7 +3,7 @@ pipeline {
stages {
stage('Build') {
steps {
sh 'docker run -v $PWD/UPF_Firewall:/UPF_Firewall logicalhacking:isabelle2018 isabelle build -D /UPF_Firewall'
sh 'docker run -v $PWD/UPF_Firewall:/UPF_Firewall logicalhacking:isabelle2019 isabelle build -D /UPF_Firewall'
}
}
}

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
section {* A Simple DMZ Setup *}
section \<open>A Simple DMZ Setup\<close>
theory
DMZ
imports

View File

@ -35,14 +35,14 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* DMZ Datatype *}
subsection \<open>DMZ Datatype\<close>
theory
DMZDatatype
imports
"../../UPF-Firewall"
begin
text{*
text\<open>
This is the fourth scenario, slightly more complicated than the previous one, as we now also
model specific servers within one network. Therefore, we could not use anymore the modelling
using datatype synonym, but only use the one where an address is modelled as an
@ -51,7 +51,7 @@ text{*
Just for comparison, this theory is the same scenario with datatype synonym anyway, but with four
distinct networks instead of one contained in another. As there is no corresponding network model
included, we need to define a custom one.
*}
\<close>
datatype Adr = Intranet | Internet | Mail | Web | DMZ
instance Adr::adr ..
@ -98,10 +98,10 @@ definition
Internet_mail_port ++
Internet_web_port"
text {*
text \<open>
We only want to create test cases which are sent between the three main networks: e.g. not
between the mailserver and the dmz. Therefore, the constraint looks as follows. \
*}
\<close>
definition
not_in_same_net :: "(Networks,DummyContent) packet \<Rightarrow> bool" where

View File

@ -35,14 +35,14 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* DMZ: Integer *}
subsection \<open>DMZ: Integer\<close>
theory
DMZInteger
imports
"../../UPF-Firewall"
begin
text{* This scenario is slightly more complicated than the SimpleDMZ
text\<open>This scenario is slightly more complicated than the SimpleDMZ
one, as we now also model specific servers within one
network. Therefore, we cannot use anymore the modelling using
datatype synonym, but only use the one where an address is modelled as an
@ -69,7 +69,7 @@ The scenario is the following:
\end{itemize}
\end{labeling}
*}
\<close>
definition
intranet::"adr\<^sub>i\<^sub>p net" where
@ -111,10 +111,10 @@ definition
Internet_mail_port ++
Internet_web_port"
text {*
text \<open>
We only want to create test cases which are sent between the three main networks:
e.g. not between the mailserver and the dmz. Therefore, the constraint looks as follows.
*}
\<close>
definition
not_in_same_net :: "(adr\<^sub>i\<^sub>p,DummyContent) packet \<Rightarrow> bool" where

View File

@ -35,12 +35,12 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
chapter {* Examples *}
chapter \<open>Examples\<close>
theory
Examples
imports
"DMZ/DMZ"
"VoIP/VoIP"
"Voice_over_IP/Voice_over_IP"
"Transformation/Transformation"
"NAT-FW/NAT-FW"
"PersonalFirewall/PersonalFirewall"

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
section {* Example: NAT *}
section \<open>Example: NAT\<close>
theory
"NAT-FW"
imports

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
section {* Personal Firewall *}
section \<open>Personal Firewall\<close>
theory
PersonalFirewall
imports

View File

@ -35,20 +35,20 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Personal Firewall: Datatype *}
subsection \<open>Personal Firewall: Datatype\<close>
theory
PersonalFirewallDatatype
imports
"../../UPF-Firewall"
begin
text{*
text\<open>
The most basic firewall scenario; there is a personal PC on one side and the Internet on the
other. There are two policies: the first one allows all traffic from the PC to the Internet and
denies all coming into the PC. The second policy only allows specific ports from the PC. This
scenario comes in three variants: the first one specifies the allowed protocols directly, the
second together with their respective port numbers, the third one only with the port numbers.
*}
\<close>
datatype Adr = pc | internet
@ -68,12 +68,12 @@ definition
not_in_same_net :: "(DatatypeTwoNets,DummyContent) packet \<Rightarrow> bool" where
"not_in_same_net x = ((src x \<sqsubset> PC \<longrightarrow> dest x \<sqsubset> Internet) \<and> (src x \<sqsubset> Internet \<longrightarrow> dest x \<sqsubset> PC))"
text {*
text \<open>
Definitions of the policies
In fact, the short definitions wouldn't have to be written down - they
are the automatically simplified versions of their big counterparts.
*}
\<close>
definition
strictPolicy :: "(DatatypeTwoNets,DummyContent) FWPolicy" where

View File

@ -35,24 +35,24 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection{* Personal Firewall: Integer *}
subsection\<open>Personal Firewall: Integer\<close>
theory
PersonalFirewallInt
imports
"../../UPF-Firewall"
begin
text{*
text\<open>
The most basic firewall scenario; there is a personal PC on one side and the Internet on the
other. There are two policies: the first one allows all traffic from the PC to the Internet and
denies all coming into the PC. The second policy only allows specific ports from the PC. This
scenario comes in three variants: the first one specifies the allowed protocols directly, the
second together with their respective port numbers, the third one only with the port numbers.
*}
\<close>
text{*
text\<open>
Definitions of the subnets
*}
\<close>
definition
PC :: "(adr\<^sub>i\<^sub>p net)" where
@ -66,9 +66,9 @@ definition
not_in_same_net :: "(adr\<^sub>i\<^sub>p,DummyContent) packet \<Rightarrow> bool" where
"not_in_same_net x = ((src x \<sqsubset> PC \<longrightarrow> dest x \<sqsubset> Internet) \<and> (src x \<sqsubset> Internet \<longrightarrow> dest x \<sqsubset> PC))"
text {*
text \<open>
Definitions of the policies
*}
\<close>
definition
strictPolicy :: "(adr\<^sub>i\<^sub>p,DummyContent) FWPolicy" where

View File

@ -35,24 +35,24 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Personal Firewall IPv4 *}
subsection \<open>Personal Firewall IPv4\<close>
theory
PersonalFirewallIpv4
imports
"../../UPF-Firewall"
begin
text{*
text\<open>
The most basic firewall scenario; there is a personal PC on one side and the Internet on the
other. There are two policies: the first one allows all traffic from the PC to the Internet and
denies all coming into the PC. The second policy only allows specific ports from the PC. This
scenario comes in three variants: the first one specifies the allowed protocols directly, the
second together with their respective port numbers, the third one only with the port numbers.
*}
\<close>
text{*
text\<open>
Definitions of the subnets
*}
\<close>
definition
PC :: "(ipv4 net)" where
@ -66,9 +66,9 @@ definition
not_in_same_net :: "(ipv4,DummyContent) packet \<Rightarrow> bool" where
"not_in_same_net x = ((src x \<sqsubset> PC \<longrightarrow> dest x \<sqsubset> Internet) \<and> (src x \<sqsubset> Internet \<longrightarrow> dest x \<sqsubset> PC))"
text {*
text \<open>
Definitions of the policies
*}
\<close>
definition
strictPolicy :: "(ipv4,DummyContent) FWPolicy" where
"strictPolicy = deny_all ++ allow_all_from_to PC Internet"

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
section {* Demonstrating Policy Transformations *}
section \<open>Demonstrating Policy Transformations\<close>
theory
Transformation
imports

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Transformation Example 1 *}
subsection \<open>Transformation Example 1\<close>
theory
Transformation01
imports

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Transforamtion Example 2 *}
subsection \<open>Transforamtion Example 2\<close>
theory
Transformation02
imports

View File

@ -35,20 +35,20 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
section {* Voice over IP *}
section \<open>Voice over IP\<close>
theory
VoIP
Voice_over_IP
imports
"../../UPF-Firewall"
begin
text{*
text\<open>
In this theory we generate the test data for correct runs of the FTP protocol. As usual, we
start with definining the networks and the policy. We use a rather simple policy which allows
only FTP connections starting from the Intranet and going to the Internet, and deny everything
else.
*}
\<close>
definition
intranet :: "adr\<^sub>i\<^sub>p net" where
@ -66,9 +66,9 @@ definition
voip_policy :: "(adr\<^sub>i\<^sub>p,address voip_msg) FWPolicy" where
"voip_policy = A\<^sub>U"
text{*
text\<open>
The next two constants check if an address is in the Intranet or in the Internet respectively.
*}
\<close>
definition
is_in_intranet :: "address \<Rightarrow> bool" where
@ -82,9 +82,9 @@ definition
is_in_internet :: "address \<Rightarrow> bool" where
"is_in_internet a = (a > 4)"
text{*
text\<open>
The next definition is our starting state: an empty trace and the just defined policy.
*}
\<close>
definition
"\<sigma>_0_voip" :: "(adr\<^sub>i\<^sub>p, address voip_msg) history \<times>
@ -92,10 +92,10 @@ definition
where
"\<sigma>_0_voip = ([],voip_policy)"
text{*
text\<open>
Next we state the conditions we have on our trace: a normal behaviour FTP run from the intranet
to some server in the internet on port 21.
*}
\<close>
definition "accept_voip" :: "(adr\<^sub>i\<^sub>p, address voip_msg) history \<Rightarrow> bool" where
"accept_voip t = (\<exists> c s g i p1 p2. t \<in> NB_voip c s g i p1 p2 \<and> is_in_intranet c
@ -107,10 +107,10 @@ fun packet_with_id where
|"packet_with_id (x#xs) i =
(if id x = i then (x#(packet_with_id xs i)) else (packet_with_id xs i))"
text{*
text\<open>
The depth of the test case generation corresponds to the maximal length of generated traces,
4 is the minimum to get a full FTP protocol run.
*}
\<close>
fun ids1 where
"ids1 i (x#xs) = (id x = i \<and> ids1 i xs)"

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Elementary Firewall Policy Transformation Rules *}
subsection \<open>Elementary Firewall Policy Transformation Rules\<close>
theory
ElementaryRules
imports
@ -43,10 +43,10 @@ theory
begin
text{*
text\<open>
This theory contains those elementary transformation rules which are presented in the ICST
2010 paper~\cite{brucker.ea:firewall:2010}. They are not used elsewhere.
*}
\<close>
lemma elem1:
"C (AllowPortFromTo x y p \<oplus> DenyAllFromTo x y) = C (DenyAllFromTo x y)"

View File

@ -34,7 +34,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
section {* Firewall Policy Normalisation *}
section \<open>Firewall Policy Normalisation\<close>
theory
FWNormalisation
imports

View File

@ -35,14 +35,14 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Policy Normalisation: Core Definitions *}
subsection \<open>Policy Normalisation: Core Definitions\<close>
theory
FWNormalisationCore
imports
"../PacketFilter/PacketFilter"
begin
text{*
text\<open>
This theory contains all the definitions used for policy normalisation as described
in~\cite{brucker.ea:icst:2010,brucker.ea:formal-fw-testing:2014}.
@ -71,17 +71,17 @@ text{*
normalization procedure does not aim to minimize the number of rules. While it does remove
unnecessary ones, it also adds new ones, enabling a policy to be split into several independent
parts.
*}
\<close>
text{*
text\<open>
Policy transformations are functions that map policies to policies. We decided to represent
policy transformations as \emph{syntactic rules}; this choice paves the way for expressing the
entire normalisation process inside HOL by functions manipulating abstract policy syntax.
*}
\<close>
subsubsection{* Basics *}
text{* We define a very simple policy language: *}
subsubsection\<open>Basics\<close>
text\<open>We define a very simple policy language:\<close>
datatype ('\<alpha>,'\<beta>) Combinators =
DenyAll
@ -89,18 +89,18 @@ datatype ('\<alpha>,'\<beta>) Combinators =
| AllowPortFromTo '\<alpha> '\<alpha> '\<beta>
| Conc "(('\<alpha>,'\<beta>) Combinators)" "(('\<alpha>,'\<beta>) Combinators)" (infixr "\<oplus>" 80)
text{*
text\<open>
And define the semantic interpretation of it. For technical reasons, we fix here the type to
policies over IntegerPort addresses. However, we could easily provide definitions for other
address types as well, using a generic constants for the type definition and a primitive
recursive definition for each desired address model.
*}
\<close>
subsubsection{* Auxiliary definitions and functions. *}
text{*
subsubsection\<open>Auxiliary definitions and functions.\<close>
text\<open>
This section defines several functions which are useful later for the combinators, invariants,
and proofs.
*}
\<close>
fun srcNet where
"srcNet (DenyAllFromTo x y) = x"
@ -153,10 +153,10 @@ fun bothNet where
|"bothNet (AllowPortFromTo a b p) = {a,b}"
|"bothNet (v \<oplus> va) = undefined "
text{*
text\<open>
$Nets\_List$ provides from a list of rules a list where the entries are the appearing sets of
source and destination network of each rule.
*}
\<close>
definition Nets_List
where
@ -209,21 +209,21 @@ fun firstList where
"firstList (x#xs) = (first_bothNet x)"
|"firstList [] = {}"
subsubsection{* Invariants *}
subsubsection\<open>Invariants\<close>
text{* If there is a DenyAll, it is at the first position *}
text\<open>If there is a DenyAll, it is at the first position\<close>
fun wellformed_policy1:: "(('\<alpha>, '\<beta>) Combinators) list \<Rightarrow> bool" where
"wellformed_policy1 [] = True"
| "wellformed_policy1 (x#xs) = (DenyAll \<notin> (set xs))"
text{* There is a DenyAll at the first position *}
text\<open>There is a DenyAll at the first position\<close>
fun wellformed_policy1_strong:: "(('\<alpha>, '\<beta>) Combinators) list \<Rightarrow> bool"
where
"wellformed_policy1_strong [] = False"
| "wellformed_policy1_strong (x#xs) = (x=DenyAll \<and> (DenyAll \<notin> (set xs)))"
text{* All two networks are either disjoint or equal. *}
text\<open>All two networks are either disjoint or equal.\<close>
definition netsDistinct where "netsDistinct a b = (\<not> (\<exists> x. x \<sqsubset> a \<and> x \<sqsubset> b))"
definition twoNetsDistinct where
@ -237,7 +237,7 @@ definition disjSD_2 where
"disjSD_2 x y = (\<forall> a b c d. ((a,b)\<in>sdnets x \<and> (c,d) \<in>sdnets y \<longrightarrow>
(twoNetsDistinct a b c d \<and> twoNetsDistinct a b d c)))"
text{* The policy is given as a list of single rules. *}
text\<open>The policy is given as a list of single rules.\<close>
fun singleCombinators where
"singleCombinators [] = True"
|"singleCombinators ((x\<oplus>y)#xs) = False"
@ -246,7 +246,7 @@ fun singleCombinators where
definition onlyTwoNets where
"onlyTwoNets x = ((\<exists> a b. (sdnets x = {(a,b)})) \<or> (\<exists> a b. sdnets x = {(a,b),(b,a)}))"
text{* Each entry of the list contains rules between two networks only. *}
text\<open>Each entry of the list contains rules between two networks only.\<close>
fun OnlyTwoNets where
"OnlyTwoNets (DenyAll#xs) = OnlyTwoNets xs"
|"OnlyTwoNets (x#xs) = (onlyTwoNets x \<and> OnlyTwoNets xs)"
@ -274,12 +274,12 @@ fun NetsCollected2 where
NetsCollected2 xs))"
|"NetsCollected2 [] = True"
subsubsection{* Transformations *}
subsubsection\<open>Transformations\<close>
text {*
text \<open>
The following two functions transform a policy into a list of single rules and vice-versa (by
staying on the combinator level).
*}
\<close>
fun policy2list::"('\<alpha>, '\<beta>) Combinators \<Rightarrow>
(('\<alpha>, '\<beta>) Combinators) list" where
@ -292,7 +292,7 @@ fun list2FWpolicy::"(('\<alpha>, '\<beta>) Combinators) list \<Rightarrow>
|"list2FWpolicy (x#[]) = x"
|"list2FWpolicy (x#y) = x \<oplus> (list2FWpolicy y)"
text{* Remove all the rules appearing before a DenyAll. There are two alternative versions. *}
text\<open>Remove all the rules appearing before a DenyAll. There are two alternative versions.\<close>
fun removeShadowRules1 where
"removeShadowRules1 (x#xs) = (if (DenyAll \<in> set xs)
@ -311,9 +311,9 @@ definition removeShadowRules1_alternative where
"removeShadowRules1_alternative p =
rev (removeShadowRules1_alternative_rev (rev p))"
text{*
text\<open>
Remove all the rules which allow a port, but are shadowed by a deny between these subnets.
*}
\<close>
fun removeShadowRules2:: "(('\<alpha>, '\<beta>) Combinators) list \<Rightarrow>
(('\<alpha>, '\<beta>) Combinators) list"
@ -325,10 +325,10 @@ where
| "removeShadowRules2 (x#y) = x#(removeShadowRules2 y)"
| "removeShadowRules2 [] = []"
text{*
text\<open>
Sorting a policies: We first need to define an ordering on rules. This ordering depends
on the $Nets\_List$ of a policy.
*}
\<close>
fun smaller :: "('\<alpha>, '\<beta>) Combinators \<Rightarrow>
('\<alpha>, '\<beta>) Combinators \<Rightarrow>
@ -343,7 +343,7 @@ where
else
(position (bothNet x) l <= position (bothNet y) l)))"
text{* We provide two different sorting algorithms: Quick Sort (qsort) and Insertion Sort (sort) *}
text\<open>We provide two different sorting algorithms: Quick Sort (qsort) and Insertion Sort (sort)\<close>
fun qsort where
"qsort [] l = []"
@ -377,10 +377,10 @@ fun separate where
else (x#(separate(y#z))))"
|"separate x = x"
text {*
text \<open>
Insert the DenyAllFromTo rules, such that traffic between two networks can be tested
individually.
*}
\<close>
fun insertDenies where
"insertDenies (x#xs) = (case x of DenyAll \<Rightarrow> (DenyAll#(insertDenies xs))
@ -389,11 +389,11 @@ fun insertDenies where
(insertDenies xs))"
| "insertDenies [] = []"
text{*
text\<open>
Remove duplicate rules. This is especially necessary as insertDenies might have inserted
duplicate rules. The second function is supposed to work on a list of policies. Only
rules which are duplicated within the same policy are removed.
*}
\<close>
fun removeDuplicates where
@ -405,7 +405,7 @@ fun removeAllDuplicates where
"removeAllDuplicates (x#xs) = ((removeDuplicates (x))#(removeAllDuplicates xs))"
|"removeAllDuplicates x = x"
text {* Insert a DenyAll at the beginning of a policy. *}
text \<open>Insert a DenyAll at the beginning of a policy.\<close>
fun insertDeny where
"insertDeny (DenyAll#xs) = DenyAll#xs"
|"insertDeny xs = DenyAll#xs"
@ -423,12 +423,12 @@ fun list2policyR::"(('\<alpha>, '\<beta>) Combinators) list \<Rightarrow>
|"list2policyR [] = undefined "
text{*
text\<open>
We provide the definitions for two address representations.
*}
\<close>
subsubsection{* IntPort *}
subsubsection\<open>IntPort\<close>
fun C :: "(adr\<^sub>i\<^sub>p net, port) Combinators \<Rightarrow> (adr\<^sub>i\<^sub>p,DummyContent) packet \<mapsto> unit"
where
@ -455,20 +455,20 @@ lemma check: "rev (policy2list (rotatePolicy p)) = policy2list p"
by (simp_all)
text{*
text\<open>
All rules appearing at the left of a DenyAllFromTo, have disjunct domains from it
(except DenyAll).
*}
\<close>
fun (sequential) wellformed_policy2 where
"wellformed_policy2 [] = True"
| "wellformed_policy2 (DenyAll#xs) = wellformed_policy2 xs"
| "wellformed_policy2 (x#xs) = ((\<forall> c a b. c = DenyAllFromTo a b \<and> c \<in> set xs \<longrightarrow>
Map.dom (C x) \<inter> Map.dom (C c) = {}) \<and> wellformed_policy2 xs)"
text{*
text\<open>
An allow rule is disjunct with all rules appearing at the right of it. This invariant is not
necessary as it is a consequence from others, but facilitates some proofs.
*}
\<close>
fun (sequential) wellformed_policy3::"((adr\<^sub>i\<^sub>p net,port) Combinators) list \<Rightarrow> bool" where
"wellformed_policy3 [] = True"
@ -515,16 +515,16 @@ definition
(qsort (removeShadowRules2 (remdups ((rm_MT_rules C) (insertDeny
(removeShadowRules1 (policy2list p)))))) ((l)))))"
text{*
text\<open>
Of course, normalize is equal to normalize', the latter looks nicer though.
*}
\<close>
lemma "normalize = normalize'"
by (rule ext, simp add: normalize_def normalize'_def sort'_def)
declare C.simps [simp del]
subsubsection{* TCP\_UDP\_IntegerPort *}
subsubsection\<open>TCP\_UDP\_IntegerPort\<close>
fun Cp :: "(adr\<^sub>i\<^sub>p\<^sub>p net, protocol \<times> port) Combinators \<Rightarrow>
(adr\<^sub>i\<^sub>p\<^sub>p,DummyContent) packet \<mapsto> unit"
@ -543,20 +543,20 @@ where
|"Dp (AllowPortFromTo x y p) = Cp (AllowPortFromTo x y p)"
|"Dp (x \<oplus> y) = Cp (y \<oplus> x)"
text{*
text\<open>
All rules appearing at the left of a DenyAllFromTo, have disjunct domains from it
(except DenyAll).
*}
\<close>
fun (sequential) wellformed_policy2Pr where
"wellformed_policy2Pr [] = True"
| "wellformed_policy2Pr (DenyAll#xs) = wellformed_policy2Pr xs"
| "wellformed_policy2Pr (x#xs) = ((\<forall> c a b. c = DenyAllFromTo a b \<and> c \<in> set xs \<longrightarrow>
Map.dom (Cp x) \<inter> Map.dom (Cp c) = {}) \<and> wellformed_policy2Pr xs)"
text{*
text\<open>
An allow rule is disjunct with all rules appearing at the right of it. This invariant is not
necessary as it is a consequence from others, but facilitates some proofs.
*}
\<close>
fun (sequential) wellformed_policy3Pr::"((adr\<^sub>i\<^sub>p\<^sub>p net, protocol \<times> port) Combinators) list \<Rightarrow> bool" where
"wellformed_policy3Pr [] = True"
@ -605,17 +605,17 @@ definition
(qsort (removeShadowRules2 (remdups ((rm_MT_rules Cp) (insertDeny
(removeShadowRules1 (policy2list p)))))) ((l)))))"
text{*
text\<open>
Of course, normalize is equal to normalize', the latter looks nicer though.
*}
\<close>
lemma "normalizePr = normalizePr'"
by (rule ext, simp add: normalizePr_def normalizePr'_def sort'_def)
text{*
text\<open>
The following definition helps in creating the test specification for the individual parts
of a normalized policy.
*}
\<close>
definition makeFUTPr where
"makeFUTPr FUT p x n =
(packet_Nets x (fst (normBothNets (bothNets p)!n))

View File

@ -35,17 +35,17 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection{* Normalisation Proofs (Generic) *}
subsection\<open>Normalisation Proofs (Generic)\<close>
theory
NormalisationGenericProofs
imports
FWNormalisationCore
begin
text {*
text \<open>
This theory contains the generic proofs of the normalisation procedure, i.e. those which
are independent from the concrete semantical interpretation function.
*}
\<close>
lemma domNMT: "dom X \<noteq> {} \<Longrightarrow> X \<noteq> \<emptyset>"
by auto
@ -1631,7 +1631,7 @@ next
case DenyAll thus ?thesis using assms DenyAll DenyAllFromTo by simp
next
case (DenyAllFromTo e f) thus ?thesis using assms DenyAllFromTo
by (metis DenyAllFromTo `a = DenyAllFromTo c d` bothNet.simps(2) smalleraux2)
by (metis DenyAllFromTo \<open>a = DenyAllFromTo c d\<close> bothNet.simps(2) smalleraux2)
next
case (AllowPortFromTo e f g) thus ?thesis
using assms DenyAllFromTo AllowPortFromTo by simp (metis eq_imp_le pos_noteq)
@ -1647,7 +1647,7 @@ next
using assms by simp (metis AllowPortFromTo DenyAllFromTo bothNet.simps(3) smalleraux2a)
next
case (AllowPortFromTo e f g) thus ?thesis
using assms by(simp)(metis AllowPortFromTo `a = AllowPortFromTo c d p`
using assms by(simp)(metis AllowPortFromTo \<open>a = AllowPortFromTo c d p\<close>
bothNet.simps(3) smalleraux2c)
next
case (Conc e f) thus ?thesis using assms by simp
@ -1747,11 +1747,11 @@ next
case DenyAll thus ?thesis by simp
next
case (DenyAllFromTo e f) thus ?thesis using DenyAllFromTo
by (auto simp: eq_imp_le `a = DenyAllFromTo c d`)
by (auto simp: eq_imp_le \<open>a = DenyAllFromTo c d\<close>)
next
case (AllowPortFromTo e f p) thus ?thesis using `a = DenyAllFromTo c d` by simp
case (AllowPortFromTo e f p) thus ?thesis using \<open>a = DenyAllFromTo c d\<close> by simp
next
case (Conc e f) thus ?thesis using Conc `a = DenyAllFromTo c d` by simp
case (Conc e f) thus ?thesis using Conc \<open>a = DenyAllFromTo c d\<close> by simp
qed
next
case (AllowPortFromTo c d p) thus ?thesis
@ -1760,7 +1760,7 @@ next
next
case (DenyAllFromTo e f) thus ?thesis using AllowPortFromTo by simp
next
case (AllowPortFromTo e f p2) thus ?thesis using `a = AllowPortFromTo c d p` by simp
case (AllowPortFromTo e f p2) thus ?thesis using \<open>a = AllowPortFromTo c d p\<close> by simp
next
case (Conc e f) thus ?thesis using AllowPortFromTo by simp
qed
@ -1913,14 +1913,14 @@ proof -
fix aa::"('a, 'b) Combinators"
assume 6: "aa \<in> set p"
show "first_bothNet a \<noteq> first_bothNet aa"
apply(insert 1 2 3 4 5 6 `a = DenyAll`, simp_all)
apply(insert 1 2 3 4 5 6 \<open>a = DenyAll\<close>, simp_all)
using fMTaux noDA by blast
next
case (DenyAllFromTo x21 x22)
fix aa::"('a, 'b) Combinators"
assume 6: "first_bothNet a \<noteq> firstList p" and 7 :"aa \<in> set p"
show "first_bothNet a \<noteq> first_bothNet aa"
apply(insert 1 2 3 4 5 6 7 `a = DenyAllFromTo x21 x22`)
apply(insert 1 2 3 4 5 6 7 \<open>a = DenyAllFromTo x21 x22\<close>)
apply(case_tac aa, simp_all)
apply (meson NCSaux1)
apply (meson NCSaux2)
@ -1930,7 +1930,7 @@ proof -
fix aa::"('a, 'b) Combinators"
assume 6: "first_bothNet a \<noteq> firstList p" and 7 :"aa \<in> set p"
show "first_bothNet a \<noteq> first_bothNet aa"
apply(insert 1 2 3 4 6 7 `a = AllowPortFromTo x31 x32 x33`)
apply(insert 1 2 3 4 6 7 \<open>a = AllowPortFromTo x31 x32 x33\<close>)
apply(case_tac aa, simp_all)
apply (meson NCSaux3)
apply (meson NCSaux4)
@ -1939,7 +1939,7 @@ proof -
case (Conc x41 x42)
fix aa::"('a, 'b) Combinators"
show "first_bothNet a \<noteq> first_bothNet aa"
by(insert 3 4 `a = x41 \<oplus> x42`,simp)
by(insert 3 4 \<open>a = x41 \<oplus> x42\<close>,simp)
qed
qed
@ -2271,7 +2271,7 @@ next
(\<forall>aa\<in>set p. first_bothNet a \<noteq> first_bothNet aa) \<and> NetsCollected p \<longrightarrow>
noDenyAll1 (a # p) \<longrightarrow> allNetsDistinct (a # p) \<longrightarrow>
(\<forall>s. s \<in> set p \<longrightarrow> disjSD_2 a s) \<and> separated p"
apply(insert Cons.hyps `a = DenyAll`)
apply(insert Cons.hyps \<open>a = DenyAll\<close>)
apply (intro impI,drule mp, erule OTNConc,drule mp)
apply (case_tac p, simp_all)
apply (case_tac a, simp_all)
@ -2283,7 +2283,7 @@ next
(\<forall>aa\<in>set p. first_bothNet a \<noteq> first_bothNet aa) \<and> NetsCollected p \<longrightarrow>
noDenyAll1 (a # p) \<longrightarrow> allNetsDistinct (a # p) \<longrightarrow> (\<forall>s. s \<in> set p \<longrightarrow>
disjSD_2 a s) \<and> separated p"
apply(insert Cons.hyps `a \<noteq> DenyAll`)
apply(insert Cons.hyps \<open>a \<noteq> DenyAll\<close>)
by (metis NetsCollected.simps(1) NetsCollected2.simps(1) separated.simps(1) separatedNC)
qed
qed

View File

@ -35,16 +35,16 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Normalisation Proofs: Integer Protocol *}
subsection \<open>Normalisation Proofs: Integer Protocol\<close>
theory
NormalisationIPPProofs
imports
NormalisationIntegerPortProof
begin
text{*
text\<open>
Normalisation proofs which are specific to the IntegerProtocol address representation.
*}
\<close>
lemma ConcAssoc: "Cp((A \<oplus> B) \<oplus> D) = Cp(A \<oplus> (B \<oplus> D))"
by (simp add: Cp.simps)
@ -980,7 +980,7 @@ next
case True thus ?thesis using Cons by (simp) (rule impI, simp)
next
case False thus ?thesis
using Cons False `\<not> not_MT Cp ys` apply (simp)
using Cons False \<open>\<not> not_MT Cp ys\<close> apply (simp)
apply (intro conjI impI| simp)+
apply (subgoal_tac "rm_MT_rules Cp ys = []")
apply (subgoal_tac "x \<notin> dom (Cp (list2FWpolicy ys))")
@ -1547,7 +1547,7 @@ proof (cases "p")
case Nil then show ?thesis by simp
next
case (Cons a list) show ?thesis
apply (insert `p = a # list`, simp_all)
apply (insert \<open>p = a # list\<close>, simp_all)
proof (cases "a = DenyAll")
case True
assume * : "a = DenyAll"
@ -1706,7 +1706,7 @@ lemmas domain_reasoningPr = domDA ConcAssoc2 domSubset1 domSubset2
domSubsetDistr2 domSubsetDistrA domSubsetDistrD coerc_assoc ConcAssoc
ConcAssoc3
text {* The following lemmas help with the normalisation *}
text \<open>The following lemmas help with the normalisation\<close>
lemma list2policyR_Start[rule_format]: "p \<in> dom (Cp a) \<longrightarrow>
Cp (list2policyR (a # list)) p = Cp a p"
by (induct "a # list" rule:list2policyR.induct)
@ -1824,9 +1824,9 @@ lemma norm_notMT_manual: "DenyAll \<in> set (policy2list p) \<Longrightarrow> no
unfolding normalize_manual_orderPr_def
by (simp add: idNMT rADnMT wellformed1_alternative_sorted wp1ID wp1_alternativesep wp1n_RS2)
text{*
text\<open>
As an example, how this theorems can be used for a concrete normalisation instantiation.
*}
\<close>
lemma normalizePrNAT:
"DenyAll \<in> set (policy2list Filter) \<Longrightarrow>

View File

@ -35,16 +35,16 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Normalisation Proofs: Integer Port *}
subsection \<open>Normalisation Proofs: Integer Port\<close>
theory
NormalisationIntegerPortProof
imports
NormalisationGenericProofs
begin
text{*
text\<open>
Normalisation proofs which are specific to the IntegerPort address representation.
*}
\<close>
lemma ConcAssoc: "C((A \<oplus> B) \<oplus> D) = C(A \<oplus> (B \<oplus> D))"
by (simp add: C.simps)
@ -967,7 +967,7 @@ next
case True thus ?thesis using Cons by (simp) (rule impI, simp)
next
case False thus ?thesis
using Cons False `\<not> not_MT C ys` apply (simp)
using Cons False \<open>\<not> not_MT C ys\<close> apply (simp)
by (metis SR3nMT l2p_aux list2FWpolicy.simps(2) nMT_domMT nlpaux)
qed
qed
@ -1693,7 +1693,7 @@ lemmas domain_reasoning = domDA ConcAssoc2 domSubset1 domSubset2
domSubsetDistr2 domSubsetDistrA domSubsetDistrD coerc_assoc ConcAssoc
ConcAssoc3
text {* The following lemmas help with the normalisation *}
text \<open>The following lemmas help with the normalisation\<close>
lemma list2policyR_Start[rule_format]: "p \<in> dom (C a) \<longrightarrow>
C (list2policyR (a # list)) p = C a p"
by (induct "a # list" rule:list2policyR.induct) (auto simp: C.simps dom_def map_add_def)
@ -1798,9 +1798,9 @@ lemma p_eq2_manual[rule_format]:
lemma norm_notMT_manual: "DenyAll \<in> set (policy2list p) \<Longrightarrow> normalize_manual_order p l \<noteq> []"
by (simp add: RS2_NMT idNMT normalize_manual_order_def rADnMT sepnMT sortnMT wp1ID)
text{*
text\<open>
As an example, how this theorems can be used for a concrete normalisation instantiation.
*}
\<close>
lemma normalizeNAT:
"DenyAll \<in> set (policy2list Filter) \<Longrightarrow> allNetsDistinct (policy2list Filter) \<Longrightarrow>
@ -1813,11 +1813,11 @@ lemma normalizeNAT:
lemma domSimpl[simp]: "dom (C (A \<oplus> DenyAll)) = dom (C (DenyAll))"
by (simp add: PLemmas)
text {*
text \<open>
The followin theorems can be applied when prepending the usual normalisation with an
additional step and using another semantical interpretation function. This is a general recipe
which can be applied whenever one nees to combine several normalisation strategies.
*}
\<close>
lemma CRotate_eq_rotateC: "CRotate p = C (rotatePolicy p)"
by (induct p rule: rotatePolicy.induct) (simp_all add: C.simps map_add_def)

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection{* Network Address Translation *}
subsection\<open>Network Address Translation\<close>
theory
NAT
imports
@ -91,10 +91,10 @@ definition srcPat2pool_IntProtocol ::
"srcPat2pool_IntProtocol srcs transl =
{x. (fst (src x)) \<in> srcs} \<triangleleft> (A\<^sub>f (srcPat2poolPort_Protocol_t transl))"
text{*
text\<open>
The following lemmas are used for achieving a normalized output format of packages after
applying NAT. This is used, e.g., by our firewall execution tool.
*}
\<close>
lemma datasimp: "{(i, (s1, s2, s3), aba).
\<forall>a aa b ba. aba = ((a, aa, b), ba) \<longrightarrow> i = i1 \<and> s1 = i101 \<and>

View File

@ -35,17 +35,17 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Datatype Addresses *}
subsection \<open>Datatype Addresses\<close>
theory
DatatypeAddress
imports
NetworkCore
begin
text{*
text\<open>
A theory describing a network consisting of three subnetworks. Hosts within a network are not
distinguished.
*}
\<close>
datatype DatatypeAddress = dmz_adr | intranet_adr | internet_adr

View File

@ -35,17 +35,17 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Datatype Addresses with Ports *}
subsection \<open>Datatype Addresses with Ports\<close>
theory
DatatypePort
imports
NetworkCore
begin
text{*
text\<open>
A theory describing a network consisting of three subnetworks, including port numbers modelled
as Integers. Hosts within a network are not distinguished.
*}
\<close>
datatype DatatypeAddress = dmz_adr | intranet_adr | internet_adr

View File

@ -35,16 +35,16 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Formalizing IPv4 Addresses *}
subsection \<open>Formalizing IPv4 Addresses\<close>
theory
IPv4
imports
NetworkCore
begin
text{*
text\<open>
A theory describing IPv4 addresses with ports. The host address is a four-tuple of Integers,
the port number is a single Integer.
*}
\<close>
type_synonym
ipv4_ip = "(int \<times> int \<times> int \<times> int)"

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* IPv4 with Ports and Protocols *}
subsection \<open>IPv4 with Ports and Protocols\<close>
theory
IPv4_TCPUDP
imports IPv4

View File

@ -35,14 +35,14 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Integer Addresses *}
subsection \<open>Integer Addresses\<close>
theory
IntegerAddress
imports
NetworkCore
begin
text{* A theory where addresses are modelled as Integers.*}
text\<open>A theory where addresses are modelled as Integers.\<close>
type_synonym
adr\<^sub>i = "int"

View File

@ -35,17 +35,17 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection{* Integer Addresses with Ports *}
subsection\<open>Integer Addresses with Ports\<close>
theory
IntegerPort
imports
NetworkCore
begin
text{*
text\<open>
A theory describing addresses which are modelled as a pair of Integers - the first being the
host address, the second the port number.
*}
\<close>
type_synonym
address = int

View File

@ -35,15 +35,15 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Integer Addresses with Ports and Protocols *}
subsection \<open>Integer Addresses with Ports and Protocols\<close>
theory
IntegerPort_TCPUDP
imports
NetworkCore
begin
text{* A theory describing addresses which are modelled as a pair of Integers - the first being
the host address, the second the port number.*}
text\<open>A theory describing addresses which are modelled as a pair of Integers - the first being
the host address, the second the port number.\<close>
type_synonym
address = int
@ -92,7 +92,7 @@ lemma src_port: "src_port (a,x::adr\<^sub>i\<^sub>p\<^sub>p,d,e) = fst (snd x)"
lemma dest_port: "dest_port (a,d,x::adr\<^sub>i\<^sub>p\<^sub>p,e) = fst (snd x)"
by (simp add: dest_port_int_TCPUDP_def in_subnet)
text {* Common test constraints: *}
text \<open>Common test constraints:\<close>
definition port_positive :: "(adr\<^sub>i\<^sub>p\<^sub>p,'b) packet \<Rightarrow> bool" where
"port_positive x = (dest_port x > (0::port))"

View File

@ -35,14 +35,14 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection{* Packets and Networks *}
subsection\<open>Packets and Networks\<close>
theory
NetworkCore
imports
Main
begin
text{*
text\<open>
In networks based e.g. on TCP/IP, a message from A to B is encapsulated in \emph{packets}, which
contain the content of the message and routing information. The routing information mainly
contains its source and its destination address.
@ -50,15 +50,15 @@ text{*
In the case of stateless packet filters, a firewall bases its decision upon this routing
information and, in the stateful case, on the content. Thus, we model a packet as a four-tuple of
the mentioned elements, together with an id field.
*}
\<close>
text{* The ID is an integer: *}
text\<open>The ID is an integer:\<close>
type_synonym id = int
text{*
text\<open>
To enable different representations of addresses (e.g. IPv4 and IPv6, with or without ports),
we model them as an unconstrained type class and directly provide several instances:
*}
\<close>
class adr
type_synonym '\<alpha> src = "'\<alpha>"
@ -70,36 +70,36 @@ instance nat ::adr ..
instance "fun" :: (adr,adr) adr ..
instance prod :: (adr,adr) adr ..
text{*
text\<open>
The content is also specified with an unconstrained generic type:
*}
\<close>
type_synonym '\<beta> content = "'\<beta>"
text {*
text \<open>
For applications where the concrete representation of the content field does not matter (usually
the case for stateless packet filters), we provide a default type which can be used in those
cases:
*}
\<close>
datatype DummyContent = data
text{* Finally, a packet is:*}
text\<open>Finally, a packet is:\<close>
type_synonym ('\<alpha>,'\<beta>) packet = "id \<times> '\<alpha> src \<times> '\<alpha> dest \<times> '\<beta> content"
text{*
text\<open>
Protocols (e.g. http) are not modelled explicitly. In the case of stateless packet filters, they
are only visible by the destination port of a packet, which are modelled as part of the address.
Additionally, stateful firewalls often determine the protocol by the content of a packet.
*}
\<close>
definition src :: "('\<alpha>::adr,'\<beta>) packet \<Rightarrow> '\<alpha>"
where "src = fst o snd "
text{*
text\<open>
Port numbers (which are part of an address) are also modelled in a generic way. The integers and
the naturals are typical representations of port numbers.
*}
\<close>
class port
@ -108,13 +108,13 @@ instance nat :: port ..
instance "fun" :: (port,port) port ..
instance "prod" :: (port,port) port ..
text{*
text\<open>
A packet therefore has two parameters, the first being the address, the second the content. For
the sake of simplicity, we do not allow to have a different address representation format for the
source and the destination of a packet.
To access the different parts of a packet directly, we define a couple of projectors:
*}
\<close>
definition id :: "('\<alpha>::adr,'\<beta>) packet \<Rightarrow> id"
where "id = fst"
@ -135,28 +135,28 @@ lemma either2[simp]: "(a \<noteq> tcp) = (a = udp)"
lemma either3[simp]: "(a \<noteq> udp) = (a = tcp)"
by (case_tac "a",simp_all)
text{*
text\<open>
The following two constants give the source and destination port number of a packet. Address
representations using port numbers need to provide a definition for these types.
*}
\<close>
consts src_port :: "('\<alpha>::adr,'\<beta>) packet \<Rightarrow> '\<gamma>::port"
consts dest_port :: "('\<alpha>::adr,'\<beta>) packet \<Rightarrow> '\<gamma>::port"
consts src_protocol :: "('\<alpha>::adr,'\<beta>) packet \<Rightarrow> protocol"
consts dest_protocol :: "('\<alpha>::adr,'\<beta>) packet \<Rightarrow> protocol"
text{* A subnetwork (or simply a network) is a set of sets of addresses.*}
text\<open>A subnetwork (or simply a network) is a set of sets of addresses.\<close>
type_synonym '\<alpha> net = "'\<alpha> set set"
text{* The relation {in\_subnet} (@{text "\<sqsubset>"}) checks if an address is in a specific network. *}
text\<open>The relation {in\_subnet} (\<open>\<sqsubset>\<close>) checks if an address is in a specific network.\<close>
definition
in_subnet :: "'\<alpha>::adr \<Rightarrow> '\<alpha> net \<Rightarrow> bool" (infixl "\<sqsubset>" 100) where
"in_subnet a S = (\<exists> s \<in> S. a \<in> s)"
text{* The following lemmas will be useful later. *}
text\<open>The following lemmas will be useful later.\<close>
lemma in_subnet:
"(a, e) \<sqsubset> {{(x1,y). P x1 y}} = P a e"
@ -170,10 +170,10 @@ lemma dest_in_subnet:
"dest (q,r,((a),e),t) \<sqsubset> {{(x1,y). P x1 y}} = P a e"
by (simp add: in_subnet_def in_subnet dest_def)
text{*
text\<open>
Address models should provide a definition for the following constant, returning a network
consisting of the input address only.
*}
\<close>
consts subnet_of :: "'\<alpha>::adr \<Rightarrow> '\<alpha> net"

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
section{* Network Models *}
section\<open>Network Models\<close>
theory
NetworkModels
imports
@ -50,7 +50,7 @@ theory
IPv4_TCPUDP
begin
text{*
text\<open>
One can think of many different possible address representations. In this distribution, we include
seven different variants:
\begin{itemize}
@ -74,6 +74,6 @@ text{*
The theories of each pf the networks are relatively small. It suffices to provide the required
types, a couple of lemmas, and - if required - a definition for the source and destination ports
of a packet.
*}
\<close>
end

View File

@ -34,7 +34,7 @@
* (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
section {* Network Policies: Packet Filter *}
section \<open>Network Policies: Packet Filter\<close>
theory
PacketFilter
imports

View File

@ -35,17 +35,17 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Policy Combinators *}
subsection \<open>Policy Combinators\<close>
theory
PolicyCombinators
imports
PolicyCore
begin
text{* In order to ease the specification of a concrete policy, we
text\<open>In order to ease the specification of a concrete policy, we
define some combinators. Using these combinators, the specification
of a policy gets very easy, and can be done similarly as in tools
like IPTables. *}
like IPTables.\<close>
definition
allow_all_from :: "'\<alpha>::adr net \<Rightarrow> (('\<alpha>,'\<beta>) packet \<mapsto> unit)" where
@ -73,9 +73,9 @@ definition
"deny_all_from_to src_net dest_net = {pa. src pa \<sqsubset> src_net \<and> dest pa \<sqsubset> dest_net} \<triangleleft> D\<^sub>U"
text{* All these combinators and the default rules are put into one
single lemma called @{text PolicyCombinators} to facilitate proving
over policies. *}
text\<open>All these combinators and the default rules are put into one
single lemma called \<open>PolicyCombinators\<close> to facilitate proving
over policies.\<close>
lemmas PolicyCombinators = allow_all_from_def deny_all_from_def

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Policy Core *}
subsection \<open>Policy Core\<close>
theory
PolicyCore
imports
@ -44,25 +44,25 @@ theory
begin
text{* A policy is seen as a partial mapping from packet to packet out. *}
text\<open>A policy is seen as a partial mapping from packet to packet out.\<close>
type_synonym ('\<alpha>, '\<beta>) FWPolicy = "('\<alpha>, '\<beta>) packet \<mapsto> unit"
text{*
text\<open>
When combining several rules, the firewall is supposed to apply the
first matching one. In our setting this means the first rule which
maps the packet in question to @{text "Some (packet out)"}. This is
exactly what happens when using the map-add operator (@{text "rule1
++ rule2"}). The only difference is that the rules must be given in
maps the packet in question to \<open>Some (packet out)\<close>. This is
exactly what happens when using the map-add operator (\<open>rule1
++ rule2\<close>). The only difference is that the rules must be given in
reverse order.
*}
\<close>
text{*
The constant @{text p_accept} is @{text "True"} iff the policy
text\<open>
The constant \<open>p_accept\<close> is \<open>True\<close> iff the policy
accepts the packet.
*}
\<close>
definition
p_accept :: "('\<alpha>, '\<beta>) packet \<Rightarrow> ('\<alpha>, '\<beta>) FWPolicy \<Rightarrow> bool" where

View File

@ -35,14 +35,14 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Policy Combinators with Ports *}
subsection \<open>Policy Combinators with Ports\<close>
theory
PortCombinators
imports
PolicyCombinators
begin
text{*
text\<open>
This theory defines policy combinators for those network models which
have ports. They are provided in addition to the the ones defined in the
PolicyCombinators theory.
@ -52,7 +52,7 @@ text{*
\item $src\_port :: ('\alpha,'\beta) packet \Rightarrow ('\gamma::port)$
\item $dest\_port :: ('\alpha,'\beta) packet \Rightarrow ('\gamma::port)$
\end{itemize}
*}
\<close>
definition
allow_all_from_port :: "'\<alpha>::adr net \<Rightarrow> ('\<gamma>::port) \<Rightarrow> (('\<alpha>,'\<beta>) packet \<mapsto> unit)" where
@ -160,9 +160,9 @@ definition
"allow_all_from_port_tos src_net s_port dest_net
= {pa. dest_port pa \<in> s_port} \<triangleleft> allow_all_from_to src_net dest_net"
text{*
text\<open>
As before, we put all the rules into one lemma called PortCombinators to ease writing later.
*}
\<close>
lemmas PortCombinatorsCore =
allow_all_from_port_def deny_all_from_port_def allow_all_to_port_def

View File

@ -35,16 +35,16 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Ports *}
subsection \<open>Ports\<close>
theory Ports
imports
Main
begin
text{*
text\<open>
This theory can be used if we want to specify the port numbers by names denoting their default
Integer values. If you want to use them, please add @{text Ports} to the simplifier.
*}
Integer values. If you want to use them, please add \<open>Ports\<close> to the simplifier.
\<close>
definition http::int where "http = 80"
@ -70,7 +70,7 @@ lemma ftp1: "x \<noteq> 21 \<Longrightarrow> x \<noteq> ftp"
lemma ftp2: "x \<noteq> 21 \<Longrightarrow> ftp \<noteq> x"
by (simp add: ftp_def)
text{* And so on for all desired port numbers. *}
text\<open>And so on for all desired port numbers.\<close>
lemmas Ports = http1 http2 ftp1 ftp2 smtp1 smtp2

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Policy Combinators with Ports and Protocols *}
subsection \<open>Policy Combinators with Ports and Protocols\<close>
theory
ProtocolPortCombinators
@ -43,7 +43,7 @@ theory
PortCombinators
begin
text{*
text\<open>
This theory defines policy combinators for those network models which
have ports. They are provided in addition to the the ones defined in the
PolicyCombinators theory.
@ -53,7 +53,7 @@ text{*
\item $src\_port :: ('\alpha,'\beta) packet \Rightarrow ('\gamma::port)$
\item $dest\_port :: ('\alpha,'\beta) packet \Rightarrow ('\gamma::port)$
\end{itemize}
*}
\<close>
definition
allow_all_from_port_prot :: "protocol \<Rightarrow> '\<alpha>::adr net \<Rightarrow> ('\<gamma>::port) \<Rightarrow> (('\<alpha>,'\<beta>) packet \<mapsto> unit)" where
@ -159,7 +159,7 @@ definition
"deny_from_to_ports_prot p ports src_net dest_net =
{pa. dest_protocol pa = p} \<triangleleft> deny_from_to_ports ports src_net dest_net"
text{* As before, we put all the rules into one lemma to ease writing later. *}
text\<open>As before, we put all the rules into one lemma to ease writing later.\<close>
lemmas ProtocolCombinatorsCore =
allow_all_from_port_prot_def deny_all_from_port_prot_def allow_all_to_port_prot_def

View File

@ -1,7 +1,7 @@
chapter AFP
session "UPF_Firewall-devel" (AFP) = UPF +
description {* Formal Network Models and Their Application to Firewall Policies *}
description "Formal Network Models and Their Application to Firewall Policies"
options [timeout = 600]
theories
"Examples/Examples"

View File

@ -35,46 +35,46 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* The File Transfer Protocol (ftp) *}
subsection \<open>The File Transfer Protocol (ftp)\<close>
theory
FTP
imports
StatefulCore
begin
subsubsection{* The protocol syntax *}
text{*
subsubsection\<open>The protocol syntax\<close>
text\<open>
The File Transfer Protocol FTP is a well known example of a protocol which uses dynamic ports and
is therefore a natural choice to use as an example for our model.
We model only a simplified version of the FTP protocol over IntegerPort addresses, still
containing all messages that matter for our purposes. It consists of the following four messages:
\begin{enumerate}
\item @{text "init"}: The client contacts the server indicating
\item \<open>init\<close>: The client contacts the server indicating
his wish to get some data.
\item @{text "ftp_port_request p"}: The client, usually after having
\item \<open>ftp_port_request p\<close>: The client, usually after having
received an acknowledgement of the server, indicates a port
number on which he wants to receive the data.
\item @{text "ftp_ftp_data"}: The server sends the requested data over
\item \<open>ftp_ftp_data\<close>: The server sends the requested data over
the new channel. There might be an arbitrary number of such
messages, including zero.
\item @{text "ftp_close"}: The client closes the connection. The
\item \<open>ftp_close\<close>: The client closes the connection. The
dynamic port gets closed again.
\end{enumerate}
The content field of a packet therefore now consists of either one of those four messages or a
default one.
*}
\<close>
datatype msg = ftp_init | ftp_port_request port | ftp_data | ftp_close | ftp_other
text{*
text\<open>
We now also make use of the ID field of a packet. It is used as session ID and we make the
assumption that they are all unique among different protocol runs.
At first, we need some predicates which check if a packet is a specific FTP message and has the
correct session ID.
*}
\<close>
definition
is_init :: "id \<Rightarrow> (adr\<^sub>i\<^sub>p, msg)packet \<Rightarrow> bool" where
@ -104,8 +104,8 @@ fun are_ftp_other where
"are_ftp_other i (x#xs) = (is_ftp_other i x \<and> are_ftp_other i xs)"
|"are_ftp_other i [] = True"
subsubsection{* The protocol policy specification *}
text{*
subsubsection\<open>The protocol policy specification\<close>
text\<open>
We now have to model the respective state transitions. It is important to note that state
transitions themselves allow all packets which are allowed by the policy, not only those which
are allowed by the protocol. Their only task is to change the policy. As an alternative, we could
@ -116,7 +116,7 @@ text{*
cases, one is enough. In our example, only messages 2 and 4 need special transitions. The default
says that if the policy accepts the packet, it is added to the history, otherwise it is simply
dropped. The policy remains the same in both cases.
*}
\<close>
fun last_opened_port where
"last_opened_port i ((j,s,d,ftp_port_request p)#xs) = (if i=j then p else last_opened_port i xs)"
@ -153,11 +153,11 @@ definition TRPolicy ::" (adr\<^sub>i\<^sub>p,msg)packet \<times> (adr\<^sub>i
definition TRPolicy\<^sub>M\<^sub>o\<^sub>n
where "TRPolicy\<^sub>M\<^sub>o\<^sub>n = policy2MON(TRPolicy)"
text{* If required to contain the policy in the output *}
text\<open>If required to contain the policy in the output\<close>
definition TRPolicy\<^sub>M\<^sub>o\<^sub>n'
where "TRPolicy\<^sub>M\<^sub>o\<^sub>n' = policy2MON (((\<lambda>(x,y,z). (z,(y,z))) o_f TRPolicy ))"
text{*
text\<open>
Now we specify our test scenario in more detail. We could test:
\begin{itemize}
\item one correct FTP-Protocol run,
@ -168,17 +168,17 @@ text{*
\end{itemize}
We only do the the simplest case here: one correct protocol run.
*}
\<close>
text{*
text\<open>
There are four different states which are modelled as a datatype.
*}
\<close>
datatype ftp_states = S0 | S1 | S2 | S3
text{*
The following constant is @{text "True"} for all sets which are correct FTP runs for a given
text\<open>
The following constant is \<open>True\<close> for all sets which are correct FTP runs for a given
source and destination address, ID, and data-port number.
*}
\<close>
fun
@ -197,13 +197,13 @@ fun
definition is_single_ftp_run :: "adr\<^sub>i\<^sub>p src \<Rightarrow> adr\<^sub>i\<^sub>p dest \<Rightarrow> id \<Rightarrow> port \<Rightarrow> (adr\<^sub>i\<^sub>p,msg) history set"
where "is_single_ftp_run s d i p = {x. (is_ftp S0 s d i p x)}"
text{*
text\<open>
The following constant then returns a set of all the historys which denote such a normal
behaviour FTP run, again for a given source and destination address, ID, and data-port.
The following definition returns the set of all possible interleaving of two correct FTP protocol
runs.
*}
\<close>
definition
ftp_2_interleaved :: "adr\<^sub>i\<^sub>p src \<Rightarrow> adr\<^sub>i\<^sub>p dest \<Rightarrow> id \<Rightarrow> port \<Rightarrow>
adr\<^sub>i\<^sub>p src \<Rightarrow> adr\<^sub>i\<^sub>p dest \<Rightarrow> id \<Rightarrow> port \<Rightarrow>

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* FTP and VoIP Protocol *}
subsection \<open>FTP and VoIP Protocol\<close>
theory
FTPVOIP
imports
@ -56,13 +56,13 @@ datatype ftpvoip = ARQ
| other
text{*
text\<open>
We now also make use of the ID field of a packet. It is used as session ID and we make the
assumption that they are all unique among different protocol runs.
At first, we need some predicates which check if a packet is a specific FTP message and has
the correct session ID.
*}
\<close>
definition
FTPVOIP_is_init :: "id \<Rightarrow> (adr\<^sub>i\<^sub>p, ftpvoip ) packet \<Rightarrow> bool" where
@ -140,19 +140,19 @@ definition
"FTPVOIP_is_setup i port p = (id p = i \<and> content p = Setup port)"
text{*
We need also an operator @{text ports_open} to get access to the two
text\<open>
We need also an operator \<open>ports_open\<close> to get access to the two
dynamic ports.
*}
\<close>
definition
FTPVOIP_ports_open :: "id \<Rightarrow> port \<times> port \<Rightarrow> (adr\<^sub>i\<^sub>p, ftpvoip) history \<Rightarrow> bool" where
"FTPVOIP_ports_open i p L = ((not_before (FTPVOIP_is_fin i) (FTPVOIP_is_setup i (fst p)) L) \<and>
not_before (FTPVOIP_is_fin i) (FTPVOIP_is_connect i (snd p)) L)"
text{*
text\<open>
As we do not know which entity closes the connection, we define an
operator which checks if the closer is the caller.
*}
\<close>
fun
FTPVOIP_src_is_initiator :: "id \<Rightarrow> adr\<^sub>i\<^sub>p \<Rightarrow> (adr\<^sub>i\<^sub>p,ftpvoip) history \<Rightarrow> bool" where
"FTPVOIP_src_is_initiator i a [] = False"
@ -243,11 +243,11 @@ subnet_of_int_def id_def FTPVOIP_port_open_def
datatype ftp_states2 = FS0 | FS1 | FS2 | FS3
datatype voip_states2 = V0 | V1 | V2 | V3 | V4 | V5
text{*
The constant @{text "is_voip"} checks if a trace corresponds to a
text\<open>
The constant \<open>is_voip\<close> checks if a trace corresponds to a
legal VoIP protocol, given the IP-addresses of the three entities,
the ID, and the two dynamic ports.
*}
\<close>
fun FTPVOIP_is_voip :: "voip_states2 \<Rightarrow> address \<Rightarrow> address \<Rightarrow> address \<Rightarrow> id \<Rightarrow> port \<Rightarrow>
port \<Rightarrow> (adr\<^sub>i\<^sub>p, ftpvoip) history \<Rightarrow> bool"
@ -276,11 +276,11 @@ where
FTPVOIP_is_voip V1 s d g i p1 p2 InL)))))) x)"
text{*
Finally, @{text "NB_voip"} returns the set of protocol traces which
text\<open>
Finally, \<open>NB_voip\<close> returns the set of protocol traces which
correspond to a correct protocol run given the three addresses, the
ID, and the two dynamic ports.
*}
\<close>
definition
FTPVOIP_NB_voip :: "address \<Rightarrow> address \<Rightarrow> address \<Rightarrow> id \<Rightarrow> port \<Rightarrow> port \<Rightarrow>
(adr\<^sub>i\<^sub>p, ftpvoip) history set" where

View File

@ -35,18 +35,18 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* FTP enriched with a security policy *}
subsection \<open>FTP enriched with a security policy\<close>
theory
FTP_WithPolicy
imports
FTP
begin
text{* FTP where the policy is part of the output. *}
text\<open>FTP where the policy is part of the output.\<close>
definition POL :: "'a \<Rightarrow> 'a" where "POL x = x"
text{* Variant 2 takes the policy into the output *}
text\<open>Variant 2 takes the policy into the output\<close>
fun FTP_STP ::
"((id \<rightharpoonup> port), adr\<^sub>i\<^sub>p, msg) FWStateTransitionP"
where

View File

@ -35,27 +35,27 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Termporal Combinators *}
subsection \<open>Termporal Combinators\<close>
theory
LTL_alike
imports
Main
begin
text{*
text\<open>
In the following, we present a small embbeding of temporal combinators, that may help to
formulate typical temporal properties in traces and protocols concisely. It is based on
\emph{finite} lists, therefore the properties of this logic are not fully compatible with
LTL based on Kripke-structures. For the purpose of this demonstration, however, the difference
does not matter.
*}
\<close>
fun nxt :: "('\<alpha> list \<Rightarrow> bool) \<Rightarrow> '\<alpha> list \<Rightarrow> bool" ("N")
where
"nxt p [] = False"
| "nxt p (a # S) = (p S)"
text{* Predicate $p$ holds at first position. *}
text\<open>Predicate $p$ holds at first position.\<close>
fun atom :: "('\<alpha> \<Rightarrow> bool) \<Rightarrow> '\<alpha> list \<Rightarrow> bool" ("\<guillemotleft>_\<guillemotright>")
where
@ -71,10 +71,10 @@ where
"always p [] = True"
| "always p (a # S) = ((p (a # S)) \<and> always p S)"
text{*
text\<open>
Always is a generalization of the \verb+list_all+ combinator from the List-library; if arguing
locally, this paves the way to a wealth of library lemmas.
*}
\<close>
lemma always_is_listall : "(\<box> \<guillemotleft>p\<guillemotright>) (t) = list_all (p) (t)"
by(induct "t", simp_all)
@ -84,29 +84,29 @@ where
| "eventually p (a # S) = ((p (a # S)) \<or> eventually p S)"
text{*
text\<open>
Eventually is a generalization of the \verb+list_ex+ combinator from the List-library; if arguing
locally, this paves the way to a wealth of library lemmas.
*}
\<close>
lemma eventually_is_listex : "(\<diamondsuit> \<guillemotleft>p\<guillemotright>) (t) = list_ex (p) (t)"
by(induct "t", simp_all)
text{*
text\<open>
The next two constants will help us later in defining the state transitions. The constant
@{text "before"} is @{text "True"} if for all elements which appear before the first element
for which @{text q} holds, @{text p} must hold.
*}
\<open>before\<close> is \<open>True\<close> if for all elements which appear before the first element
for which \<open>q\<close> holds, \<open>p\<close> must hold.
\<close>
fun before :: "('\<alpha> \<Rightarrow> bool) \<Rightarrow> ('\<alpha> \<Rightarrow> bool) \<Rightarrow> '\<alpha> list \<Rightarrow> bool"
where
"before p q [] = False"
| "before p q (a # S) = (q a \<or> (p a \<and> (before p q S)))"
text{*
Analogously there is an operator @{text not_before} which returns
@{text "True"} if for all elements which appear before the first
element for which @{text q} holds, @{text p} must not hold.
*}
text\<open>
Analogously there is an operator \<open>not_before\<close> which returns
\<open>True\<close> if for all elements which appear before the first
element for which \<open>q\<close> holds, \<open>p\<close> must not hold.
\<close>
fun not_before :: "('\<alpha> \<Rightarrow> bool) \<Rightarrow> ('\<alpha> \<Rightarrow> bool) \<Rightarrow> '\<alpha> list \<Rightarrow> bool"
where
@ -122,13 +122,13 @@ lemma not_before_superfluous:
done
done
text{*General "before":*}
text\<open>General "before":\<close>
fun until :: "('\<alpha> list \<Rightarrow> bool) \<Rightarrow> ('\<alpha> list \<Rightarrow> bool) \<Rightarrow> '\<alpha> list \<Rightarrow> bool" (infixl "U" 66)
where
"until p q [] = False"
| "until p q (a # S) = (\<exists> s t. a # S= s @ t \<and> p s \<and> q t)"
text{* This leads to this amazingly tricky proof:*}
text\<open>This leads to this amazingly tricky proof:\<close>
lemma before_vs_until:
"(before p q) = ((\<box>\<guillemotleft>p\<guillemotright>) U \<guillemotleft>q\<guillemotright>)"
proof -

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* Stateful Protocols: Foundations *}
subsection \<open>Stateful Protocols: Foundations\<close>
theory
StatefulCore
imports
@ -43,7 +43,7 @@ theory
LTL_alike
begin
text{*
text\<open>
The simple system of a stateless packet filter is not enough to model all common real-world
scenarios. Some protocols need further actions in order to be secured. A prominent example is
the File Transfer Protocol (FTP), which is a popular means to move files across the Internet.
@ -65,16 +65,16 @@ text{*
At first we hence need a state. It is a tuple from some memory to be refined later and the
current policy.
*}
\<close>
type_synonym ('\<alpha>,'\<beta>,'\<gamma>) FWState = "'\<alpha> \<times> (('\<beta>,'\<gamma>) packet \<mapsto> unit)"
text{* Having a state, we need of course some state transitions. Such
text\<open>Having a state, we need of course some state transitions. Such
a transition can happen every time a new packet arrives. State
transitions can be modelled using a state-exception monad.
*}
\<close>
type_synonym ('\<alpha>,'\<beta>,'\<gamma>) FWStateTransitionP =
@ -83,7 +83,7 @@ type_synonym ('\<alpha>,'\<beta>,'\<gamma>) FWStateTransitionP =
type_synonym ('\<alpha>,'\<beta>,'\<gamma>) FWStateTransition =
"(('\<beta>,'\<gamma>) packet \<times> ('\<alpha>,'\<beta>,'\<gamma>) FWState) \<rightharpoonup> ('\<alpha>,'\<beta>,'\<gamma>) FWState"
text{* The memory could be modelled as a list of accepted packets. *}
text\<open>The memory could be modelled as a list of accepted packets.\<close>
type_synonym ('\<beta>,'\<gamma>) history = "('\<beta>,'\<gamma>) packet list"

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
section {* Stateful Network Protocols *}
section \<open>Stateful Network Protocols\<close>
theory
StatefulFW
imports

View File

@ -35,12 +35,12 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
subsection {* A simple voice-over-ip model *}
subsection \<open>A simple voice-over-ip model\<close>
theory VOIP
imports StatefulCore
begin
text{*
text\<open>
After the FTP-Protocol which was rather simple we show the strength
of the model with a more current and especially much more
@ -69,14 +69,14 @@ text{*
\begin{itemize}
\item Admission Request (ARQ)
\item Admission Reject (ARJ)
\item Admission Confirm (ACF) @{text "'a"}
\item Admission Confirm (ACF) \<open>'a\<close>
\end{itemize}
\item Call Signaling (Q.931, port 1720) The caller and the callee
agree on the dynamic ports over which the call will take
place.
\begin{itemize}
\item Setup @{text "port"}
\item Connect @{text "port"}
\item Setup \<open>port\<close>
\item Connect \<open>port\<close>
\end{itemize}
\item Stream (dynamic ports). The call itself. In reality, several
connections are used here.
@ -96,14 +96,14 @@ text{*
we are not interested in the messages from the callee to its
gatekeeper. Incoming calls are not modelled either, they would
require a different set of state transitions.
*}
\<close>
text{*
text\<open>
The content of a packet now consists of one of the seven messages or
a default one. It is parameterized with the type of the address that
the gatekeeper returns.
*}
\<close>
datatype 'a voip_msg = ARQ
@ -114,11 +114,11 @@ datatype 'a voip_msg = ARQ
| Stream
| Fin
| other
text{*
text\<open>
As before, we need operators which check if a packet contains a
specific content and ID, respectively if such a packet has appeared
in the trace.
*}
\<close>
definition
@ -139,10 +139,10 @@ definition
"is_setup i port p = (id p = i \<and> content p = Setup port)"
text{*
We need also an operator @{text ports_open} to get access to the two
text\<open>
We need also an operator \<open>ports_open\<close> to get access to the two
dynamic ports.
*}
\<close>
definition
ports_open :: "id \<Rightarrow> port \<times> port \<Rightarrow> (adr\<^sub>i\<^sub>p, 'a voip_msg) history \<Rightarrow> bool" where
"ports_open i p L = ((not_before (is_fin i) (is_setup i (fst p)) L) \<and>
@ -151,10 +151,10 @@ definition
text{*
text\<open>
As we do not know which entity closes the connection, we define an
operator which checks if the closer is the caller.
*}
\<close>
fun
src_is_initiator :: "id \<Rightarrow> adr\<^sub>i\<^sub>p \<Rightarrow> (adr\<^sub>i\<^sub>p,'b voip_msg) history \<Rightarrow> bool" where
"src_is_initiator i a [] = False"
@ -165,11 +165,11 @@ fun
text{*
text\<open>
The first state transition is for those messages which do not change
the policy. In this scenario, this only happens for the Stream
messages.
*}
\<close>
definition subnet_of_adr where
"subnet_of_adr x = {{(a,b). a = x}}"
@ -259,16 +259,16 @@ definition VOIP_TRPolicy where
"VOIP_TRPolicy = policy2MON (
((VOIP_STA,VOIP_STD) \<Otimes>\<^sub>\<nabla> applyPolicy) o (\<lambda> (x,(y,z)). ((x,z),(x,(y,z)))))"
text{*
text\<open>
For a full protocol run, six states are needed.
*}
\<close>
datatype voip_states = S0 | S1 | S2 | S3 | S4 | S5
text{*
The constant @{text "is_voip"} checks if a trace corresponds to a
text\<open>
The constant \<open>is_voip\<close> checks if a trace corresponds to a
legal VoIP protocol, given the IP-addresses of the three entities,
the ID, and the two dynamic ports.
*}
\<close>
fun is_voip :: "voip_states \<Rightarrow> address \<Rightarrow> address \<Rightarrow> address \<Rightarrow> id \<Rightarrow> port \<Rightarrow>
port \<Rightarrow> (adr\<^sub>i\<^sub>p, address voip_msg) history \<Rightarrow> bool"
@ -297,11 +297,11 @@ where
is_voip S1 s d g i p1 p2 InL)))))) x)"
text{*
Finally, @{text "NB_voip"} returns the set of protocol traces which
text\<open>
Finally, \<open>NB_voip\<close> returns the set of protocol traces which
correspond to a correct protocol run given the three addresses, the
ID, and the two dynamic ports.
*}
\<close>
definition
NB_voip :: "address \<Rightarrow> address \<Rightarrow> address \<Rightarrow> id \<Rightarrow> port \<Rightarrow> port \<Rightarrow>
(adr\<^sub>i\<^sub>p, address voip_msg) history set" where

View File

@ -35,7 +35,7 @@
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*****************************************************************************)
chapter {* UPF Firewall *}
chapter \<open>UPF Firewall\<close>
theory
"UPF-Firewall"
imports
@ -44,7 +44,7 @@ theory
"FWNormalisation/FWNormalisation"
"StatefulFW/StatefulFW"
begin
text{*
text\<open>
This is the main entry point for specifications of firewall policies.
*}
\<close>
end

View File

@ -9,7 +9,7 @@
# {\providecommand{\isbn}{\textsc{isbn}} }
# {\providecommand{\Cpp}{C++} }
# {\providecommand{\Specsharp}{Spec\#} }
# {\providecommand{\doi}[1]{\href{http://dx.doi.org/#1}{doi:
# {\providecommand{\doi}[1]{\href{https://doi.org/#1}{doi:
{\urlstyle{rm}\nolinkurl{#1}}}}} }
@STRING{conf-tphols="\acs{tphols}" }
@STRING{iso = {International Organization for Standardization} }

View File

@ -172,7 +172,7 @@
\input{Transformation01.tex}
\input{Transformation02.tex}
\input{NAT-FW.tex}
\input{VoIP.tex}
\input{Voice_over_IP.tex}
% </session>
%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%