(***************************************************************************** * Copyright (c) 2005-2010 ETH Zurich, Switzerland * 2008-2015 Achim D. Brucker, Germany * 2009-2017 Université Paris-Sud, France * 2015-2017 The University of Sheffield, UK * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *****************************************************************************) subsection\Packets and Networks\ theory NetworkCore imports Main begin text\ 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. 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. \ text\The ID is an integer:\ type_synonym id = int text\ 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: \ class adr type_synonym '\ src = "'\" type_synonym '\ dest = "'\" instance int ::adr .. instance nat ::adr .. instance "fun" :: (adr,adr) adr .. instance prod :: (adr,adr) adr .. text\ The content is also specified with an unconstrained generic type: \ type_synonym '\ content = "'\" text \ 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: \ datatype DummyContent = data text\Finally, a packet is:\ type_synonym ('\,'\) packet = "id \ '\ src \ '\ dest \ '\ content" text\ 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. \ definition src :: "('\::adr,'\) packet \ '\" where "src = fst o snd " text\ 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. \ class port instance int ::port .. instance nat :: port .. instance "fun" :: (port,port) port .. instance "prod" :: (port,port) port .. text\ 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: \ definition id :: "('\::adr,'\) packet \ id" where "id = fst" definition dest :: "('\::adr,'\) packet \ '\ dest" where "dest = fst o snd o snd" definition content :: "('\::adr,'\) packet \ '\ content" where "content = snd o snd o snd" datatype protocol = tcp | udp lemma either: "\a \ tcp;a \ udp\ \ False" by (case_tac "a",simp_all) lemma either2[simp]: "(a \ tcp) = (a = udp)" by (case_tac "a",simp_all) lemma either3[simp]: "(a \ udp) = (a = tcp)" by (case_tac "a",simp_all) text\ 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. \ consts src_port :: "('\::adr,'\) packet \ '\::port" consts dest_port :: "('\::adr,'\) packet \ '\::port" consts src_protocol :: "('\::adr,'\) packet \ protocol" consts dest_protocol :: "('\::adr,'\) packet \ protocol" text\A subnetwork (or simply a network) is a set of sets of addresses.\ type_synonym '\ net = "'\ set set" text\The relation {in\_subnet} (\\\) checks if an address is in a specific network.\ definition in_subnet :: "'\::adr \ '\ net \ bool" (infixl "\" 100) where "in_subnet a S = (\ s \ S. a \ s)" text\The following lemmas will be useful later.\ lemma in_subnet: "(a, e) \ {{(x1,y). P x1 y}} = P a e" by (simp add: in_subnet_def) lemma src_in_subnet: "src(q,(a,e),r,t) \ {{(x1,y). P x1 y}} = P a e" by (simp add: in_subnet_def in_subnet src_def) lemma dest_in_subnet: "dest (q,r,((a),e),t) \ {{(x1,y). P x1 y}} = P a e" by (simp add: in_subnet_def in_subnet dest_def) text\ Address models should provide a definition for the following constant, returning a network consisting of the input address only. \ consts subnet_of :: "'\::adr \ '\ net" lemmas packet_defs = in_subnet_def id_def content_def src_def dest_def end