Theory ProtocolPortCombinators
subsection ‹Policy Combinators with Ports and Protocols›
theory
ProtocolPortCombinators
imports
PortCombinators
begin
text‹
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.
This theory requires from the network models a definition for the two following constants:
\begin{itemize}
\item $src\_port :: ('\alpha,'\beta) packet \Rightarrow ('\gamma::port)$
\item $dest\_port :: ('\alpha,'\beta) packet \Rightarrow ('\gamma::port)$
\end{itemize}
›
definition
allow_all_from_port_prot :: "protocol ⇒ 'α::adr net ⇒ ('γ::port) ⇒ (('α,'β) packet ↦ unit)" where
"allow_all_from_port_prot p src_net s_port =
{pa. dest_protocol pa = p} ◃ allow_all_from_port src_net s_port"
definition
deny_all_from_port_prot :: "protocol =>'α::adr net ⇒ 'γ::port ⇒ (('α,'β) packet ↦ unit)" where
"deny_all_from_port_prot p src_net s_port =
{pa. dest_protocol pa = p} ◃ deny_all_from_port src_net s_port"
definition
allow_all_to_port_prot :: "protocol =>'α::adr net ⇒ 'γ::port ⇒ (('α,'β) packet ↦ unit)" where
"allow_all_to_port_prot p dest_net d_port =
{pa. dest_protocol pa = p} ◃ allow_all_to_port dest_net d_port"
definition
deny_all_to_port_prot :: "protocol =>'α::adr net ⇒ 'γ::port ⇒ (('α,'β) packet ↦ unit)" where
"deny_all_to_port_prot p dest_net d_port =
{pa. dest_protocol pa = p} ◃ deny_all_to_port dest_net d_port"
definition
allow_all_from_port_to_prot:: "protocol =>'α::adr net ⇒ 'γ::port ⇒ 'α::adr net ⇒ (('α,'β) packet ↦ unit)"
where
"allow_all_from_port_to_prot p src_net s_port dest_net =
{pa. dest_protocol pa = p} ◃ allow_all_from_port_to src_net s_port dest_net"
definition
deny_all_from_port_to_prot::"protocol ⇒ 'α::adr net ⇒ 'γ::port ⇒ 'α::adr net ⇒ (('α,'β) packet ↦ unit)"
where
"deny_all_from_port_to_prot p src_net s_port dest_net =
{pa. dest_protocol pa = p} ◃ deny_all_from_port_to src_net s_port dest_net"
definition
allow_all_from_port_to_port_prot::"protocol ⇒ 'α::adr net ⇒ 'γ::port ⇒ 'α::adr net ⇒ 'γ::port ⇒
(('α,'β) packet ↦ unit)" where
"allow_all_from_port_to_port_prot p src_net s_port dest_net d_port =
{pa. dest_protocol pa = p} ◃ allow_all_from_port_to_port src_net s_port dest_net d_port "
definition
deny_all_from_port_to_port_prot :: "protocol =>'α::adr net ⇒ 'γ::port ⇒ 'α::adr net ⇒
'γ::port ⇒ (('α,'β) packet ↦ unit)" where
"deny_all_from_port_to_port_prot p src_net s_port dest_net d_port =
{pa. dest_protocol pa = p} ◃ deny_all_from_port_to_port src_net s_port dest_net d_port"
definition
allow_all_from_to_port_prot :: "protocol =>'α::adr net ⇒ 'α::adr net ⇒
'γ::port ⇒ (('α,'β) packet ↦ unit)" where
"allow_all_from_to_port_prot p src_net dest_net d_port =
{pa. dest_protocol pa = p} ◃ allow_all_from_to_port src_net dest_net d_port "
definition
deny_all_from_to_port_prot :: "protocol =>'α::adr net ⇒ 'α::adr net ⇒ 'γ::port ⇒
(('α,'β) packet ↦ unit)" where
"deny_all_from_to_port_prot p src_net dest_net d_port =
{pa. dest_protocol pa = p} ◃ deny_all_from_to_port src_net dest_net d_port"
definition
allow_from_port_to_prot :: "protocol =>'γ::port ⇒ 'α::adr net ⇒ 'α::adr net ⇒ (('α,'β) packet ↦ unit)"
where
"allow_from_port_to_prot p port src_net dest_net =
{pa. dest_protocol pa = p} ◃ allow_from_port_to port src_net dest_net"
definition
deny_from_port_to_prot :: "protocol =>'γ::port ⇒ 'α::adr net ⇒ 'α::adr net ⇒ (('α,'β) packet ↦ unit)"
where
"deny_from_port_to_prot p port src_net dest_net =
{pa. dest_protocol pa = p} ◃ deny_from_port_to port src_net dest_net"
definition
allow_from_to_port_prot :: "protocol =>'γ::port ⇒ 'α::adr net ⇒ 'α::adr net ⇒ (('α,'β) packet ↦ unit)"
where
"allow_from_to_port_prot p port src_net dest_net =
{pa. dest_protocol pa = p} ◃ allow_from_to_port port src_net dest_net"
definition
deny_from_to_port_prot :: "protocol =>'γ::port ⇒ 'α::adr net ⇒ 'α::adr net ⇒ (('α,'β) packet ↦ unit)"
where
"deny_from_to_port_prot p port src_net dest_net =
{pa. dest_protocol pa = p} ◃ deny_from_to_port port src_net dest_net"
definition
allow_from_ports_to_prot :: "protocol =>'γ::port set ⇒ 'α::adr net ⇒ 'α::adr net ⇒
(('α,'β) packet ↦ unit)" where
"allow_from_ports_to_prot p ports src_net dest_net =
{pa. dest_protocol pa = p} ◃ allow_from_ports_to ports src_net dest_net"
definition
allow_from_to_ports_prot :: "protocol =>'γ::port set ⇒ 'α::adr net ⇒ 'α::adr net ⇒
(('α,'β) packet ↦ unit)" where
"allow_from_to_ports_prot p ports src_net dest_net =
{pa. dest_protocol pa = p} ◃ allow_from_to_ports ports src_net dest_net"
definition
deny_from_ports_to_prot :: "protocol =>'γ::port set ⇒ 'α::adr net ⇒ 'α::adr net ⇒
(('α,'β) packet ↦ unit)" where
"deny_from_ports_to_prot p ports src_net dest_net =
{pa. dest_protocol pa = p} ◃ deny_from_ports_to ports src_net dest_net"
definition
deny_from_to_ports_prot :: "protocol =>'γ::port set ⇒ 'α::adr net ⇒ 'α::adr net ⇒
(('α,'β) packet ↦ unit)" where
"deny_from_to_ports_prot p ports src_net dest_net =
{pa. dest_protocol pa = p} ◃ deny_from_to_ports ports src_net dest_net"
text‹As before, we put all the rules into one lemma to ease writing later.›
lemmas ProtocolCombinatorsCore =
allow_all_from_port_prot_def deny_all_from_port_prot_def allow_all_to_port_prot_def
deny_all_to_port_prot_def allow_all_from_to_port_prot_def
deny_all_from_to_port_prot_def
allow_from_ports_to_prot_def allow_from_to_ports_prot_def
deny_from_ports_to_prot_def deny_from_to_ports_prot_def
allow_all_from_port_to_prot_def deny_all_from_port_to_prot_def
allow_from_port_to_prot_def allow_from_to_port_prot_def deny_from_to_port_prot_def
deny_from_port_to_prot_def
lemmas ProtocolCombinators = PortCombinators.PortCombinators ProtocolCombinatorsCore
end