Theory PersonalFirewallDatatype
subsection ‹Personal Firewall: Datatype›
theory
PersonalFirewallDatatype
imports
"../../UPF-Firewall"
begin
text‹
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.
›
datatype Adr = pc | internet
type_synonym DatatypeTwoNets = "Adr × int"
instance Adr::adr ..
definition
PC :: "DatatypeTwoNets net" where
"PC = {{(a,b). a = pc}}"
definition
Internet :: "DatatypeTwoNets net" where
"Internet = {{(a,b). a = internet}}"
definition
not_in_same_net :: "(DatatypeTwoNets,DummyContent) packet ⇒ bool" where
"not_in_same_net x = ((src x ⊏ PC ⟶ dest x ⊏ Internet) ∧ (src x ⊏ Internet ⟶ dest x ⊏ PC))"
text ‹
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.
›
definition
strictPolicy :: "(DatatypeTwoNets,DummyContent) FWPolicy" where
"strictPolicy = deny_all ++ allow_all_from_to PC Internet"
definition
PortPolicy :: "(DatatypeTwoNets,'b) FWPolicy" where
"PortPolicy = deny_all ++ allow_from_ports_to {80::port,24,21} PC Internet"
definition
PortPolicyBig :: "(DatatypeTwoNets,'b) FWPolicy" where
"PortPolicyBig =
allow_from_port_to (80::port) PC Internet ⨁
allow_from_port_to (24::port) PC Internet ⨁
allow_from_port_to (21::port) PC Internet ⨁
deny_all"
lemmas policyLemmas = strictPolicy_def PortPolicy_def PC_def Internet_def PortPolicyBig_def src_def
PolicyCombinators PortCombinators in_subnet_def
end