Theory IPv6
theory IPv6
imports
IP_Address
NumberWang_IPv6
begin
section ‹IPv6 Addresses›
text‹An IPv6 address is basically a 128 bit unsigned integer. RFC 4291, Section 2.›
type_synonym ipv6addr = "128 word"
text‹Conversion between natural numbers and IPv6 adresses›
definition nat_of_ipv6addr :: "ipv6addr ⇒ nat" where
"nat_of_ipv6addr a = unat a"
definition ipv6addr_of_nat :: "nat ⇒ ipv6addr" where
"ipv6addr_of_nat n = of_nat n"
lemma "ipv6addr_of_nat n = word_of_int (int n)"
by(simp add: ipv6addr_of_nat_def)
text‹The maximum IPv6 address›
definition max_ipv6_addr :: "ipv6addr" where
"max_ipv6_addr ≡ ipv6addr_of_nat ((2^128) - 1)"
lemma max_ipv6_addr_number: "max_ipv6_addr = 0xFFFFFFFFFFFFFFFFFFFFFFFFFFFFFFFF"
unfolding max_ipv6_addr_def ipv6addr_of_nat_def by(simp)
lemma "max_ipv6_addr = 340282366920938463463374607431768211455"
by(fact max_ipv6_addr_number)
lemma max_ipv6_addr_max_word: "max_ipv6_addr = - 1"
by(simp add: max_ipv6_addr_number)
lemma max_ipv6_addr_max: "∀a. a ≤ max_ipv6_addr"
by(simp add: max_ipv6_addr_max_word)
lemma UNIV_ipv6addrset: "UNIV = {0 .. max_ipv6_addr}"
by(simp add: max_ipv6_addr_max_word) fastforce
text‹identity functions›
lemma nat_of_ipv6addr_ipv6addr_of_nat_mod: "nat_of_ipv6addr (ipv6addr_of_nat n) = n mod 2^128"
by (simp add: ipv6addr_of_nat_def nat_of_ipv6addr_def unat_of_nat take_bit_eq_mod)
lemma nat_of_ipv6addr_ipv6addr_of_nat:
"n ≤ nat_of_ipv6addr max_ipv6_addr ⟹ nat_of_ipv6addr (ipv6addr_of_nat n) = n"
by (simp add: nat_of_ipv6addr_ipv6addr_of_nat_mod max_ipv6_addr_def)
lemma ipv6addr_of_nat_nat_of_ipv6addr: "ipv6addr_of_nat (nat_of_ipv6addr addr) = addr"
by(simp add: ipv6addr_of_nat_def nat_of_ipv6addr_def)
subsection‹Syntax of IPv6 Adresses›
text‹RFC 4291, Section 2.2.: Text Representation of Addresses›
text‹Quoting the RFC (note: errata exists):›
text_raw‹
\begin{verbatim}
1. The preferred form is x:x:x:x:x:x:x:x, where the 'x's are one to
four hexadecimal digits of the eight 16-bit pieces of the address.
Examples:
ABCD:EF01:2345:6789:ABCD:EF01:2345:6789
2001:DB8:0:0:8:800:200C:417A
\end{verbatim}
›