Theory IntegerPort_TCPUDP
subsection ‹Integer Addresses with Ports and Protocols›
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.›
type_synonym
address = int
type_synonym
port = int
type_synonym
adr⇩i⇩p⇩p = "address × port × protocol"
instance protocol :: adr ..
overloading src_port_int_TCPUDP ≡ "src_port :: ('α::adr,'β) packet ⇒ 'γ::port"
begin
definition
"src_port_int_TCPUDP (x::(adr⇩i⇩p⇩p,'β) packet) ≡ (fst o snd o fst o snd) x"
end
overloading dest_port_int_TCPUDP ≡ "dest_port :: ('α::adr,'β) packet ⇒ 'γ::port"
begin
definition
"dest_port_int_TCPUDP (x::(adr⇩i⇩p⇩p,'β) packet) ≡ (fst o snd o fst o snd o snd) x"
end
overloading subnet_of_int_TCPUDP ≡ "subnet_of :: 'α::adr ⇒ 'α net"
begin
definition
"subnet_of_int_TCPUDP (x::(adr⇩i⇩p⇩p)) ≡ {{(a,b,c). a = fst x}}::adr⇩i⇩p⇩p net"
end
overloading src_protocol_int_TCPUDP ≡ "src_protocol :: ('α::adr,'β) packet ⇒ protocol"
begin
definition
"src_protocol_int_TCPUDP (x::(adr⇩i⇩p⇩p,'β) packet) ≡ (snd o snd o fst o snd) x"
end
overloading dest_protocol_int_TCPUDP ≡ "dest_protocol :: ('α::adr,'β) packet ⇒ protocol"
begin
definition
"dest_protocol_int_TCPUDP (x::(adr⇩i⇩p⇩p,'β) packet) ≡ (snd o snd o fst o snd o snd) x"
end
lemma src_port: "src_port (a,x::adr⇩i⇩p⇩p,d,e) = fst (snd x)"
by (simp add: src_port_int_TCPUDP_def in_subnet)
lemma dest_port: "dest_port (a,d,x::adr⇩i⇩p⇩p,e) = fst (snd x)"
by (simp add: dest_port_int_TCPUDP_def in_subnet)
text ‹Common test constraints:›
definition port_positive :: "(adr⇩i⇩p⇩p,'b) packet ⇒ bool" where
"port_positive x = (dest_port x > (0::port))"
definition fix_values :: "(adr⇩i⇩p⇩p,DummyContent) packet ⇒ bool" where
"fix_values x = (src_port x = (1::port) ∧ src_protocol x = udp ∧ content x = data ∧ id x = 1)"
lemmas adr⇩i⇩p⇩pLemmas = src_port dest_port src_port_int_TCPUDP_def dest_port_int_TCPUDP_def
src_protocol_int_TCPUDP_def dest_protocol_int_TCPUDP_def subnet_of_int_TCPUDP_def
lemmas adr⇩i⇩p⇩pTestConstraints = port_positive_def fix_values_def
end